--- a/NEWS Mon Oct 09 22:03:05 2017 +0200
+++ b/NEWS Mon Oct 09 22:08:05 2017 +0200
@@ -44,6 +44,21 @@
higher-order quantifiers. An 'smt_explicit_application' option has been
added to control this. INCOMPATIBILITY.
+* Theory HOL-Computational_Algebra.Polynomial_Factorial does not
+provide instances of rat, real, complex as factorial rings etc.
+Import HOL-Computational_Algebra.Field_as_Ring explicitly in case of
+need. INCOMPATIBILITY.
+
+* Fact mod_mult_self4 (on nat) renamed to mod_mult_self3', to avoid clash
+with fact mod_mult_self4 (on more generic semirings). INCOMPATIBILITY.
+
+* Facts sum.commute(_restrict) and prod.commute(_restrict) renamed to
+sum.swap(_restrict) and prod.swap(_restrict), to avoid name clashes
+on interpretation of abstract locales. INCOMPATIBILITY.
+
+* Predicate pairwise_coprime abolished, use "pairwise coprime" instead.
+INCOMPATIBILITY.
+
*** System ***
--- a/src/HOL/Analysis/Bochner_Integration.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy Mon Oct 09 22:08:05 2017 +0200
@@ -368,7 +368,7 @@
if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))"
by (auto intro!: sum.cong simp: scaleR_sum_left)
also have "\<dots> = ?r"
- by (subst sum.commute)
+ by (subst sum.swap)
(auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
finally show "simple_bochner_integral M f = ?r" .
qed
--- a/src/HOL/Analysis/Bounded_Linear_Function.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Mon Oct 09 22:08:05 2017 +0200
@@ -387,7 +387,7 @@
have "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b
= (\<Sum>j\<in>Basis. if j = b then (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j))) else 0)"
using b
- by (simp add: inner_sum_left inner_Basis if_distrib cong: if_cong) (simp add: sum.commute)
+ by (simp add: inner_sum_left inner_Basis if_distrib cong: if_cong) (simp add: sum.swap)
also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"
using b by (simp add: sum.delta)
also have "\<dots> = f x \<bullet> b"
--- a/src/HOL/Analysis/Caratheodory.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Caratheodory.thy Mon Oct 09 22:08:05 2017 +0200
@@ -44,7 +44,7 @@
qed
also have "\<dots> = (SUP p. \<Sum>i<p. \<Sum>n. f (i, n))"
unfolding suminf_sum[OF summableI, symmetric]
- by (simp add: suminf_eq_SUP SUP_pair sum.commute[of _ "{..< fst _}"])
+ by (simp add: suminf_eq_SUP SUP_pair sum.swap[of _ "{..< fst _}"])
finally show ?thesis unfolding g_def
by (simp add: suminf_eq_SUP)
qed
@@ -592,7 +592,7 @@
assume "\<Union>C = \<Union>D"
with split_sum[OF C D] split_sum[OF D C]
have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
- by (simp, subst sum.commute, simp add: ac_simps) }
+ by (simp, subst sum.swap, simp add: ac_simps) }
note sum_eq = this
{ fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Oct 09 22:08:05 2017 +0200
@@ -518,14 +518,14 @@
lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply simp
done
lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
apply (vector matrix_matrix_mult_def matrix_vector_mult_def
sum_distrib_left sum_distrib_right mult.assoc)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply simp
done
@@ -556,7 +556,7 @@
lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply simp
done
@@ -659,7 +659,7 @@
apply (rule adjoint_unique)
apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
sum_distrib_right sum_distrib_left)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply (auto simp add: ac_simps)
done
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Oct 09 22:08:05 2017 +0200
@@ -180,7 +180,7 @@
have "f differentiable at x within ({a<..<c} - s)"
apply (rule differentiable_at_withinI)
using x le st
- by (metis (no_types, lifting) DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost le st(3) x)
+ by (metis (no_types, lifting) DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost_real le st(3) x)
moreover have "open ({a<..<c} - s)"
by (blast intro: open_greaterThanLessThan \<open>finite s\<close> finite_imp_closed)
ultimately show "f differentiable at x within {a..b}"
@@ -192,7 +192,7 @@
have "g differentiable at x within ({c<..<b} - t)"
apply (rule differentiable_at_withinI)
using x ge st
- by (metis DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost)
+ by (metis DiffD1 DiffD2 DiffI UnCI atLeastAtMost_diff_ends atLeastAtMost_iff at_within_interior insert_iff interior_atLeastAtMost_real)
moreover have "open ({c<..<b} - t)"
by (blast intro: open_greaterThanLessThan \<open>finite t\<close> finite_imp_closed)
ultimately show "g differentiable at x within {a..b}"
@@ -1446,7 +1446,7 @@
show ?thesis
apply (rule fundamental_theorem_of_calculus_interior_strong)
using k assms cfg *
- apply (auto simp: at_within_closed_interval)
+ apply (auto simp: at_within_Icc_at)
done
qed
@@ -4158,7 +4158,7 @@
subsection\<open>Winding number is zero "outside" a curve, in various senses\<close>
-lemma winding_number_zero_in_outside:
+proposition winding_number_zero_in_outside:
assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)"
shows "winding_number \<gamma> z = 0"
proof -
@@ -4210,7 +4210,11 @@
finally show ?thesis .
qed
-lemma winding_number_zero_outside:
+corollary winding_number_zero_const: "a \<noteq> z \<Longrightarrow> winding_number (\<lambda>t. a) z = 0"
+ by (rule winding_number_zero_in_outside)
+ (auto simp: pathfinish_def pathstart_def path_polynomial_function)
+
+corollary winding_number_zero_outside:
"\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
--- a/src/HOL/Analysis/Complex_Transcendental.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1233,7 +1233,7 @@
lemma continuous_on_Ln [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on s Ln"
by (simp add: continuous_at_imp_continuous_on continuous_within_Ln)
-lemma holomorphic_on_Ln: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> Ln holomorphic_on s"
+lemma holomorphic_on_Ln [holomorphic_intros]: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> Ln holomorphic_on s"
by (simp add: field_differentiable_within_Ln holomorphic_on_def)
lemma divide_ln_mono:
--- a/src/HOL/Analysis/Conformal_Mappings.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Conformal_Mappings.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1454,6 +1454,15 @@
qed
qed
+corollary Schwarz_Lemma':
+ assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
+ and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1"
+ shows "((\<forall>\<xi>. norm \<xi> < 1 \<longrightarrow> norm (f \<xi>) \<le> norm \<xi>) \<and> norm(deriv f 0) \<le> 1) \<and>
+ (((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z) \<or> norm(deriv f 0) = 1)
+ \<longrightarrow> (\<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1))"
+ using Schwarz_Lemma [OF assms]
+ by (metis (no_types) norm_eq_zero zero_less_one)
+
subsection\<open>The Schwarz reflection principle\<close>
lemma hol_pal_lem0:
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Oct 09 22:08:05 2017 +0200
@@ -4696,6 +4696,15 @@
finally show ?thesis .
qed
+lemma interior_atLeastLessThan [simp]:
+ fixes a::real shows "interior {a..<b} = {a<..<b}"
+ by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_semiline)
+
+lemma interior_lessThanAtMost [simp]:
+ fixes a::real shows "interior {a<..b} = {a<..<b}"
+ by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int
+ interior_interior interior_real_semiline)
+
lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}"
by (metis interior_atLeastAtMost_real interior_interior)
--- a/src/HOL/Analysis/Determinants.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Determinants.thy Mon Oct 09 22:08:05 2017 +0200
@@ -29,7 +29,7 @@
lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)"
apply (simp add: trace_def matrix_matrix_mult_def)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply (simp add: mult.commute)
done
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Oct 09 22:08:05 2017 +0200
@@ -6467,7 +6467,7 @@
qed
have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x))
(at x)" if "x \<in> {a..b} - (s \<union> {a,b})" for x
- using deriv[of x] that by (simp add: at_within_closed_interval o_def)
+ using deriv[of x] that by (simp add: at_within_Icc_at o_def)
have "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}"
using le cont_int s deriv cont_int
by (intro fundamental_theorem_of_calculus_interior_strong[of "s \<union> {a,b}"]) simp_all
--- a/src/HOL/Analysis/Linear_Algebra.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1886,7 +1886,7 @@
have "(\<Sum>x\<in>P. norm (f x)) \<le> (\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>)"
by (rule sum_mono) (rule norm_le_l1)
also have "(\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>) = (\<Sum>b\<in>Basis. \<Sum>x\<in>P. \<bar>f x \<bullet> b\<bar>)"
- by (rule sum.commute)
+ by (rule sum.swap)
also have "\<dots> \<le> of_nat (card (Basis :: 'n set)) * (2 * e)"
proof (rule sum_bounded_above)
fix i :: 'n
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Mon Oct 09 22:08:05 2017 +0200
@@ -648,7 +648,7 @@
if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
by (auto intro!: sum.cong simp: sum_distrib_left)
also have "\<dots> = ?r"
- by (subst sum.commute)
+ by (subst sum.swap)
(auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
finally show "integral\<^sup>S M f = ?r" .
qed
--- a/src/HOL/Analysis/Path_Connected.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy Mon Oct 09 22:08:05 2017 +0200
@@ -166,6 +166,9 @@
lemma simple_path_eq_arc: "pathfinish g \<noteq> pathstart g \<Longrightarrow> (simple_path g = arc g)"
by (simp add: arc_simple_path)
+lemma path_image_const [simp]: "path_image (\<lambda>t. a) = {a}"
+ by (force simp: path_image_def)
+
lemma path_image_nonempty [simp]: "path_image g \<noteq> {}"
unfolding path_image_def image_is_empty box_eq_empty
by auto
@@ -1511,13 +1514,7 @@
assumes "convex s"
shows "path_connected s"
unfolding path_connected_def
- apply rule
- apply rule
- apply (rule_tac x = "linepath x y" in exI)
- unfolding path_image_linepath
- using assms [unfolded convex_contains_segment]
- apply auto
- done
+ using assms convex_contains_segment by fastforce
lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)"
by (simp add: convex_imp_path_connected)
@@ -1528,13 +1525,7 @@
lemma path_connected_imp_connected:
assumes "path_connected S"
shows "connected S"
- unfolding connected_def not_ex
- apply rule
- apply rule
- apply (rule ccontr)
- unfolding not_not
- apply (elim conjE)
-proof -
+proof (rule connectedI)
fix e1 e2
assume as: "open e1" "open e2" "S \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> S = {}" "e1 \<inter> S \<noteq> {}" "e2 \<inter> S \<noteq> {}"
then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> S" "x2 \<in> e2 \<inter> S"
@@ -1570,31 +1561,14 @@
fix y
assume as: "y \<in> path_component_set S x"
then have "y \<in> S"
- apply -
- apply (rule path_component_mem(2))
- unfolding mem_Collect_eq
- apply auto
- done
+ by (simp add: path_component_mem(2))
then obtain e where e: "e > 0" "ball y e \<subseteq> S"
using assms[unfolded open_contains_ball]
by auto
- show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x"
- apply (rule_tac x=e in exI)
- apply (rule,rule \<open>e>0\<close>)
- apply rule
- unfolding mem_ball mem_Collect_eq
- proof -
- fix z
- assume "dist y z < e"
- 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)])
- apply (rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format])
- using \<open>e > 0\<close> as
- apply auto
- done
- qed
+have "\<And>u. dist y u < e \<Longrightarrow> path_component S x u"
+ by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component)
+ then show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x"
+ using \<open>e>0\<close> by auto
qed
lemma open_non_path_component:
@@ -1904,6 +1878,8 @@
using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast
+subsection\<open>Path components\<close>
+
lemma Union_path_component [simp]:
"Union {path_component_set S x |x. x \<in> S} = S"
apply (rule subset_antisym)
@@ -2151,7 +2127,6 @@
by (auto simp: path_connected_component)
qed
-
lemma connected_complement_bounded_convex:
fixes s :: "'a :: euclidean_space set"
assumes "bounded s" "convex s" "2 \<le> DIM('a)"
@@ -2300,6 +2275,65 @@
qed
+subsection\<open>Every annulus is a connected set\<close>
+
+lemma path_connected_2DIM_I:
+ fixes a :: "'N::euclidean_space"
+ assumes 2: "2 \<le> DIM('N)" and pc: "path_connected {r. 0 \<le> r \<and> P r}"
+ shows "path_connected {x. P(norm(x - a))}"
+proof -
+ have "{x. P(norm(x - a))} = op+a ` {x. P(norm x)}"
+ by force
+ moreover have "path_connected {x::'N. P(norm x)}"
+ proof -
+ let ?D = "{x. 0 \<le> x \<and> P x} \<times> sphere (0::'N) 1"
+ have "x \<in> (\<lambda>z. fst z *\<^sub>R snd z) ` ?D"
+ if "P (norm x)" for x::'N
+ proof (cases "x=0")
+ case True
+ with that show ?thesis
+ apply (simp add: image_iff)
+ apply (rule_tac x=0 in exI, simp)
+ using vector_choose_size zero_le_one by blast
+ next
+ case False
+ with that show ?thesis
+ by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto
+ qed
+ then have *: "{x::'N. P(norm x)} = (\<lambda>z. fst z *\<^sub>R snd z) ` ?D"
+ by auto
+ have "continuous_on ?D (\<lambda>z:: real\<times>'N. fst z *\<^sub>R snd z)"
+ by (intro continuous_intros)
+ moreover have "path_connected ?D"
+ by (metis path_connected_Times [OF pc] path_connected_sphere 2)
+ ultimately show ?thesis
+ apply (subst *)
+ apply (rule path_connected_continuous_image, auto)
+ done
+ qed
+ ultimately show ?thesis
+ using path_connected_translation by metis
+qed
+
+lemma path_connected_annulus:
+ fixes a :: "'N::euclidean_space"
+ assumes "2 \<le> DIM('N)"
+ shows "path_connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
+ "path_connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
+ "path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
+ "path_connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
+ by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms])
+
+lemma connected_annulus:
+ fixes a :: "'N::euclidean_space"
+ assumes "2 \<le> DIM('N::euclidean_space)"
+ shows "connected {x. r1 < norm(x - a) \<and> norm(x - a) < r2}"
+ "connected {x. r1 < norm(x - a) \<and> norm(x - a) \<le> r2}"
+ "connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) < r2}"
+ "connected {x. r1 \<le> norm(x - a) \<and> norm(x - a) \<le> r2}"
+ by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected)
+
+
subsection\<open>Relations between components and path components\<close>
lemma open_connected_component:
@@ -2894,11 +2928,21 @@
shows "outside s = - s"
by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2)
+lemma outside_singleton [simp]:
+ fixes x :: "'a :: {real_normed_vector, perfect_space}"
+ shows "outside {x} = -{x}"
+ by (auto simp: outside_convex)
+
lemma inside_convex:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
shows "convex s \<Longrightarrow> inside s = {}"
by (simp add: inside_outside outside_convex)
+lemma inside_singleton [simp]:
+ fixes x :: "'a :: {real_normed_vector, perfect_space}"
+ shows "inside {x} = {}"
+ by (auto simp: inside_convex)
+
lemma outside_subset_convex:
fixes s :: "'a :: {real_normed_vector, perfect_space} set"
shows "\<lbrakk>convex t; s \<subseteq> t\<rbrakk> \<Longrightarrow> - t \<subseteq> outside s"
@@ -4119,6 +4163,39 @@
by (blast intro: homotopic_loops_trans)
qed
+lemma homotopic_paths_loop_parts:
+ assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q"
+ shows "homotopic_paths S p q"
+proof -
+ have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))"
+ using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp
+ then have "path p"
+ using \<open>path q\<close> homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast
+ show ?thesis
+ proof (cases "pathfinish p = pathfinish q")
+ case True
+ have pipq: "path_image p \<subseteq> S" "path_image q \<subseteq> S"
+ by (metis Un_subset_iff paths \<open>path p\<close> \<open>path q\<close> homotopic_loops_imp_subset homotopic_paths_imp_path loops
+ path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+
+ have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))"
+ using \<open>path p\<close> \<open>path_image p \<subseteq> S\<close> homotopic_paths_rid homotopic_paths_sym by blast
+ moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))"
+ by (simp add: True \<open>path p\<close> \<open>path q\<close> pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym)
+ moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)"
+ by (simp add: True \<open>path p\<close> \<open>path q\<close> homotopic_paths_assoc pipq)
+ moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)"
+ by (simp add: \<open>path q\<close> homotopic_paths_join paths pipq)
+ moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q"
+ by (metis \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2))
+ ultimately show ?thesis
+ using homotopic_paths_trans by metis
+ next
+ case False
+ then show ?thesis
+ using \<open>path q\<close> homotopic_loops_imp_path loops path_join_path_ends by fastforce
+ qed
+qed
+
subsection\<open> Homotopy of "nearby" function, paths and loops.\<close>
--- a/src/HOL/Analysis/Starlike.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Starlike.thy Mon Oct 09 22:08:05 2017 +0200
@@ -3776,24 +3776,6 @@
subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close>
-lemma interior_atLeastAtMost [simp]:
- fixes a::real shows "interior {a..b} = {a<..<b}"
- using interior_cbox [of a b] by auto
-
-lemma interior_atLeastLessThan [simp]:
- fixes a::real shows "interior {a..<b} = {a<..<b}"
- by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost interior_Int interior_interior interior_real_semiline)
-
-lemma interior_lessThanAtMost [simp]:
- fixes a::real shows "interior {a<..b} = {a<..<b}"
- by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost interior_Int
- interior_interior interior_real_semiline)
-
-lemma at_within_closed_interval:
- fixes x::real
- shows "a < x \<Longrightarrow> x < b \<Longrightarrow> (at x within {a..b}) = at x"
- by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost)
-
lemma at_within_cbox_finite:
assumes "x \<in> box a b" "x \<notin> S" "finite S"
shows "(at x within cbox a b - S) = at x"
@@ -5087,7 +5069,6 @@
using separation_closures [of S T]
by (metis assms closure_closed disjnt_def inf_commute)
-
lemma separation_normal_local:
fixes S :: "'a::euclidean_space set"
assumes US: "closedin (subtopology euclidean U) S"
@@ -5147,6 +5128,139 @@
by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl)
qed
+subsection\<open>Connectedness of the intersection of a chain\<close>
+
+proposition connected_chain:
+ fixes \<F> :: "'a :: euclidean_space set set"
+ assumes cc: "\<And>S. S \<in> \<F> \<Longrightarrow> compact S \<and> connected S"
+ and linear: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
+ shows "connected(\<Inter>\<F>)"
+proof (cases "\<F> = {}")
+ case True then show ?thesis
+ by auto
+next
+ case False
+ then have cf: "compact(\<Inter>\<F>)"
+ by (simp add: cc compact_Inter)
+ have False if AB: "closed A" "closed B" "A \<inter> B = {}"
+ and ABeq: "A \<union> B = \<Inter>\<F>" and "A \<noteq> {}" "B \<noteq> {}" for A B
+ proof -
+ obtain U V where "open U" "open V" "A \<subseteq> U" "B \<subseteq> V" "U \<inter> V = {}"
+ using separation_normal [OF AB] by metis
+ obtain K where "K \<in> \<F>" "compact K"
+ using cc False by blast
+ then obtain N where "open N" and "K \<subseteq> N"
+ by blast
+ let ?\<C> = "insert (U \<union> V) ((\<lambda>S. N - S) ` \<F>)"
+ obtain \<D> where "\<D> \<subseteq> ?\<C>" "finite \<D>" "K \<subseteq> \<Union>\<D>"
+ proof (rule compactE [OF \<open>compact K\<close>])
+ show "K \<subseteq> \<Union>insert (U \<union> V) (op - N ` \<F>)"
+ using \<open>K \<subseteq> N\<close> ABeq \<open>A \<subseteq> U\<close> \<open>B \<subseteq> V\<close> by auto
+ show "\<And>B. B \<in> insert (U \<union> V) (op - N ` \<F>) \<Longrightarrow> open B"
+ by (auto simp: \<open>open U\<close> \<open>open V\<close> open_Un \<open>open N\<close> cc compact_imp_closed open_Diff)
+ qed
+ then have "finite(\<D> - {U \<union> V})"
+ by blast
+ moreover have "\<D> - {U \<union> V} \<subseteq> (\<lambda>S. N - S) ` \<F>"
+ using \<open>\<D> \<subseteq> ?\<C>\<close> by blast
+ ultimately obtain \<G> where "\<G> \<subseteq> \<F>" "finite \<G>" and Deq: "\<D> - {U \<union> V} = (\<lambda>S. N-S) ` \<G>"
+ using finite_subset_image by metis
+ obtain J where "J \<in> \<F>" and J: "(\<Union>S\<in>\<G>. N - S) \<subseteq> N - J"
+ proof (cases "\<G> = {}")
+ case True
+ with \<open>\<F> \<noteq> {}\<close> that show ?thesis
+ by auto
+ next
+ case False
+ have "\<And>S T. \<lbrakk>S \<in> \<G>; T \<in> \<G>\<rbrakk> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
+ by (meson \<open>\<G> \<subseteq> \<F>\<close> in_mono local.linear)
+ with \<open>finite \<G>\<close> \<open>\<G> \<noteq> {}\<close>
+ have "\<exists>J \<in> \<G>. (\<Union>S\<in>\<G>. N - S) \<subseteq> N - J"
+ proof induction
+ case (insert X \<H>)
+ show ?case
+ proof (cases "\<H> = {}")
+ case True then show ?thesis by auto
+ next
+ case False
+ then have "\<And>S T. \<lbrakk>S \<in> \<H>; T \<in> \<H>\<rbrakk> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
+ by (simp add: insert.prems)
+ with insert.IH False obtain J where "J \<in> \<H>" and J: "(\<Union>Y\<in>\<H>. N - Y) \<subseteq> N - J"
+ by metis
+ have "N - J \<subseteq> N - X \<or> N - X \<subseteq> N - J"
+ by (meson Diff_mono \<open>J \<in> \<H>\<close> insert.prems(2) insert_iff order_refl)
+ then show ?thesis
+ proof
+ assume "N - J \<subseteq> N - X" with J show ?thesis
+ by auto
+ next
+ assume "N - X \<subseteq> N - J"
+ with J have "N - X \<union> UNION \<H> (op - N) \<subseteq> N - J"
+ by auto
+ with \<open>J \<in> \<H>\<close> show ?thesis
+ by blast
+ qed
+ qed
+ qed simp
+ with \<open>\<G> \<subseteq> \<F>\<close> show ?thesis by (blast intro: that)
+ qed
+ have "K \<subseteq> \<Union>(insert (U \<union> V) (\<D> - {U \<union> V}))"
+ using \<open>K \<subseteq> \<Union>\<D>\<close> by auto
+ also have "... \<subseteq> (U \<union> V) \<union> (N - J)"
+ by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1)
+ finally have "J \<inter> K \<subseteq> U \<union> V"
+ by blast
+ moreover have "connected(J \<inter> K)"
+ by (metis Int_absorb1 \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> cc inf.orderE local.linear)
+ moreover have "U \<inter> (J \<inter> K) \<noteq> {}"
+ using ABeq \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> \<open>A \<noteq> {}\<close> \<open>A \<subseteq> U\<close> by blast
+ moreover have "V \<inter> (J \<inter> K) \<noteq> {}"
+ using ABeq \<open>J \<in> \<F>\<close> \<open>K \<in> \<F>\<close> \<open>B \<noteq> {}\<close> \<open>B \<subseteq> V\<close> by blast
+ ultimately show False
+ using connectedD [of "J \<inter> K" U V] \<open>open U\<close> \<open>open V\<close> \<open>U \<inter> V = {}\<close> by auto
+ qed
+ with cf show ?thesis
+ by (auto simp: connected_closed_set compact_imp_closed)
+qed
+
+lemma connected_chain_gen:
+ fixes \<F> :: "'a :: euclidean_space set set"
+ assumes X: "X \<in> \<F>" "compact X"
+ and cc: "\<And>T. T \<in> \<F> \<Longrightarrow> closed T \<and> connected T"
+ and linear: "\<And>S T. S \<in> \<F> \<and> T \<in> \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
+ shows "connected(\<Inter>\<F>)"
+proof -
+ have "\<Inter>\<F> = (\<Inter>T\<in>\<F>. X \<inter> T)"
+ using X by blast
+ moreover have "connected (\<Inter>T\<in>\<F>. X \<inter> T)"
+ proof (rule connected_chain)
+ show "\<And>T. T \<in> op \<inter> X ` \<F> \<Longrightarrow> compact T \<and> connected T"
+ using cc X by auto (metis inf.absorb2 inf.orderE local.linear)
+ show "\<And>S T. S \<in> op \<inter> X ` \<F> \<and> T \<in> op \<inter> X ` \<F> \<Longrightarrow> S \<subseteq> T \<or> T \<subseteq> S"
+ using local.linear by blast
+ qed
+ ultimately show ?thesis
+ by metis
+qed
+
+lemma connected_nest:
+ fixes S :: "'a::linorder \<Rightarrow> 'b::euclidean_space set"
+ assumes S: "\<And>n. compact(S n)" "\<And>n. connected(S n)"
+ and nest: "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
+ shows "connected(\<Inter> (range S))"
+ apply (rule connected_chain)
+ using S apply blast
+ by (metis image_iff le_cases nest)
+
+lemma connected_nest_gen:
+ fixes S :: "'a::linorder \<Rightarrow> 'b::euclidean_space set"
+ assumes S: "\<And>n. closed(S n)" "\<And>n. connected(S n)" "compact(S k)"
+ and nest: "\<And>m n. m \<le> n \<Longrightarrow> S n \<subseteq> S m"
+ shows "connected(\<Inter> (range S))"
+ apply (rule connected_chain_gen [of "S k"])
+ using S apply auto
+ by (meson le_cases nest subsetCE)
+
subsection\<open>Proper maps, including projections out of compact sets\<close>
lemma finite_indexed_bound:
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Oct 09 22:08:05 2017 +0200
@@ -676,6 +676,10 @@
using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute)
qed
+lemma openin_Inter [intro]:
+ assumes "finite \<F>" "\<F> \<noteq> {}" "\<And>X. X \<in> \<F> \<Longrightarrow> openin T X" shows "openin T (\<Inter>\<F>)"
+ by (metis (full_types) assms openin_INT2 image_ident)
+
subsubsection \<open>Closed sets\<close>
@@ -2404,6 +2408,11 @@
lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (subtopology euclidean S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S"
by (meson closedin_limpt closed_subset closedin_closed_trans)
+lemma connected_closed_set:
+ "closed S
+ \<Longrightarrow> connected S \<longleftrightarrow> (\<nexists>A B. closed A \<and> closed B \<and> A \<noteq> {} \<and> B \<noteq> {} \<and> A \<union> B = S \<and> A \<inter> B = {})"
+ unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast
+
lemma closedin_subset_trans:
"closedin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
closedin (subtopology euclidean T) S"
@@ -5407,6 +5416,13 @@
by auto
qed
+lemma compact_Inter:
+ fixes \<F> :: "'a :: heine_borel set set"
+ assumes com: "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" and "\<F> \<noteq> {}"
+ shows "compact(\<Inter> \<F>)"
+ using assms
+ by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed)
+
lemma compact_closure [simp]:
fixes S :: "'a::heine_borel set"
shows "compact(closure S) \<longleftrightarrow> bounded S"
@@ -5819,19 +5835,6 @@
then show ?thesis unfolding complete_def by auto
qed
-lemma nat_approx_posE:
- fixes e::real
- assumes "0 < e"
- obtains n :: nat where "1 / (Suc n) < e"
-proof atomize_elim
- have "1 / real (Suc (nat \<lceil>1/e\<rceil>)) < 1 / \<lceil>1/e\<rceil>"
- by (rule divide_strict_left_mono) (auto simp: \<open>0 < e\<close>)
- also have "1 / \<lceil>1/e\<rceil> \<le> 1 / (1/e)"
- by (rule divide_left_mono) (auto simp: \<open>0 < e\<close> ceiling_correct)
- also have "\<dots> = e" by simp
- finally show "\<exists>n. 1 / real (Suc n) < e" ..
-qed
-
lemma compact_eq_totally_bounded:
"compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>x\<in>k. ball x e))"
(is "_ \<longleftrightarrow> ?rhs")
@@ -9679,6 +9682,43 @@
simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono)
done
+lemma homeomorphic_ball01_UNIV:
+ "ball (0::'a::real_normed_vector) 1 homeomorphic (UNIV:: 'a set)"
+ (is "?B homeomorphic ?U")
+proof
+ have "x \<in> (\<lambda>z. z /\<^sub>R (1 - norm z)) ` ball 0 1" for x::'a
+ apply (rule_tac x="x /\<^sub>R (1 + norm x)" in image_eqI)
+ apply (auto simp: divide_simps)
+ using norm_ge_zero [of x] apply linarith+
+ done
+ then show "(\<lambda>z::'a. z /\<^sub>R (1 - norm z)) ` ?B = ?U"
+ by blast
+ have "x \<in> range (\<lambda>z. (1 / (1 + norm z)) *\<^sub>R z)" if "norm x < 1" for x::'a
+ apply (rule_tac x="x /\<^sub>R (1 - norm x)" in image_eqI)
+ using that apply (auto simp: divide_simps)
+ done
+ then show "(\<lambda>z::'a. z /\<^sub>R (1 + norm z)) ` ?U = ?B"
+ by (force simp: divide_simps dest: add_less_zeroD)
+ show "continuous_on (ball 0 1) (\<lambda>z. z /\<^sub>R (1 - norm z))"
+ by (rule continuous_intros | force)+
+ show "continuous_on UNIV (\<lambda>z. z /\<^sub>R (1 + norm z))"
+ apply (intro continuous_intros)
+ apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one)
+ done
+ show "\<And>x. x \<in> ball 0 1 \<Longrightarrow>
+ x /\<^sub>R (1 - norm x) /\<^sub>R (1 + norm (x /\<^sub>R (1 - norm x))) = x"
+ by (auto simp: divide_simps)
+ show "\<And>y. y /\<^sub>R (1 + norm y) /\<^sub>R (1 - norm (y /\<^sub>R (1 + norm y))) = y"
+ apply (auto simp: divide_simps)
+ apply (metis le_add_same_cancel1 norm_ge_zero not_le zero_less_one)
+ done
+qed
+
+proposition homeomorphic_ball_UNIV:
+ fixes a ::"'a::real_normed_vector"
+ assumes "0 < r" shows "ball a r homeomorphic (UNIV:: 'a set)"
+ using assms homeomorphic_ball01_UNIV homeomorphic_balls(1) homeomorphic_trans zero_less_one by blast
+
subsection\<open>Inverse function property for open/closed maps\<close>
lemma continuous_on_inverse_open_map:
@@ -11043,38 +11083,38 @@
text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
lemma continuous_disconnected_range_constant:
- assumes s: "connected s"
- and conf: "continuous_on s f"
- and fim: "f ` s \<subseteq> t"
+ assumes S: "connected S"
+ and conf: "continuous_on S f"
+ and fim: "f ` S \<subseteq> t"
and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
- shows "\<exists>a. \<forall>x \<in> s. f x = a"
-proof (cases "s = {}")
+ shows "\<exists>a. \<forall>x \<in> S. f x = a"
+proof (cases "S = {}")
case True then show ?thesis by force
next
case False
- { fix x assume "x \<in> s"
- then have "f ` s \<subseteq> {f x}"
- by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct)
+ { fix x assume "x \<in> S"
+ then have "f ` S \<subseteq> {f x}"
+ by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct)
}
with False show ?thesis
by blast
qed
lemma discrete_subset_disconnected:
- fixes s :: "'a::topological_space set"
+ fixes S :: "'a::topological_space set"
fixes t :: "'b::real_normed_vector set"
- assumes conf: "continuous_on s f"
- and no: "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
- shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}"
-proof -
- { fix x assume x: "x \<in> s"
- then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> s; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)"
+ assumes conf: "continuous_on S f"
+ and no: "\<And>x. x \<in> S \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
+ shows "f ` S \<subseteq> {y. connected_component_set (f ` S) y = {y}}"
+proof -
+ { fix x assume x: "x \<in> S"
+ then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> S; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)"
using conf no [OF x] by auto
then have e2: "0 \<le> e / 2"
by simp
- have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y
+ have "f y = f x" if "y \<in> S" and ccs: "f y \<in> connected_component_set (f ` S) (f x)" for y
apply (rule ccontr)
- using connected_closed [of "connected_component_set (f ` s) (f x)"] \<open>e>0\<close>
+ using connected_closed [of "connected_component_set (f ` S) (f x)"] \<open>e>0\<close>
apply (simp add: del: ex_simps)
apply (drule spec [where x="cball (f x) (e / 2)"])
apply (drule spec [where x="- ball(f x) e"])
@@ -11082,11 +11122,11 @@
apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less)
using centre_in_cball connected_component_refl_eq e2 x apply blast
using ccs
- apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>])
+ apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> S\<close>])
done
- moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
+ moreover have "connected_component_set (f ` S) (f x) \<subseteq> f ` S"
by (auto simp: connected_component_in)
- ultimately have "connected_component_set (f ` s) (f x) = {f x}"
+ ultimately have "connected_component_set (f ` S) (f x) = {f x}"
by (auto simp: x)
}
with assms show ?thesis
@@ -11094,22 +11134,22 @@
qed
lemma finite_implies_discrete:
- fixes s :: "'a::topological_space set"
- assumes "finite (f ` s)"
- shows "(\<forall>x \<in> s. \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))"
-proof -
- have "\<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> s" for x
- proof (cases "f ` s - {f x} = {}")
+ fixes S :: "'a::topological_space set"
+ assumes "finite (f ` S)"
+ shows "(\<forall>x \<in> S. \<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))"
+proof -
+ have "\<exists>e>0. \<forall>y. y \<in> S \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> S" for x
+ proof (cases "f ` S - {f x} = {}")
case True
with zero_less_numeral show ?thesis
by (fastforce simp add: Set.image_subset_iff cong: conj_cong)
next
case False
- then obtain z where z: "z \<in> s" "f z \<noteq> f x"
+ then obtain z where z: "z \<in> S" "f z \<noteq> f x"
by blast
- have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}"
+ have finn: "finite {norm (z - f x) |z. z \<in> f ` S - {f x}}"
using assms by simp
- then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}"
+ then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` S - {f x}}"
apply (rule finite_imp_less_Inf)
using z apply force+
done
@@ -11123,20 +11163,20 @@
text\<open>This proof requires the existence of two separate values of the range type.\<close>
lemma finite_range_constant_imp_connected:
assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
- \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a"
- shows "connected s"
+ \<lbrakk>continuous_on S f; finite(f ` S)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> S. f x = a"
+ shows "connected S"
proof -
{ fix t u
- assume clt: "closedin (subtopology euclidean s) t"
- and clu: "closedin (subtopology euclidean s) u"
- and tue: "t \<inter> u = {}" and tus: "t \<union> u = s"
- have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)"
+ assume clt: "closedin (subtopology euclidean S) t"
+ and clu: "closedin (subtopology euclidean S) u"
+ and tue: "t \<inter> u = {}" and tus: "t \<union> u = S"
+ have conif: "continuous_on S (\<lambda>x. if x \<in> t then 0 else 1)"
apply (subst tus [symmetric])
apply (rule continuous_on_cases_local)
using clt clu tue
apply (auto simp: tus continuous_on_const)
done
- have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)"
+ have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` S)"
by (rule finite_subset [of _ "{0,1}"]) auto
have "t = {} \<or> u = {}"
using assms [OF conif fi] tus [symmetric]
--- a/src/HOL/Archimedean_Field.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Archimedean_Field.thy Mon Oct 09 22:08:05 2017 +0200
@@ -526,7 +526,8 @@
shows "finite {k \<in> \<int>. \<bar>k\<bar> \<le> a}"
using finite_int_segment [of "-a" a] by (auto simp add: abs_le_iff conj_commute minus_le_iff)
-text \<open>Ceiling with numerals.\<close>
+
+subsubsection \<open>Ceiling with numerals.\<close>
lemma ceiling_zero [simp]: "\<lceil>0\<rceil> = 0"
using ceiling_of_int [of 0] by simp
@@ -595,7 +596,7 @@
by (simp add: ceiling_altdef)
-text \<open>Addition and subtraction of integers.\<close>
+subsubsection \<open>Addition and subtraction of integers.\<close>
lemma ceiling_add_of_int [simp]: "\<lceil>x + of_int z\<rceil> = \<lceil>x\<rceil> + z"
using ceiling_correct [of x] by (simp add: ceiling_def)
@@ -630,6 +631,24 @@
unfolding of_int_less_iff by simp
qed
+lemma nat_approx_posE:
+ fixes e:: "'a::{archimedean_field,floor_ceiling}"
+ assumes "0 < e"
+ obtains n :: nat where "1 / of_nat(Suc n) < e"
+proof
+ have "(1::'a) / of_nat (Suc (nat \<lceil>1/e\<rceil>)) < 1 / of_int (\<lceil>1/e\<rceil>)"
+ proof (rule divide_strict_left_mono)
+ show "(of_int \<lceil>1 / e\<rceil>::'a) < of_nat (Suc (nat \<lceil>1 / e\<rceil>))"
+ using assms by (simp add: field_simps)
+ show "(0::'a) < of_nat (Suc (nat \<lceil>1 / e\<rceil>)) * of_int \<lceil>1 / e\<rceil>"
+ using assms by (auto simp: zero_less_mult_iff pos_add_strict)
+ qed auto
+ also have "1 / of_int (\<lceil>1/e\<rceil>) \<le> 1 / (1/e)"
+ by (rule divide_left_mono) (auto simp: \<open>0 < e\<close> ceiling_correct)
+ also have "\<dots> = e" by simp
+ finally show "1 / of_nat (Suc (nat \<lceil>1 / e\<rceil>)) < e"
+ by metis
+qed
subsection \<open>Negation\<close>
--- a/src/HOL/Bali/Basis.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Bali/Basis.thy Mon Oct 09 22:08:05 2017 +0200
@@ -4,7 +4,7 @@
subsection \<open>Definitions extending HOL as logical basis of Bali\<close>
theory Basis
-imports Main "HOL-Library.Old_Recdef"
+imports Main
begin
subsubsection "misc"
--- a/src/HOL/Binomial.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Binomial.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1011,7 +1011,7 @@
by (subst binomial_fact_lemma [symmetric]) auto
lemma choose_dvd:
- "k \<le> n \<Longrightarrow> fact k * fact (n - k) dvd (fact n :: 'a::{semiring_div,linordered_semidom})"
+ "k \<le> n \<Longrightarrow> fact k * fact (n - k) dvd (fact n :: 'a::linordered_semidom)"
unfolding dvd_def
apply (rule exI [where x="of_nat (n choose k)"])
using binomial_fact_lemma [of k n, THEN arg_cong [where f = of_nat and 'b='a]]
@@ -1019,7 +1019,7 @@
done
lemma fact_fact_dvd_fact:
- "fact k * fact n dvd (fact (k + n) :: 'a::{semiring_div,linordered_semidom})"
+ "fact k * fact n dvd (fact (k + n) :: 'a::linordered_semidom)"
by (metis add.commute add_diff_cancel_left' choose_dvd le_add2)
lemma choose_mult_lemma:
--- a/src/HOL/Code_Numeral.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Code_Numeral.thy Mon Oct 09 22:08:05 2017 +0200
@@ -220,48 +220,51 @@
"of_nat (nat_of_integer k) = max 0 k"
by transfer auto
-instantiation integer :: normalization_semidom
+instantiation integer :: unique_euclidean_ring
begin
-lift_definition normalize_integer :: "integer \<Rightarrow> integer"
- is "normalize :: int \<Rightarrow> int"
- .
-
-declare normalize_integer.rep_eq [simp]
-
-lift_definition unit_factor_integer :: "integer \<Rightarrow> integer"
- is "unit_factor :: int \<Rightarrow> int"
- .
-
-declare unit_factor_integer.rep_eq [simp]
-
lift_definition divide_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
is "divide :: int \<Rightarrow> int \<Rightarrow> int"
.
declare divide_integer.rep_eq [simp]
-
-instance
- by (standard; transfer)
- (auto simp add: mult_sgn_abs sgn_mult abs_eq_iff')
-end
-
-instantiation integer :: ring_div
-begin
-
lift_definition modulo_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer"
is "modulo :: int \<Rightarrow> int \<Rightarrow> int"
.
declare modulo_integer.rep_eq [simp]
+lift_definition euclidean_size_integer :: "integer \<Rightarrow> nat"
+ is "euclidean_size :: int \<Rightarrow> nat"
+ .
+
+declare euclidean_size_integer.rep_eq [simp]
+
+lift_definition uniqueness_constraint_integer :: "integer \<Rightarrow> integer \<Rightarrow> bool"
+ is "uniqueness_constraint :: int \<Rightarrow> int \<Rightarrow> bool"
+ .
+
+declare uniqueness_constraint_integer.rep_eq [simp]
+
instance
- by (standard; transfer) simp_all
+ by (standard; transfer)
+ (use mult_le_mono2 [of 1] in \<open>auto simp add: sgn_mult_abs abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\<close>, rule div_eqI, simp_all)
end
-instantiation integer :: semiring_numeral_div
+lemma [code]:
+ "euclidean_size = nat_of_integer \<circ> abs"
+ by (simp add: fun_eq_iff nat_of_integer.rep_eq)
+
+lemma [code]:
+ "uniqueness_constraint (k :: integer) l \<longleftrightarrow> sgn k = sgn l"
+ by (simp add: integer_eq_iff)
+
+instance integer :: ring_parity
+ by (standard; transfer) (simp_all add: of_nat_div odd_iff_mod_2_eq_one)
+
+instantiation integer :: unique_euclidean_semiring_numeral
begin
definition divmod_integer :: "num \<Rightarrow> num \<Rightarrow> integer \<times> integer"
@@ -283,15 +286,15 @@
by (fact divmod_step_integer_def)
qed (transfer,
fact le_add_diff_inverse2
- semiring_numeral_div_class.div_less
- semiring_numeral_div_class.mod_less
- semiring_numeral_div_class.div_positive
- semiring_numeral_div_class.mod_less_eq_dividend
- semiring_numeral_div_class.pos_mod_bound
- semiring_numeral_div_class.pos_mod_sign
- semiring_numeral_div_class.mod_mult2_eq
- semiring_numeral_div_class.div_mult2_eq
- semiring_numeral_div_class.discrete)+
+ unique_euclidean_semiring_numeral_class.div_less
+ unique_euclidean_semiring_numeral_class.mod_less
+ unique_euclidean_semiring_numeral_class.div_positive
+ unique_euclidean_semiring_numeral_class.mod_less_eq_dividend
+ unique_euclidean_semiring_numeral_class.pos_mod_bound
+ unique_euclidean_semiring_numeral_class.pos_mod_sign
+ unique_euclidean_semiring_numeral_class.mod_mult2_eq
+ unique_euclidean_semiring_numeral_class.div_mult2_eq
+ unique_euclidean_semiring_numeral_class.discrete)+
end
@@ -434,23 +437,15 @@
"Neg m * Neg n = Pos (m * n)"
by simp_all
-lemma normalize_integer_code [code]:
- "normalize = (abs :: integer \<Rightarrow> integer)"
- by transfer simp
-
-lemma unit_factor_integer_code [code]:
- "unit_factor = (sgn :: integer \<Rightarrow> integer)"
- by transfer simp
-
definition divmod_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer \<times> integer"
where
"divmod_integer k l = (k div l, k mod l)"
-lemma fst_divmod [simp]:
+lemma fst_divmod_integer [simp]:
"fst (divmod_integer k l) = k div l"
by (simp add: divmod_integer_def)
-lemma snd_divmod [simp]:
+lemma snd_divmod_integer [simp]:
"snd (divmod_integer k l) = k mod l"
by (simp add: divmod_integer_def)
@@ -853,21 +848,9 @@
"nat_of_natural (max k l) = max (nat_of_natural k) (nat_of_natural l)"
by transfer rule
-instantiation natural :: "{semiring_div, normalization_semidom}"
+instantiation natural :: unique_euclidean_semiring
begin
-lift_definition normalize_natural :: "natural \<Rightarrow> natural"
- is "normalize :: nat \<Rightarrow> nat"
- .
-
-declare normalize_natural.rep_eq [simp]
-
-lift_definition unit_factor_natural :: "natural \<Rightarrow> natural"
- is "unit_factor :: nat \<Rightarrow> nat"
- .
-
-declare unit_factor_natural.rep_eq [simp]
-
lift_definition divide_natural :: "natural \<Rightarrow> natural \<Rightarrow> natural"
is "divide :: nat \<Rightarrow> nat \<Rightarrow> nat"
.
@@ -880,11 +863,35 @@
declare modulo_natural.rep_eq [simp]
+lift_definition euclidean_size_natural :: "natural \<Rightarrow> nat"
+ is "euclidean_size :: nat \<Rightarrow> nat"
+ .
+
+declare euclidean_size_natural.rep_eq [simp]
+
+lift_definition uniqueness_constraint_natural :: "natural \<Rightarrow> natural \<Rightarrow> bool"
+ is "uniqueness_constraint :: nat \<Rightarrow> nat \<Rightarrow> bool"
+ .
+
+declare uniqueness_constraint_natural.rep_eq [simp]
+
instance
- by standard (transfer, auto simp add: algebra_simps unit_factor_nat_def gr0_conv_Suc)+
+ by (standard; transfer)
+ (auto simp add: algebra_simps unit_factor_nat_def gr0_conv_Suc)
end
+lemma [code]:
+ "euclidean_size = nat_of_natural"
+ by (simp add: fun_eq_iff)
+
+lemma [code]:
+ "uniqueness_constraint = (\<top> :: natural \<Rightarrow> natural \<Rightarrow> bool)"
+ by (simp add: fun_eq_iff)
+
+instance natural :: semiring_parity
+ by (standard; transfer) (simp_all add: of_nat_div odd_iff_mod_2_eq_one)
+
lift_definition natural_of_integer :: "integer \<Rightarrow> natural"
is "nat :: int \<Rightarrow> nat"
.
@@ -1020,31 +1027,6 @@
"integer_of_natural (m * n) = integer_of_natural m * integer_of_natural n"
by transfer simp
-lemma [code]:
- "normalize n = n" for n :: natural
- by transfer simp
-
-lemma [code]:
- "unit_factor n = of_bool (n \<noteq> 0)" for n :: natural
-proof (cases "n = 0")
- case True
- then show ?thesis
- by simp
-next
- case False
- then have "unit_factor n = 1"
- proof transfer
- fix n :: nat
- assume "n \<noteq> 0"
- then obtain m where "n = Suc m"
- by (cases n) auto
- then show "unit_factor n = 1"
- by simp
- qed
- with False show ?thesis
- by simp
-qed
-
lemma [code abstract]:
"integer_of_natural (m div n) = integer_of_natural m div integer_of_natural n"
by transfer (simp add: zdiv_int)
--- a/src/HOL/Complex.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Complex.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1077,7 +1077,7 @@
and Complex_mult_complex_of_real: "Complex x y * complex_of_real r = Complex (x*r) (y*r)"
and complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)"
and complex_eq_cancel_iff2: "(Complex x y = complex_of_real xa) = (x = xa \<and> y = 0)"
- and complex_cn: "cnj (Complex a b) = Complex a (- b)"
+ and complex_cnj: "cnj (Complex a b) = Complex a (- b)"
and Complex_sum': "sum (\<lambda>x. Complex (f x) 0) s = Complex (sum f s) 0"
and Complex_sum: "Complex (sum f s) 0 = sum (\<lambda>x. Complex (f x) 0) s"
and complex_of_real_def: "complex_of_real r = Complex r 0"
--- a/src/HOL/Computational_Algebra/Computational_Algebra.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Computational_Algebra.thy Mon Oct 09 22:08:05 2017 +0200
@@ -12,6 +12,7 @@
Nth_Powers
Polynomial_FPS
Polynomial
+ Polynomial_Factorial
Primes
Squarefree
begin
--- a/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Euclidean_Algorithm.thy Mon Oct 09 22:08:05 2017 +0200
@@ -9,10 +9,26 @@
begin
subsection \<open>Generic construction of the (simple) euclidean algorithm\<close>
-
-context euclidean_semiring
+
+class normalization_euclidean_semiring = euclidean_semiring + normalization_semidom
begin
+lemma euclidean_size_normalize [simp]:
+ "euclidean_size (normalize a) = euclidean_size a"
+proof (cases "a = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case [simp]: False
+ have "euclidean_size (normalize a) \<le> euclidean_size (normalize a * unit_factor a)"
+ by (rule size_mult_mono) simp
+ moreover have "euclidean_size a \<le> euclidean_size (a * (1 div unit_factor a))"
+ by (rule size_mult_mono) simp
+ ultimately show ?thesis
+ by simp
+qed
+
context
begin
@@ -282,7 +298,7 @@
subsection \<open>The (simple) euclidean algorithm as gcd computation\<close>
-class euclidean_semiring_gcd = euclidean_semiring + gcd + Gcd +
+class euclidean_semiring_gcd = normalization_euclidean_semiring + gcd + Gcd +
assumes gcd_eucl: "Euclidean_Algorithm.gcd = GCD.gcd"
and lcm_eucl: "Euclidean_Algorithm.lcm = GCD.lcm"
assumes Gcd_eucl: "Euclidean_Algorithm.Gcd = GCD.Gcd"
@@ -400,7 +416,7 @@
end
lemma factorial_euclidean_semiring_gcdI:
- "OFCLASS('a::{factorial_semiring_gcd, euclidean_semiring}, euclidean_semiring_gcd_class)"
+ "OFCLASS('a::{factorial_semiring_gcd, normalization_euclidean_semiring}, euclidean_semiring_gcd_class)"
proof
interpret semiring_Gcd 1 0 times
Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm
@@ -502,7 +518,7 @@
also have "r' - r' div r * r = r' mod r" using div_mult_mod_eq [of r' r]
by (simp add: algebra_simps)
finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" .
- qed (auto simp: gcd_mod_right algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute)
+ qed (auto simp: algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute)
finally show ?thesis .
qed
qed
@@ -552,6 +568,8 @@
subsection \<open>Typical instances\<close>
+instance nat :: normalization_euclidean_semiring ..
+
instance nat :: euclidean_semiring_gcd
proof
interpret semiring_Gcd 1 0 times
@@ -584,6 +602,8 @@
by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
qed
+instance int :: normalization_euclidean_semiring ..
+
instance int :: euclidean_ring_gcd
proof
interpret semiring_Gcd 1 0 times
--- a/src/HOL/Computational_Algebra/Field_as_Ring.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Field_as_Ring.thy Mon Oct 09 22:08:05 2017 +0200
@@ -24,7 +24,7 @@
end
-instantiation real :: unique_euclidean_ring
+instantiation real :: "{unique_euclidean_ring, normalization_euclidean_semiring}"
begin
definition [simp]: "normalize_real = (normalize_field :: real \<Rightarrow> _)"
@@ -55,7 +55,7 @@
end
-instantiation rat :: unique_euclidean_ring
+instantiation rat :: "{unique_euclidean_ring, normalization_euclidean_semiring}"
begin
definition [simp]: "normalize_rat = (normalize_field :: rat \<Rightarrow> _)"
@@ -86,7 +86,7 @@
end
-instantiation complex :: unique_euclidean_ring
+instantiation complex :: "{unique_euclidean_ring, normalization_euclidean_semiring}"
begin
definition [simp]: "normalize_complex = (normalize_field :: complex \<Rightarrow> _)"
--- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1181,7 +1181,7 @@
end
-instantiation fps :: (field) ring_div
+instantiation fps :: (field) idom_modulo
begin
definition fps_mod_def:
@@ -1224,39 +1224,6 @@
from assms show "f mod g = 0" by (intro fps_mod_eq_zero) auto
qed
-context
-begin
-private lemma fps_divide_cancel_aux1:
- assumes "h$0 \<noteq> (0 :: 'a :: field)"
- shows "(h * f) div (h * g) = f div g"
-proof (cases "g = 0")
- assume "g \<noteq> 0"
- from assms have "h \<noteq> 0" by auto
- note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
- from assms have [simp]: "subdegree h = 0" by (simp add: subdegree_eq_0_iff)
-
- have "(h * f) div (h * g) =
- fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
- by (simp add: fps_divide_def Let_def)
- also have "h * f * inverse (fps_shift (subdegree g) (h*g)) =
- (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
- by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
- also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
- finally show "(h * f) div (h * g) = f div g" by (simp_all add: fps_divide_def Let_def)
-qed (simp_all add: fps_divide_def)
-
-private lemma fps_divide_cancel_aux2:
- "(f * fps_X^m) div (g * fps_X^m) = f div (g :: 'a :: field fps)"
-proof (cases "g = 0")
- assume [simp]: "g \<noteq> 0"
- have "(f * fps_X^m) div (g * fps_X^m) =
- fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*fps_X^m))*fps_X^m)"
- by (simp add: fps_divide_def Let_def algebra_simps)
- also have "... = f div g"
- by (simp add: fps_shift_times_fps_X_power'' fps_divide_def Let_def)
- finally show ?thesis .
-qed (simp_all add: fps_divide_def)
-
instance proof
fix f g :: "'a fps"
define n where "n = subdegree g"
@@ -1284,39 +1251,9 @@
finally show ?thesis by simp
qed
qed (auto simp: fps_mod_def fps_divide_def Let_def)
-next
-
- fix f g h :: "'a fps"
- assume "h \<noteq> 0"
- show "(h * f) div (h * g) = f div g"
- proof -
- define m where "m = subdegree h"
- define h' where "h' = fps_shift m h"
- have h_decomp: "h = h' * fps_X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
- from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
- have "(h * f) div (h * g) = (h' * f * fps_X^m) div (h' * g * fps_X^m)"
- by (simp add: h_decomp algebra_simps)
- also have "... = f div g" by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
- finally show ?thesis .
- qed
-
-next
- fix f g h :: "'a fps"
- assume [simp]: "h \<noteq> 0"
- define n h' where dfs: "n = subdegree h" "h' = fps_shift n h"
- have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
- by (simp add: fps_divide_def Let_def dfs[symmetric] algebra_simps fps_shift_add)
- also have "h * inverse h' = (inverse h' * h') * fps_X^n"
- by (subst subdegree_decompose) (simp_all add: dfs)
- also have "... = fps_X^n" by (subst inverse_mult_eq_1) (simp_all add: dfs)
- also have "fps_shift n (g * fps_X^n) = g" by simp
- also have "fps_shift n (f * inverse h') = f div h"
- by (simp add: fps_divide_def Let_def dfs)
- finally show "(f + g * h) div h = g + f div h" by simp
qed
end
-end
lemma subdegree_mod:
assumes "f \<noteq> 0" "subdegree f < subdegree g"
@@ -1411,6 +1348,40 @@
definition fps_euclidean_size_def:
"euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)"
+context
+begin
+
+private lemma fps_divide_cancel_aux1:
+ assumes "h$0 \<noteq> (0 :: 'a :: field)"
+ shows "(h * f) div (h * g) = f div g"
+proof (cases "g = 0")
+ assume "g \<noteq> 0"
+ from assms have "h \<noteq> 0" by auto
+ note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
+ from assms have [simp]: "subdegree h = 0" by (simp add: subdegree_eq_0_iff)
+
+ have "(h * f) div (h * g) =
+ fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
+ by (simp add: fps_divide_def Let_def)
+ also have "h * f * inverse (fps_shift (subdegree g) (h*g)) =
+ (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
+ by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
+ also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
+ finally show "(h * f) div (h * g) = f div g" by (simp_all add: fps_divide_def Let_def)
+qed (simp_all add: fps_divide_def)
+
+private lemma fps_divide_cancel_aux2:
+ "(f * fps_X^m) div (g * fps_X^m) = f div (g :: 'a :: field fps)"
+proof (cases "g = 0")
+ assume [simp]: "g \<noteq> 0"
+ have "(f * fps_X^m) div (g * fps_X^m) =
+ fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*fps_X^m))*fps_X^m)"
+ by (simp add: fps_divide_def Let_def algebra_simps)
+ also have "... = f div g"
+ by (simp add: fps_shift_times_fps_X_power'' fps_divide_def Let_def)
+ finally show ?thesis .
+qed (simp_all add: fps_divide_def)
+
instance proof
fix f g :: "'a fps" assume [simp]: "g \<noteq> 0"
show "euclidean_size f \<le> euclidean_size (f * g)"
@@ -1420,10 +1391,42 @@
apply (rule disjE[OF le_less_linear[of "subdegree g" "subdegree f"]])
apply (simp_all add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod)
done
+ show "(h * f) div (h * g) = f div g" if "h \<noteq> 0"
+ for f g h :: "'a fps"
+ proof -
+ define m where "m = subdegree h"
+ define h' where "h' = fps_shift m h"
+ have h_decomp: "h = h' * fps_X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
+ from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
+ have "(h * f) div (h * g) = (h' * f * fps_X^m) div (h' * g * fps_X^m)"
+ by (simp add: h_decomp algebra_simps)
+ also have "... = f div g"
+ by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
+ finally show ?thesis .
+ qed
+ show "(f + g * h) div h = g + f div h"
+ if "h \<noteq> 0" for f g h :: "'a fps"
+ proof -
+ define n h' where dfs: "n = subdegree h" "h' = fps_shift n h"
+ have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
+ by (simp add: fps_divide_def Let_def dfs [symmetric] algebra_simps fps_shift_add that)
+ also have "h * inverse h' = (inverse h' * h') * fps_X^n"
+ by (subst subdegree_decompose) (simp_all add: dfs)
+ also have "... = fps_X^n"
+ by (subst inverse_mult_eq_1) (simp_all add: dfs that)
+ also have "fps_shift n (g * fps_X^n) = g" by simp
+ also have "fps_shift n (f * inverse h') = f div h"
+ by (simp add: fps_divide_def Let_def dfs)
+ finally show ?thesis by simp
+ qed
qed (simp_all add: fps_euclidean_size_def)
end
+end
+
+instance fps :: (field) normalization_euclidean_semiring ..
+
instantiation fps :: (field) euclidean_ring_gcd
begin
definition fps_gcd_def: "(gcd :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.gcd"
@@ -3133,7 +3136,7 @@
done
also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
unfolding sum_distrib_left
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply (rule sum.cong, rule refl)+
apply simp
done
@@ -3314,7 +3317,7 @@
done
have "?r = sum (\<lambda>i. sum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
apply (simp add: fps_mult_nth sum_distrib_left)
- apply (subst sum.commute)
+ apply (subst sum.swap)
apply (rule sum.cong)
apply (auto simp add: field_simps)
done
--- a/src/HOL/Computational_Algebra/Polynomial.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Polynomial.thy Mon Oct 09 22:08:05 2017 +0200
@@ -12,6 +12,7 @@
HOL.Deriv
"HOL-Library.More_List"
"HOL-Library.Infinite_Set"
+ Factorial_Ring
begin
subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
@@ -596,6 +597,10 @@
lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c"
by (cases "c = 0") (simp_all add: degree_monom_eq)
+lemma last_coeffs_eq_coeff_degree:
+ "last (coeffs p) = lead_coeff p" if "p \<noteq> 0"
+ using that by (simp add: coeffs_def)
+
subsection \<open>Addition and subtraction\<close>
@@ -1110,7 +1115,7 @@
also have "\<dots> = (\<Sum>i\<le>k. (if n = i then c * coeff p (k - i) else 0))"
by (intro sum.cong) simp_all
also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))"
- by (simp add: sum.delta')
+ by simp
finally show ?thesis .
qed
@@ -1162,8 +1167,9 @@
lemma coeffs_map_poly':
assumes "\<And>x. x \<noteq> 0 \<Longrightarrow> f x \<noteq> 0"
shows "coeffs (map_poly f p) = map f (coeffs p)"
- using assms by (simp add: coeffs_map_poly no_trailing_map strip_while_idem_iff)
- (metis comp_def no_leading_def no_trailing_coeffs)
+ using assms
+ by (auto simp add: coeffs_map_poly strip_while_idem_iff
+ last_coeffs_eq_coeff_degree no_trailing_unfold last_map)
lemma set_coeffs_map_poly:
"(\<And>x. f x = 0 \<longleftrightarrow> x = 0) \<Longrightarrow> set (coeffs (map_poly f p)) = f ` set (coeffs p)"
@@ -1494,9 +1500,6 @@
for p :: "'a::linordered_idom poly"
by (induct p) (auto simp: pos_poly_pCons)
-lemma last_coeffs_eq_coeff_degree: "p \<noteq> 0 \<Longrightarrow> last (coeffs p) = coeff p (degree p)"
- by (simp add: coeffs_def)
-
lemma pos_poly_coeffs [code]: "pos_poly p \<longleftrightarrow> (let as = coeffs p in as \<noteq> [] \<and> last as > 0)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
@@ -2879,167 +2882,6 @@
qed
-subsection \<open>Content and primitive part of a polynomial\<close>
-
-definition content :: "'a::semiring_gcd poly \<Rightarrow> 'a"
- where "content p = gcd_list (coeffs p)"
-
-lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0"
- by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps)
-
-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: "[:c:] dvd p \<longleftrightarrow> c dvd content p"
- for c :: "'a::semiring_gcd"
-proof (cases "p = 0")
- case True
- then show ?thesis by simp
-next
- case False
- have "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)"
- by (rule const_poly_dvd_iff)
- also have "\<dots> \<longleftrightarrow> (\<forall>a\<in>set (coeffs p). c dvd a)"
- proof safe
- fix n :: nat
- assume "\<forall>a\<in>set (coeffs p). c dvd a"
- then show "c dvd coeff p n"
- by (cases "n \<le> 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 "\<dots> \<longleftrightarrow> c dvd content p"
- by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff)
- finally show ?thesis .
-qed
-
-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"
-proof (cases "p = 0")
- case True
- then show ?thesis
- by simp
-next
- case False
- then show ?thesis
- by (cases "n \<le> degree p")
- (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd)
-qed
-
-lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
- by (simp add: content_def Gcd_fin_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) \<longleftrightarrow> content p = 1"
-proof
- assume "is_unit (content p)"
- then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
- then show "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_fin_mult)
-
-lemma content_eq_zero_iff [simp]: "content p = 0 \<longleftrightarrow> p = 0"
- by (auto simp: content_def simp: poly_eq_iff coeffs_def)
-
-definition primitive_part :: "'a :: semiring_gcd poly \<Rightarrow> 'a poly"
- where "primitive_part p = map_poly (\<lambda>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]: "smult (content p) (primitive_part p) = p"
- for p :: "'a :: semiring_gcd poly"
-proof (cases "p = 0")
- case True
- then show ?thesis by simp
-next
- case False
- then show ?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
-
-lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
-proof (cases "p = 0")
- case True
- then show ?thesis by simp
-next
- case False
- then have "primitive_part p = map_poly (\<lambda>x. x div content p) p"
- by (simp add: primitive_part_def)
- also from False have "\<dots> = 0 \<longleftrightarrow> 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
-
-lemma content_primitive_part [simp]:
- assumes "p \<noteq> 0"
- shows "content (primitive_part p) = 1"
-proof -
- have "p = smult (content p) (primitive_part p)"
- by simp
- also have "content \<dots> = content (primitive_part p) * content p"
- by (simp del: content_times_primitive_part add: ac_simps)
- finally have "1 * content p = content (primitive_part p) * content p"
- by simp
- then have "1 * content p div content p = content (primitive_part p) * content p div content p"
- by simp
- with assms show ?thesis
- by simp
-qed
-
-lemma content_decompose:
- obtains p' :: "'a::semiring_gcd poly" where "p = smult (content p) p'" "content p' = 1"
-proof (cases "p = 0")
- case True
- then show ?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 (rule 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 \<Longrightarrow> 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 \<Longrightarrow> 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 True
- then show ?thesis by simp
-next
- case False
- have "p = smult (content p) (primitive_part p)"
- by simp
- also from False have "degree \<dots> = degree (primitive_part p)"
- by (subst degree_smult_eq) simp_all
- finally show ?thesis ..
-qed
-
-
subsection \<open>Division of polynomials\<close>
subsubsection \<open>Division in general\<close>
@@ -3647,16 +3489,6 @@
finally show ?thesis .
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
-
inductive eucl_rel_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly \<Rightarrow> bool"
where
eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)"
@@ -3805,40 +3637,7 @@
for x :: "'a::field poly"
by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly])
-instance poly :: (field) ring_div
-proof
- fix x y z :: "'a poly"
- assume "y \<noteq> 0"
- with eucl_rel_poly [of x y] have "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
- by (simp add: eucl_rel_poly_iff distrib_right)
- then show "(x + z * y) div y = z + x div y"
- by (rule div_poly_eq)
-next
- fix x y z :: "'a poly"
- assume "x \<noteq> 0"
- show "(x * y) div (x * z) = y div z"
- proof (cases "y \<noteq> 0 \<and> z \<noteq> 0")
- have "\<And>x::'a poly. eucl_rel_poly x 0 (0, x)"
- by (rule eucl_rel_poly_by_0)
- then have [simp]: "\<And>x::'a poly. x div 0 = 0"
- by (rule div_poly_eq)
- have "\<And>x::'a poly. eucl_rel_poly 0 x (0, 0)"
- by (rule eucl_rel_poly_0)
- then have [simp]: "\<And>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 \<noteq> 0" and "z \<noteq> 0" by auto
- with \<open>x \<noteq> 0\<close> have "\<And>q r. eucl_rel_poly y z (q, r) \<Longrightarrow> eucl_rel_poly (x * y) (x * z) (q, x * r)"
- by (auto simp: 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
+instance poly :: (field) idom_modulo ..
lemma div_pCons_eq:
"pCons a p div q =
@@ -4500,6 +4299,387 @@
qed
qed
+
+subsection \<open>Primality and irreducibility in polynomial rings\<close>
+
+lemma prod_mset_const_poly: "(\<Prod>x\<in>#A. [:f x:]) = [:prod_mset (image_mset f A):]"
+ by (induct A) (simp_all add: ac_simps)
+
+lemma irreducible_const_poly_iff:
+ fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
+ shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
+proof
+ assume A: "irreducible c"
+ show "irreducible [:c:]"
+ proof (rule irreducibleI)
+ fix a b assume ab: "[:c:] = a * b"
+ hence "degree [:c:] = degree (a * b)" by (simp only: )
+ also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
+ hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
+ finally have "degree a = 0" "degree b = 0" by auto
+ then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
+ from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
+ hence "c = a' * b'" by (simp add: ab' mult_ac)
+ from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
+ with ab' show "a dvd 1 \<or> b dvd 1"
+ by (auto simp add: is_unit_const_poly_iff)
+ qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
+next
+ assume A: "irreducible [:c:]"
+ then have "c \<noteq> 0" and "\<not> c dvd 1"
+ by (auto simp add: irreducible_def is_unit_const_poly_iff)
+ then show "irreducible c"
+ proof (rule irreducibleI)
+ fix a b assume ab: "c = a * b"
+ hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
+ from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
+ then show "a dvd 1 \<or> b dvd 1"
+ by (auto simp add: is_unit_const_poly_iff)
+ qed
+qed
+
+lemma lift_prime_elem_poly:
+ assumes "prime_elem (c :: 'a :: semidom)"
+ shows "prime_elem [:c:]"
+proof (rule prime_elemI)
+ fix a b assume *: "[:c:] dvd a * b"
+ from * have dvd: "c dvd coeff (a * b) n" for n
+ by (subst (asm) const_poly_dvd_iff) blast
+ {
+ define m where "m = (GREATEST m. \<not>c dvd coeff b m)"
+ assume "\<not>[:c:] dvd b"
+ hence A: "\<exists>i. \<not>c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast
+ have B: "\<forall>i. \<not>c dvd coeff b i \<longrightarrow> i \<le> degree b"
+ by (auto intro: le_degree)
+ have coeff_m: "\<not>c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B])
+ have "i \<le> m" if "\<not>c dvd coeff b i" for i
+ unfolding m_def by (rule Greatest_le_nat[OF that B])
+ hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force
+
+ have "c dvd coeff a i" for i
+ proof (induction i rule: nat_descend_induct[of "degree a"])
+ case (base i)
+ thus ?case by (simp add: coeff_eq_0)
+ next
+ case (descend i)
+ let ?A = "{..i+m} - {i}"
+ have "c dvd coeff (a * b) (i + m)" by (rule dvd)
+ also have "coeff (a * b) (i + m) = (\<Sum>k\<le>i + m. coeff a k * coeff b (i + m - k))"
+ by (simp add: coeff_mult)
+ also have "{..i+m} = insert i ?A" by auto
+ also have "(\<Sum>k\<in>\<dots>. coeff a k * coeff b (i + m - k)) =
+ coeff a i * coeff b m + (\<Sum>k\<in>?A. coeff a k * coeff b (i + m - k))"
+ (is "_ = _ + ?S")
+ by (subst sum.insert) simp_all
+ finally have eq: "c dvd coeff a i * coeff b m + ?S" .
+ moreover have "c dvd ?S"
+ proof (rule dvd_sum)
+ fix k assume k: "k \<in> {..i+m} - {i}"
+ show "c dvd coeff a k * coeff b (i + m - k)"
+ proof (cases "k < i")
+ case False
+ with k have "c dvd coeff a k" by (intro descend.IH) simp
+ thus ?thesis by simp
+ next
+ case True
+ hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp
+ thus ?thesis by simp
+ qed
+ qed
+ ultimately have "c dvd coeff a i * coeff b m"
+ by (simp add: dvd_add_left_iff)
+ with assms coeff_m show "c dvd coeff a i"
+ by (simp add: prime_elem_dvd_mult_iff)
+ qed
+ hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast
+ }
+ then show "[:c:] dvd a \<or> [:c:] dvd b" by blast
+next
+ from assms show "[:c:] \<noteq> 0" and "\<not> [:c:] dvd 1"
+ by (simp_all add: prime_elem_def is_unit_const_poly_iff)
+qed
+
+lemma prime_elem_const_poly_iff:
+ fixes c :: "'a :: semidom"
+ shows "prime_elem [:c:] \<longleftrightarrow> prime_elem c"
+proof
+ assume A: "prime_elem [:c:]"
+ show "prime_elem c"
+ proof (rule prime_elemI)
+ fix a b assume "c dvd a * b"
+ hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac)
+ from A and this have "[:c:] dvd [:a:] \<or> [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD)
+ thus "c dvd a \<or> c dvd b" by simp
+ qed (insert A, auto simp: prime_elem_def is_unit_poly_iff)
+qed (auto intro: lift_prime_elem_poly)
+
+
+subsection \<open>Content and primitive part of a polynomial\<close>
+
+definition content :: "'a::semiring_gcd poly \<Rightarrow> 'a"
+ where "content p = gcd_list (coeffs p)"
+
+lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0"
+ by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps)
+
+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: "[:c:] dvd p \<longleftrightarrow> c dvd content p"
+ for c :: "'a::semiring_gcd"
+proof (cases "p = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ have "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)"
+ by (rule const_poly_dvd_iff)
+ also have "\<dots> \<longleftrightarrow> (\<forall>a\<in>set (coeffs p). c dvd a)"
+ proof safe
+ fix n :: nat
+ assume "\<forall>a\<in>set (coeffs p). c dvd a"
+ then show "c dvd coeff p n"
+ by (cases "n \<le> 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 "\<dots> \<longleftrightarrow> c dvd content p"
+ by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff)
+ finally show ?thesis .
+qed
+
+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"
+proof (cases "p = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ then show ?thesis
+ by (cases "n \<le> degree p")
+ (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd)
+qed
+
+lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
+ by (simp add: content_def Gcd_fin_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) \<longleftrightarrow> content p = 1"
+proof
+ assume "is_unit (content p)"
+ then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
+ then show "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_fin_mult)
+
+lemma content_eq_zero_iff [simp]: "content p = 0 \<longleftrightarrow> p = 0"
+ by (auto simp: content_def simp: poly_eq_iff coeffs_def)
+
+definition primitive_part :: "'a :: semiring_gcd poly \<Rightarrow> 'a poly"
+ where "primitive_part p = map_poly (\<lambda>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]: "smult (content p) (primitive_part p) = p"
+ for p :: "'a :: semiring_gcd poly"
+proof (cases "p = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then show ?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
+
+lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
+proof (cases "p = 0")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then have "primitive_part p = map_poly (\<lambda>x. x div content p) p"
+ by (simp add: primitive_part_def)
+ also from False have "\<dots> = 0 \<longleftrightarrow> 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
+
+lemma content_primitive_part [simp]:
+ assumes "p \<noteq> 0"
+ shows "content (primitive_part p) = 1"
+proof -
+ have "p = smult (content p) (primitive_part p)"
+ by simp
+ also have "content \<dots> = content (primitive_part p) * content p"
+ by (simp del: content_times_primitive_part add: ac_simps)
+ finally have "1 * content p = content (primitive_part p) * content p"
+ by simp
+ then have "1 * content p div content p = content (primitive_part p) * content p div content p"
+ by simp
+ with assms show ?thesis
+ by simp
+qed
+
+lemma content_decompose:
+ obtains p' :: "'a::semiring_gcd poly" where "p = smult (content p) p'" "content p' = 1"
+proof (cases "p = 0")
+ case True
+ then show ?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 (rule 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 \<Longrightarrow> 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 \<Longrightarrow> 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 True
+ then show ?thesis by simp
+next
+ case False
+ have "p = smult (content p) (primitive_part p)"
+ by simp
+ also from False have "degree \<dots> = degree (primitive_part p)"
+ by (subst degree_smult_eq) simp_all
+ finally show ?thesis ..
+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_mult:
+ fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
+ shows "content (p * q) = content p * content q"
+proof -
+ from content_decompose[of p] guess p' . note p = this
+ from content_decompose[of q] guess q' . note q = this
+ have "content (p * q) = content p * content q * content (p' * q')"
+ by (subst p, subst q) (simp add: mult_ac normalize_mult)
+ also have "content (p' * q') = 1"
+ proof (cases "p' * q' = 0")
+ case True
+ with \<open>content p' = 1\<close> \<open>content q' = 1\<close> show ?thesis
+ by auto
+ next
+ case False
+ from \<open>content p' = 1\<close> \<open>content q' = 1\<close>
+ have "p' \<noteq> 0" "q' \<noteq> 0"
+ by auto
+ then have "p' * q' \<noteq> 0"
+ by auto
+ have "is_unit (content (p' * q'))"
+ proof (rule ccontr)
+ assume "\<not> is_unit (content (p' * q'))"
+ with False have "\<exists>p. p dvd content (p' * q') \<and> prime p"
+ by (intro prime_divisor_exists) simp_all
+ then obtain p where "p dvd content (p' * q')" "prime p"
+ by blast
+ from \<open>p dvd content (p' * q')\<close> have "[:p:] dvd p' * q'"
+ by (simp add: const_poly_dvd_iff_dvd_content)
+ moreover from \<open>prime p\<close> have "prime_elem [:p:]"
+ by (simp add: lift_prime_elem_poly)
+ ultimately have "[:p:] dvd p' \<or> [:p:] dvd q'"
+ by (simp add: prime_elem_dvd_mult_iff)
+ with \<open>content p' = 1\<close> \<open>content q' = 1\<close> have "is_unit p"
+ by (simp add: const_poly_dvd_iff_dvd_content)
+ with \<open>prime p\<close> show False
+ by simp
+ qed
+ then have "normalize (content (p' * q')) = 1"
+ by (simp add: is_unit_normalize del: normalize_content)
+ then show ?thesis
+ by simp
+ qed
+ finally show ?thesis
+ by simp
+qed
+
+lemma primitive_part_mult:
+ fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+ shows "primitive_part (p * q) = primitive_part p * primitive_part q"
+proof -
+ have "primitive_part (p * q) = p * q div [:content (p * q):]"
+ by (simp add: primitive_part_def div_const_poly_conv_map_poly)
+ also have "\<dots> = (p div [:content p:]) * (q div [:content q:])"
+ by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac)
+ also have "\<dots> = primitive_part p * primitive_part q"
+ by (simp add: primitive_part_def div_const_poly_conv_map_poly)
+ finally show ?thesis .
+qed
+
+lemma primitive_part_smult:
+ fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+ shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)"
+proof -
+ have "smult a p = [:a:] * p" by simp
+ also have "primitive_part \<dots> = smult (unit_factor a) (primitive_part p)"
+ by (subst primitive_part_mult) simp_all
+ finally show ?thesis .
+qed
+
+lemma primitive_part_dvd_primitive_partI [intro]:
+ fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
+ shows "p dvd q \<Longrightarrow> primitive_part p dvd primitive_part q"
+ by (auto elim!: dvdE simp: primitive_part_mult)
+
+lemma content_prod_mset:
+ fixes A :: "'a :: {factorial_semiring, semiring_Gcd} poly multiset"
+ shows "content (prod_mset A) = prod_mset (image_mset content A)"
+ by (induction A) (simp_all add: content_mult mult_ac)
+
+lemma content_prod_eq_1_iff:
+ fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
+ shows "content (p * q) = 1 \<longleftrightarrow> content p = 1 \<and> content q = 1"
+proof safe
+ assume A: "content (p * q) = 1"
+ {
+ fix p q :: "'a poly" assume "content p * content q = 1"
+ hence "1 = content p * content q" by simp
+ hence "content p dvd 1" by (rule dvdI)
+ hence "content p = 1" by simp
+ } note B = this
+ from A B[of p q] B [of q p] show "content p = 1" "content q = 1"
+ by (simp_all add: content_mult mult_ac)
+qed (auto simp: content_mult)
+
+
no_notation cCons (infixr "##" 65)
end
--- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1,62 +1,21 @@
(* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy
- Author: Brian Huffman
- Author: Clemens Ballarin
- Author: Amine Chaieb
- Author: Florian Haftmann
Author: Manuel Eberl
*)
+section \<open>Polynomials, fractions and rings\<close>
+
theory Polynomial_Factorial
imports
Complex_Main
Polynomial
Normalized_Fraction
- Field_as_Ring
begin
-subsection \<open>Various facts about polynomials\<close>
-
-lemma prod_mset_const_poly: " (\<Prod>x\<in>#A. [:f x:]) = [:prod_mset (image_mset f A):]"
- by (induct A) (simp_all add: ac_simps)
-
-lemma irreducible_const_poly_iff:
- fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
- shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
-proof
- assume A: "irreducible c"
- show "irreducible [:c:]"
- proof (rule irreducibleI)
- fix a b assume ab: "[:c:] = a * b"
- hence "degree [:c:] = degree (a * b)" by (simp only: )
- also from A ab have "a \<noteq> 0" "b \<noteq> 0" by auto
- hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq)
- finally have "degree a = 0" "degree b = 0" by auto
- then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE)
- from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: )
- hence "c = a' * b'" by (simp add: ab' mult_ac)
- from A and this have "a' dvd 1 \<or> b' dvd 1" by (rule irreducibleD)
- with ab' show "a dvd 1 \<or> b dvd 1"
- by (auto simp add: is_unit_const_poly_iff)
- qed (insert A, auto simp: irreducible_def is_unit_poly_iff)
-next
- assume A: "irreducible [:c:]"
- then have "c \<noteq> 0" and "\<not> c dvd 1"
- by (auto simp add: irreducible_def is_unit_const_poly_iff)
- then show "irreducible c"
- proof (rule irreducibleI)
- fix a b assume ab: "c = a * b"
- hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac)
- from A and this have "[:a:] dvd 1 \<or> [:b:] dvd 1" by (rule irreducibleD)
- then show "a dvd 1 \<or> b dvd 1"
- by (auto simp add: is_unit_const_poly_iff)
- qed
-qed
-
-
subsection \<open>Lifting elements into the field of fractions\<close>
-definition to_fract :: "'a :: idom \<Rightarrow> 'a fract" where "to_fract x = Fract x 1"
- \<comment> \<open>FIXME: name \<open>of_idom\<close>, abbreviation\<close>
+definition to_fract :: "'a :: idom \<Rightarrow> 'a fract"
+ where "to_fract x = Fract x 1"
+ \<comment> \<open>FIXME: more idiomatic name, abbreviation\<close>
lemma to_fract_0 [simp]: "to_fract 0 = 0"
by (simp add: to_fract_def eq_fract Zero_fract_def)
@@ -429,39 +388,6 @@
finally show ?thesis by simp
qed
-lemma primitive_part_mult:
- fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
- shows "primitive_part (p * q) = primitive_part p * primitive_part q"
-proof -
- have "primitive_part (p * q) = p * q div [:content (p * q):]"
- by (simp add: primitive_part_def div_const_poly_conv_map_poly)
- also have "\<dots> = (p div [:content p:]) * (q div [:content q:])"
- by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac)
- also have "\<dots> = primitive_part p * primitive_part q"
- by (simp add: primitive_part_def div_const_poly_conv_map_poly)
- finally show ?thesis .
-qed
-
-lemma primitive_part_smult:
- fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
- shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)"
-proof -
- have "smult a p = [:a:] * p" by simp
- also have "primitive_part \<dots> = smult (unit_factor a) (primitive_part p)"
- by (subst primitive_part_mult) simp_all
- finally show ?thesis .
-qed
-
-lemma primitive_part_dvd_primitive_partI [intro]:
- fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide} poly"
- shows "p dvd q \<Longrightarrow> primitive_part p dvd primitive_part q"
- by (auto elim!: dvdE simp: primitive_part_mult)
-
-lemma content_prod_mset:
- fixes A :: "'a :: {factorial_semiring, semiring_Gcd} poly multiset"
- shows "content (prod_mset A) = prod_mset (image_mset content A)"
- by (induction A) (simp_all add: content_mult mult_ac)
-
lemma fract_poly_dvdD:
fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
assumes "fract_poly p dvd fract_poly q" "content p = 1"
@@ -481,104 +407,67 @@
thus ?thesis by (rule dvdI)
qed
-lemma content_prod_eq_1_iff:
- fixes p q :: "'a :: {factorial_semiring, semiring_Gcd} poly"
- shows "content (p * q) = 1 \<longleftrightarrow> content p = 1 \<and> content q = 1"
-proof safe
- assume A: "content (p * q) = 1"
- {
- fix p q :: "'a poly" assume "content p * content q = 1"
- hence "1 = content p * content q" by simp
- hence "content p dvd 1" by (rule dvdI)
- hence "content p = 1" by simp
- } note B = this
- from A B[of p q] B [of q p] show "content p = 1" "content q = 1"
- by (simp_all add: content_mult mult_ac)
-qed (auto simp: content_mult)
-
end
subsection \<open>Polynomials over a field are a Euclidean ring\<close>
-definition unit_factor_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
- "unit_factor_field_poly p = [:lead_coeff p:]"
-
-definition normalize_field_poly :: "'a :: field poly \<Rightarrow> 'a poly" where
- "normalize_field_poly p = smult (inverse (lead_coeff p)) p"
-
-definition euclidean_size_field_poly :: "'a :: field poly \<Rightarrow> nat" where
- "euclidean_size_field_poly p = (if p = 0 then 0 else 2 ^ degree p)"
-
-lemma dvd_field_poly: "dvd.dvd (op * :: 'a :: field poly \<Rightarrow> _) = op dvd"
- by (intro ext) (simp_all add: dvd.dvd_def dvd_def)
+context
+begin
interpretation field_poly:
- unique_euclidean_ring where zero = "0 :: 'a :: field poly"
- and one = 1 and plus = plus and uminus = uminus and minus = minus
+ normalization_euclidean_semiring where zero = "0 :: 'a :: field poly"
+ and one = 1 and plus = plus and minus = minus
and times = times
- and normalize = normalize_field_poly and unit_factor = unit_factor_field_poly
- and euclidean_size = euclidean_size_field_poly
- and uniqueness_constraint = top
+ and normalize = "\<lambda>p. smult (inverse (lead_coeff p)) p"
+ and unit_factor = "\<lambda>p. [:lead_coeff p:]"
+ and euclidean_size = "\<lambda>p. if p = 0 then 0 else 2 ^ degree p"
and divide = divide and modulo = modulo
-proof (standard, unfold dvd_field_poly)
- 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)
-next
- fix p :: "'a poly" assume "is_unit p"
- then show "unit_factor_field_poly p = p"
- by (elim is_unit_polyE) (auto simp: unit_factor_field_poly_def monom_0 one_poly_def field_simps)
-next
- fix p :: "'a poly" assume "p \<noteq> 0"
- thus "is_unit (unit_factor_field_poly p)"
- by (simp add: unit_factor_field_poly_def is_unit_pCons_iff)
-next
- fix p q s :: "'a poly" assume "s \<noteq> 0"
- moreover assume "euclidean_size_field_poly p < euclidean_size_field_poly q"
- ultimately show "euclidean_size_field_poly (p * s) < euclidean_size_field_poly (q * s)"
- by (auto simp add: euclidean_size_field_poly_def degree_mult_eq)
-next
- fix p q r :: "'a poly" assume "p \<noteq> 0"
- moreover assume "euclidean_size_field_poly r < euclidean_size_field_poly p"
- ultimately show "(q * p + r) div p = q"
- by (cases "r = 0")
- (auto simp add: unit_factor_field_poly_def euclidean_size_field_poly_def div_poly_less)
-qed (auto simp: unit_factor_field_poly_def normalize_field_poly_def lead_coeff_mult
- euclidean_size_field_poly_def Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le)
+ rewrites "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd"
+ and "comm_monoid_mult.prod_mset times 1 = prod_mset"
+ and "comm_semiring_1.irreducible times 1 0 = irreducible"
+ and "comm_semiring_1.prime_elem times 1 0 = prime_elem"
+proof -
+ show "dvd.dvd (times :: 'a poly \<Rightarrow> _) = Rings.dvd"
+ by (simp add: dvd_dict)
+ show "comm_monoid_mult.prod_mset times 1 = prod_mset"
+ by (simp add: prod_mset_dict)
+ show "comm_semiring_1.irreducible times 1 0 = irreducible"
+ by (simp add: irreducible_dict)
+ show "comm_semiring_1.prime_elem times 1 0 = prime_elem"
+ by (simp add: prime_elem_dict)
+ show "class.normalization_euclidean_semiring divide plus minus (0 :: 'a poly) times 1
+ modulo (\<lambda>p. if p = 0 then 0 else 2 ^ degree p)
+ (\<lambda>p. [:lead_coeff p:]) (\<lambda>p. smult (inverse (lead_coeff p)) p)"
+ proof (standard, fold dvd_dict)
+ fix p :: "'a poly"
+ show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p"
+ by (cases "p = 0") simp_all
+ next
+ fix p :: "'a poly" assume "is_unit p"
+ then show "[:lead_coeff p:] = p"
+ by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps)
+ next
+ fix p :: "'a poly" assume "p \<noteq> 0"
+ then show "is_unit [:lead_coeff p:]"
+ by (simp add: is_unit_pCons_iff)
+ qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le)
+qed
lemma field_poly_irreducible_imp_prime:
- assumes "irreducible (p :: 'a :: field poly)"
- shows "prime_elem p"
-proof -
- have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
- from field_poly.irreducible_imp_prime_elem[of p] assms
- show ?thesis unfolding irreducible_def prime_elem_def dvd_field_poly
- comm_semiring_1.irreducible_def[OF A] comm_semiring_1.prime_elem_def[OF A] by blast
-qed
+ "prime_elem p" if "irreducible p" for p :: "'a :: field poly"
+ using that by (fact field_poly.irreducible_imp_prime_elem)
lemma field_poly_prod_mset_prime_factorization:
- assumes "(x :: 'a :: field poly) \<noteq> 0"
- shows "prod_mset (field_poly.prime_factorization x) = normalize_field_poly x"
-proof -
- have A: "class.comm_monoid_mult op * (1 :: 'a poly)" ..
- have "comm_monoid_mult.prod_mset op * (1 :: 'a poly) = prod_mset"
- by (intro ext) (simp add: comm_monoid_mult.prod_mset_def[OF A] prod_mset_def)
- with field_poly.prod_mset_prime_factorization[OF assms] show ?thesis by simp
-qed
+ "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p"
+ if "p \<noteq> 0" for p :: "'a :: field poly"
+ using that by (fact field_poly.prod_mset_prime_factorization)
lemma field_poly_in_prime_factorization_imp_prime:
- assumes "(p :: 'a :: field poly) \<in># field_poly.prime_factorization x"
- shows "prime_elem p"
-proof -
- have A: "class.comm_semiring_1 op * 1 op + (0 :: 'a poly)" ..
- have B: "class.normalization_semidom op div op + op - (0 :: 'a poly) op * 1
- unit_factor_field_poly normalize_field_poly" ..
- from field_poly.in_prime_factors_imp_prime [of p x] assms
- show ?thesis unfolding prime_elem_def dvd_field_poly
- comm_semiring_1.prime_elem_def[OF A] normalization_semidom.prime_def[OF B] by blast
-qed
+ "prime_elem p" if "p \<in># field_poly.prime_factorization x"
+ for p :: "'a :: field poly"
+ by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime)
+ (fact that)
subsection \<open>Primality and irreducibility in polynomial rings\<close>
@@ -658,9 +547,6 @@
qed
qed
-context
-begin
-
private lemma irreducible_imp_prime_poly:
fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
assumes "irreducible p"
@@ -686,7 +572,6 @@
qed (insert assms, auto simp: irreducible_def)
qed
-
lemma degree_primitive_part_fract [simp]:
"degree (primitive_part_fract p) = degree p"
proof -
@@ -749,14 +634,9 @@
shows "b \<noteq> 0 \<Longrightarrow> coprime a b \<Longrightarrow> prime_elem [:a,b:]"
by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly)
-end
-
subsection \<open>Prime factorisation of polynomials\<close>
-context
-begin
-
private lemma poly_prime_factorization_exists_content_1:
fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide} poly"
assumes "p \<noteq> 0" "content p = 1"
@@ -775,11 +655,12 @@
finally have content_e: "content e = 1"
by simp
- 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 monom_0 degree_map_poly coeff_map_poly)
- also from assms have "normalize_field_poly (fract_poly p) = prod_mset ?P"
+ from \<open>p \<noteq> 0\<close> have "fract_poly p = [:lead_coeff (fract_poly p):] *
+ smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)"
+ by simp
+ also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]"
+ by (simp add: monom_0 degree_map_poly coeff_map_poly)
+ also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P"
by (subst field_poly_prod_mset_prime_factorization) simp_all
also have "\<dots> = prod_mset (image_mset id ?P)" by simp
also have "image_mset id ?P =
@@ -858,7 +739,7 @@
end
-instantiation poly :: ("{field,factorial_ring_gcd}") unique_euclidean_ring
+instantiation poly :: ("{field,factorial_ring_gcd}") "{unique_euclidean_ring, normalization_euclidean_semiring}"
begin
definition euclidean_size_poly :: "'a poly \<Rightarrow> nat"
@@ -867,16 +748,24 @@
definition uniqueness_constraint_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
where [simp]: "uniqueness_constraint_poly = top"
-instance
- by standard
- (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq intro!: degree_mod_less' degree_mult_right_le
+instance proof
+ show "(q * p + r) div p = q" if "p \<noteq> 0"
+ and "euclidean_size r < euclidean_size p" for q p r :: "'a poly"
+ proof -
+ from \<open>p \<noteq> 0\<close> eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)"
+ by (simp add: eucl_rel_poly_iff distrib_right)
+ then have "(r + q * p) div p = q + r div p"
+ by (rule div_poly_eq)
+ with that show ?thesis
+ by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps)
+ qed
+qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq intro!: degree_mod_less' degree_mult_right_le
split: if_splits)
end
-instance poly :: ("{field,factorial_ring_gcd}") euclidean_ring_gcd
- by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI)
- standard
+instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd}") euclidean_ring_gcd
+ by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard
subsection \<open>Polynomial GCD\<close>
--- a/src/HOL/Decision_Procs/Cooper.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Mon Oct 09 22:08:05 2017 +0200
@@ -6,7 +6,6 @@
imports
Complex_Main
"HOL-Library.Code_Target_Numeral"
- "HOL-Library.Old_Recdef"
begin
(* Periodicity of dvd *)
@@ -15,50 +14,74 @@
(**** SHADOW SYNTAX AND SEMANTICS ****)
(*********************************************************************************)
-datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
+datatype (plugins del: size) num = C int | Bound nat | CN nat int num
+ | Neg num | Add num num | Sub num num
| Mul int num
-primrec num_size :: "num \<Rightarrow> nat" \<comment> \<open>A size for num to make inductive proofs simpler\<close>
+instantiation num :: size
+begin
+
+primrec size_num :: "num \<Rightarrow> nat"
where
- "num_size (C c) = 1"
-| "num_size (Bound n) = 1"
-| "num_size (Neg a) = 1 + num_size a"
-| "num_size (Add a b) = 1 + num_size a + num_size b"
-| "num_size (Sub a b) = 3 + num_size a + num_size b"
-| "num_size (CN n c a) = 4 + num_size a"
-| "num_size (Mul c a) = 1 + num_size a"
+ "size_num (C c) = 1"
+| "size_num (Bound n) = 1"
+| "size_num (Neg a) = 1 + size_num a"
+| "size_num (Add a b) = 1 + size_num a + size_num b"
+| "size_num (Sub a b) = 3 + size_num a + size_num b"
+| "size_num (CN n c a) = 4 + size_num a"
+| "size_num (Mul c a) = 1 + size_num a"
+
+instance ..
+
+end
primrec Inum :: "int list \<Rightarrow> num \<Rightarrow> int"
where
"Inum bs (C c) = c"
-| "Inum bs (Bound n) = bs!n"
-| "Inum bs (CN n c a) = c * (bs!n) + (Inum bs a)"
-| "Inum bs (Neg a) = -(Inum bs a)"
+| "Inum bs (Bound n) = bs ! n"
+| "Inum bs (CN n c a) = c * (bs ! n) + Inum bs a"
+| "Inum bs (Neg a) = - Inum bs a"
| "Inum bs (Add a b) = Inum bs a + Inum bs b"
| "Inum bs (Sub a b) = Inum bs a - Inum bs b"
-| "Inum bs (Mul c a) = c* Inum bs a"
+| "Inum bs (Mul c a) = c * Inum bs a"
-datatype fm =
- T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
- NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
- | Closed nat | NClosed nat
+datatype (plugins del: size) fm = T | F
+ | Lt num | Le num | Gt num | Ge num | Eq num | NEq num
+ | Dvd int num | NDvd int num
+ | NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm
+ | E fm | A fm | Closed nat | NClosed nat
+instantiation fm :: size
+begin
-fun fmsize :: "fm \<Rightarrow> nat" \<comment> \<open>A size for fm\<close>
+primrec size_fm :: "fm \<Rightarrow> nat"
where
- "fmsize (NOT p) = 1 + fmsize p"
-| "fmsize (And p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
-| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
-| "fmsize (E p) = 1 + fmsize p"
-| "fmsize (A p) = 4+ fmsize p"
-| "fmsize (Dvd i t) = 2"
-| "fmsize (NDvd i t) = 2"
-| "fmsize p = 1"
+ "size_fm (NOT p) = 1 + size_fm p"
+| "size_fm (And p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+| "size_fm (E p) = 1 + size_fm p"
+| "size_fm (A p) = 4 + size_fm p"
+| "size_fm (Dvd i t) = 2"
+| "size_fm (NDvd i t) = 2"
+| "size_fm T = 1"
+| "size_fm F = 1"
+| "size_fm (Lt _) = 1"
+| "size_fm (Le _) = 1"
+| "size_fm (Gt _) = 1"
+| "size_fm (Ge _) = 1"
+| "size_fm (Eq _) = 1"
+| "size_fm (NEq _) = 1"
+| "size_fm (Closed _) = 1"
+| "size_fm (NClosed _) = 1"
-lemma fmsize_pos: "fmsize p > 0"
- by (induct p rule: fmsize.induct) simp_all
+instance ..
+
+end
+
+lemma fmsize_pos [simp]: "size p > 0" for p :: fm
+ by (induct p) simp_all
primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool" \<comment> \<open>Semantics of formulae (fm)\<close>
where
@@ -79,10 +102,10 @@
| "Ifm bbs bs (Iff p q) \<longleftrightarrow> Ifm bbs bs p = Ifm bbs bs q"
| "Ifm bbs bs (E p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) p)"
| "Ifm bbs bs (A p) \<longleftrightarrow> (\<forall>x. Ifm bbs (x # bs) p)"
-| "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs!n"
-| "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs!n"
+| "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs ! n"
+| "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs ! n"
-function (sequential) prep :: "fm \<Rightarrow> fm"
+fun prep :: "fm \<Rightarrow> fm"
where
"prep (E T) = T"
| "prep (E F) = F"
@@ -107,10 +130,6 @@
| "prep (Imp p q) = prep (Or (NOT p) q)"
| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
| "prep p = p"
- by pat_completeness auto
-
-termination
- by (relation "measure fmsize") (simp_all add: fmsize_pos)
lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p"
by (induct p arbitrary: bs rule: prep.induct) auto
@@ -424,34 +443,24 @@
definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
where "lex_bnd t s = lex_ns (bnds t) (bnds s)"
-consts numadd:: "num \<times> num \<Rightarrow> num"
-recdef numadd "measure (\<lambda>(t, s). num_size t + num_size s)"
- "numadd (CN n1 c1 r1, CN n2 c2 r2) =
+fun numadd:: "num \<Rightarrow> num \<Rightarrow> num"
+where
+ "numadd (CN n1 c1 r1) (CN n2 c2 r2) =
(if n1 = n2 then
let c = c1 + c2
- in if c = 0 then numadd (r1, r2) else CN n1 c (numadd (r1, r2))
- else if n1 \<le> n2 then CN n1 c1 (numadd (r1, Add (Mul c2 (Bound n2)) r2))
- else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1, r2)))"
- "numadd (CN n1 c1 r1, t) = CN n1 c1 (numadd (r1, t))"
- "numadd (t, CN n2 c2 r2) = CN n2 c2 (numadd (t, r2))"
- "numadd (C b1, C b2) = C (b1 + b2)"
- "numadd (a, b) = Add a b"
+ in if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)
+ else if n1 \<le> n2 then CN n1 c1 (numadd r1 (Add (Mul c2 (Bound n2)) r2))
+ else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1) r2))"
+| "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"
+| "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"
+| "numadd (C b1) (C b2) = C (b1 + b2)"
+| "numadd a b = Add a b"
-lemma numadd: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
- apply (induct t s rule: numadd.induct)
- apply (simp_all add: Let_def)
- subgoal for n1 c1 r1 n2 c2 r2
- apply (cases "c1 + c2 = 0")
- apply (cases "n1 \<le> n2")
- apply simp_all
- apply (cases "n1 = n2")
- apply (simp_all add: algebra_simps)
- apply (simp add: distrib_right[symmetric])
- done
- done
+lemma numadd: "Inum bs (numadd t s) = Inum bs (Add t s)"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
-lemma numadd_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd (t, s))"
- by (induct t s rule: numadd.induct) (auto simp add: Let_def)
+lemma numadd_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd t s)"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def)
fun nummul :: "int \<Rightarrow> num \<Rightarrow> num"
where
@@ -460,16 +469,16 @@
| "nummul i t = Mul i t"
lemma nummul: "Inum bs (nummul i t) = Inum bs (Mul i t)"
- by (induct t arbitrary: i rule: nummul.induct) (auto simp add: algebra_simps numadd)
+ by (induct t arbitrary: i rule: nummul.induct) (simp_all add: algebra_simps)
lemma nummul_nb: "numbound0 t \<Longrightarrow> numbound0 (nummul i t)"
- by (induct t arbitrary: i rule: nummul.induct) (auto simp add: numadd_nb)
+ by (induct t arbitrary: i rule: nummul.induct) (simp_all add: numadd_nb)
definition numneg :: "num \<Rightarrow> num"
where "numneg t = nummul (- 1) t"
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"
- where "numsub s t = (if s = t then C 0 else numadd (s, numneg t))"
+ where "numsub s t = (if s = t then C 0 else numadd s (numneg t))"
lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)"
using numneg_def nummul by simp
@@ -488,7 +497,7 @@
"simpnum (C j) = C j"
| "simpnum (Bound n) = CN n 1 (C 0)"
| "simpnum (Neg t) = numneg (simpnum t)"
-| "simpnum (Add t s) = numadd (simpnum t, simpnum s)"
+| "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
| "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))"
| "simpnum t = t"
@@ -587,7 +596,7 @@
lemma iff_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)"
using iff_def by (unfold iff_def, cases "p = q", auto simp add: not_bn)
-function (sequential) simpfm :: "fm \<Rightarrow> fm"
+fun simpfm :: "fm \<Rightarrow> fm"
where
"simpfm (And p q) = conj (simpfm p) (simpfm q)"
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
@@ -609,8 +618,6 @@
else if \<bar>i\<bar> = 1 then F
else let a' = simpnum a in case a' of C v \<Rightarrow> if \<not>( i dvd v) then T else F | _ \<Rightarrow> NDvd i a')"
| "simpfm p = p"
- by pat_completeness auto
-termination by (relation "measure fmsize") auto
lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs bs p"
proof (induct p rule: simpfm.induct)
@@ -819,7 +826,7 @@
done
text \<open>Generic quantifier elimination\<close>
-function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
@@ -829,8 +836,6 @@
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
| "qelim p = (\<lambda>y. simpfm p)"
- by pat_completeness auto
-termination by (relation "measure fmsize") auto
lemma qelim_ci:
assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)"
@@ -990,7 +995,7 @@
lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p"
by (induct p rule: iszlfm.induct) auto
-function (sequential) zlfm :: "fm \<Rightarrow> fm" \<comment> \<open>Linearity transformation for fm\<close>
+fun zlfm :: "fm \<Rightarrow> fm" \<comment> \<open>Linearity transformation for fm\<close>
where
"zlfm (And p q) = And (zlfm p) (zlfm q)"
| "zlfm (Or p q) = Or (zlfm p) (zlfm q)"
@@ -1058,10 +1063,6 @@
| "zlfm (NOT (Closed P)) = NClosed P"
| "zlfm (NOT (NClosed P)) = Closed P"
| "zlfm p = p"
- by pat_completeness auto
-
-termination
- by (relation "measure fmsize") (simp_all add: fmsize_pos)
lemma zlfm_I:
assumes qfp: "qfree p"
--- a/src/HOL/Decision_Procs/Ferrack.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Oct 09 22:08:05 2017 +0200
@@ -4,7 +4,7 @@
theory Ferrack
imports Complex_Main Dense_Linear_Order DP_Library
- "HOL-Library.Code_Target_Numeral" "HOL-Library.Old_Recdef"
+ "HOL-Library.Code_Target_Numeral"
begin
section \<open>Quantifier elimination for \<open>\<real> (0, 1, +, <)\<close>\<close>
@@ -13,19 +13,25 @@
(**** SHADOW SYNTAX AND SEMANTICS ****)
(*********************************************************************************)
-datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
+datatype (plugins del: size) num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
| Mul int num
- (* A size for num to make inductive proofs simpler*)
-primrec num_size :: "num \<Rightarrow> nat"
+instantiation num :: size
+begin
+
+primrec size_num :: "num \<Rightarrow> nat"
where
- "num_size (C c) = 1"
-| "num_size (Bound n) = 1"
-| "num_size (Neg a) = 1 + num_size a"
-| "num_size (Add a b) = 1 + num_size a + num_size b"
-| "num_size (Sub a b) = 3 + num_size a + num_size b"
-| "num_size (Mul c a) = 1 + num_size a"
-| "num_size (CN n c a) = 3 + num_size a "
+ "size_num (C c) = 1"
+| "size_num (Bound n) = 1"
+| "size_num (Neg a) = 1 + size_num a"
+| "size_num (Add a b) = 1 + size_num a + size_num b"
+| "size_num (Sub a b) = 3 + size_num a + size_num b"
+| "size_num (Mul c a) = 1 + size_num a"
+| "size_num (CN n c a) = 3 + size_num a "
+
+instance ..
+
+end
(* Semantics of numeral terms (num) *)
primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real"
@@ -38,26 +44,37 @@
| "Inum bs (Sub a b) = Inum bs a - Inum bs b"
| "Inum bs (Mul c a) = (real_of_int c) * Inum bs a"
(* FORMULAE *)
-datatype fm =
+datatype (plugins del: size) fm =
T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num|
NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
+instantiation fm :: size
+begin
- (* A size for fm *)
-fun fmsize :: "fm \<Rightarrow> nat"
+primrec size_fm :: "fm \<Rightarrow> nat"
where
- "fmsize (NOT p) = 1 + fmsize p"
-| "fmsize (And p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
-| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
-| "fmsize (E p) = 1 + fmsize p"
-| "fmsize (A p) = 4+ fmsize p"
-| "fmsize p = 1"
- (* several lemmas about fmsize *)
+ "size_fm (NOT p) = 1 + size_fm p"
+| "size_fm (And p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+| "size_fm (E p) = 1 + size_fm p"
+| "size_fm (A p) = 4 + size_fm p"
+| "size_fm T = 1"
+| "size_fm F = 1"
+| "size_fm (Lt _) = 1"
+| "size_fm (Le _) = 1"
+| "size_fm (Gt _) = 1"
+| "size_fm (Ge _) = 1"
+| "size_fm (Eq _) = 1"
+| "size_fm (NEq _) = 1"
-lemma fmsize_pos: "fmsize p > 0"
- by (induct p rule: fmsize.induct) simp_all
+instance ..
+
+end
+
+lemma size_fm_pos [simp]: "size p > 0" for p :: fm
+ by (induct p) simp_all
(* Semantics of formulae (fm) *)
primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool"
@@ -646,33 +663,24 @@
lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
-consts numadd:: "num \<times> num \<Rightarrow> num"
-recdef numadd "measure (\<lambda>(t,s). size t + size s)"
- "numadd (CN n1 c1 r1,CN n2 c2 r2) =
+fun numadd:: "num \<Rightarrow> num \<Rightarrow> num"
+where
+ "numadd (CN n1 c1 r1) (CN n2 c2 r2) =
(if n1 = n2 then
(let c = c1 + c2
- in (if c = 0 then numadd(r1,r2) else CN n1 c (numadd (r1, r2))))
- else if n1 \<le> n2 then (CN n1 c1 (numadd (r1,CN n2 c2 r2)))
- else (CN n2 c2 (numadd (CN n1 c1 r1, r2))))"
- "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"
- "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t, r2))"
- "numadd (C b1, C b2) = C (b1 + b2)"
- "numadd (a,b) = Add a b"
+ in (if c = 0 then numadd r1 r2 else CN n1 c (numadd r1 r2)))
+ else if n1 \<le> n2 then (CN n1 c1 (numadd r1 (CN n2 c2 r2)))
+ else (CN n2 c2 (numadd (CN n1 c1 r1) r2)))"
+| "numadd (CN n1 c1 r1) t = CN n1 c1 (numadd r1 t)"
+| "numadd t (CN n2 c2 r2) = CN n2 c2 (numadd t r2)"
+| "numadd (C b1) (C b2) = C (b1 + b2)"
+| "numadd a b = Add a b"
-lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
- apply (induct t s rule: numadd.induct)
- apply (simp_all add: Let_def)
- apply (case_tac "c1 + c2 = 0")
- apply (case_tac "n1 \<le> n2")
- apply simp_all
- apply (case_tac "n1 = n2")
- apply (simp_all add: algebra_simps)
- apply (simp only: distrib_right[symmetric])
- apply simp
- done
+lemma numadd [simp]: "Inum bs (numadd t s) = Inum bs (Add t s)"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
-lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
- by (induct t s rule: numadd.induct) (auto simp add: Let_def)
+lemma numadd_nb [simp]: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd t s)"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def)
fun nummul:: "num \<Rightarrow> int \<Rightarrow> num"
where
@@ -690,7 +698,7 @@
where "numneg t = nummul t (- 1)"
definition numsub :: "num \<Rightarrow> num \<Rightarrow> num"
- where "numsub s t = (if s = t then C 0 else numadd (s, numneg t))"
+ where "numsub s t = (if s = t then C 0 else numadd s (numneg t))"
lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
using numneg_def by simp
@@ -709,10 +717,10 @@
"simpnum (C j) = C j"
| "simpnum (Bound n) = CN n 1 (C 0)"
| "simpnum (Neg t) = numneg (simpnum t)"
-| "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
+| "simpnum (Add t s) = numadd (simpnum t) (simpnum s)"
| "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
| "simpnum (Mul i t) = (if i = 0 then C 0 else nummul (simpnum t) i)"
-| "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0), simpnum t))"
+| "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0)) (simpnum t))"
lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
by (induct t) simp_all
@@ -726,8 +734,8 @@
| "nozerocoeff (CN n c t) = (c \<noteq> 0 \<and> nozerocoeff t)"
| "nozerocoeff t = True"
-lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))"
- by (induct a b rule: numadd.induct) (auto simp add: Let_def)
+lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd a b)"
+ by (induct a b rule: numadd.induct) (simp_all add: Let_def)
lemma nummul_nz : "\<And>i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
by (induct a rule: nummul.induct) (auto simp add: Let_def numadd_nz)
@@ -856,7 +864,7 @@
proof (cases "?g > 1")
case False
then show ?thesis
- using assms by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)
+ using assms by (auto simp add: Let_def simp_num_pair_def)
next
case g1: True
then have g0: "?g > 0" by simp
@@ -868,7 +876,7 @@
proof cases
case 1
then show ?thesis
- using assms g1 by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)
+ using assms g1 by (auto simp add: Let_def simp_num_pair_def)
next
case g'1: 2
have gpdg: "?g' dvd ?g" by simp
@@ -879,7 +887,7 @@
by simp
then show ?thesis
using assms g1 g'1
- by (auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)
+ by (auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)
qed
qed
qed
@@ -985,7 +993,7 @@
case 2
then show ?thesis using sa by (cases ?sa) (simp_all add: Let_def)
qed
-qed (induct p rule: simpfm.induct, simp_all add: conj disj imp iff not)
+qed (induct p rule: simpfm.induct, simp_all)
lemma simpfm_bound0: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
@@ -1019,7 +1027,7 @@
then have nb: "numbound0 a" by simp
then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
then show ?case by (cases "simpnum a") (auto simp add: Let_def)
-qed (auto simp add: disj_def imp_def iff_def conj_def not_bn)
+qed (auto simp add: disj_def imp_def iff_def conj_def)
lemma simpfm_qf: "qfree p \<Longrightarrow> qfree (simpfm p)"
apply (induct p rule: simpfm.induct)
@@ -1027,38 +1035,37 @@
apply (case_tac "simpnum a", auto)+
done
-consts prep :: "fm \<Rightarrow> fm"
-recdef prep "measure fmsize"
+fun prep :: "fm \<Rightarrow> fm"
+where
"prep (E T) = T"
- "prep (E F) = F"
- "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
- "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
- "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
- "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
- "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
- "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
- "prep (E p) = E (prep p)"
- "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
- "prep (A p) = prep (NOT (E (NOT p)))"
- "prep (NOT (NOT p)) = prep p"
- "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
- "prep (NOT (A p)) = prep (E (NOT p))"
- "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
- "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
- "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
- "prep (NOT p) = not (prep p)"
- "prep (Or p q) = disj (prep p) (prep q)"
- "prep (And p q) = conj (prep p) (prep q)"
- "prep (Imp p q) = prep (Or (NOT p) q)"
- "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
- "prep p = p"
- (hints simp add: fmsize_pos)
+| "prep (E F) = F"
+| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
+| "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
+| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
+| "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
+| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+| "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+| "prep (E p) = E (prep p)"
+| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
+| "prep (A p) = prep (NOT (E (NOT p)))"
+| "prep (NOT (NOT p)) = prep p"
+| "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
+| "prep (NOT (A p)) = prep (E (NOT p))"
+| "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
+| "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
+| "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
+| "prep (NOT p) = not (prep p)"
+| "prep (Or p q) = disj (prep p) (prep q)"
+| "prep (And p q) = conj (prep p) (prep q)"
+| "prep (Imp p q) = prep (Or (NOT p) q)"
+| "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+| "prep p = p"
lemma prep: "\<And>bs. Ifm bs (prep p) = Ifm bs p"
by (induct p rule: prep.induct) auto
(* Generic quantifier elimination *)
-function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
@@ -1068,16 +1075,13 @@
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
| "qelim p = (\<lambda>y. simpfm p)"
- by pat_completeness auto
-termination qelim by (relation "measure fmsize") simp_all
lemma qelim_ci:
assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
shows "\<And>bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
using qe_inv DJ_qe[OF qe_inv]
by (induct p rule: qelim.induct)
- (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf
- simpfm simpfm_qf simp del: simpfm.simps)
+ (auto simp add: simpfm simpfm_qf simp del: simpfm.simps)
fun minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
where
@@ -1116,7 +1120,7 @@
| "isrlfm p = (isatom p \<and> (bound0 p))"
(* splits the bounded from the unbounded part*)
-function (sequential) rsplit0 :: "num \<Rightarrow> int \<times> num"
+fun rsplit0 :: "num \<Rightarrow> int \<times> num"
where
"rsplit0 (Bound 0) = (1,C 0)"
| "rsplit0 (Add a b) = (let (ca,ta) = rsplit0 a; (cb,tb) = rsplit0 b in (ca + cb, Add ta tb))"
@@ -1126,8 +1130,6 @@
| "rsplit0 (CN 0 c a) = (let (ca,ta) = rsplit0 a in (c + ca, ta))"
| "rsplit0 (CN n c a) = (let (ca,ta) = rsplit0 a in (ca, CN n c ta))"
| "rsplit0 t = (0,t)"
- by pat_completeness auto
-termination rsplit0 by (relation "measure num_size") simp_all
lemma rsplit0: "Inum bs ((case_prod (CN 0)) (rsplit0 t)) = Inum bs t \<and> numbound0 (snd (rsplit0 t))"
proof (induct t rule: rsplit0.induct)
@@ -1222,39 +1224,38 @@
lemma disj_lin: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
by (auto simp add: disj_def)
-consts rlfm :: "fm \<Rightarrow> fm"
-recdef rlfm "measure fmsize"
+fun rlfm :: "fm \<Rightarrow> fm"
+where
"rlfm (And p q) = conj (rlfm p) (rlfm q)"
- "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
- "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
- "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))"
- "rlfm (Lt a) = case_prod lt (rsplit0 a)"
- "rlfm (Le a) = case_prod le (rsplit0 a)"
- "rlfm (Gt a) = case_prod gt (rsplit0 a)"
- "rlfm (Ge a) = case_prod ge (rsplit0 a)"
- "rlfm (Eq a) = case_prod eq (rsplit0 a)"
- "rlfm (NEq a) = case_prod neq (rsplit0 a)"
- "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
- "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
- "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
- "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
- "rlfm (NOT (NOT p)) = rlfm p"
- "rlfm (NOT T) = F"
- "rlfm (NOT F) = T"
- "rlfm (NOT (Lt a)) = rlfm (Ge a)"
- "rlfm (NOT (Le a)) = rlfm (Gt a)"
- "rlfm (NOT (Gt a)) = rlfm (Le a)"
- "rlfm (NOT (Ge a)) = rlfm (Lt a)"
- "rlfm (NOT (Eq a)) = rlfm (NEq a)"
- "rlfm (NOT (NEq a)) = rlfm (Eq a)"
- "rlfm p = p"
- (hints simp add: fmsize_pos)
+| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
+| "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
+| "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))"
+| "rlfm (Lt a) = case_prod lt (rsplit0 a)"
+| "rlfm (Le a) = case_prod le (rsplit0 a)"
+| "rlfm (Gt a) = case_prod gt (rsplit0 a)"
+| "rlfm (Ge a) = case_prod ge (rsplit0 a)"
+| "rlfm (Eq a) = case_prod eq (rsplit0 a)"
+| "rlfm (NEq a) = case_prod neq (rsplit0 a)"
+| "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
+| "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
+| "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
+| "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
+| "rlfm (NOT (NOT p)) = rlfm p"
+| "rlfm (NOT T) = F"
+| "rlfm (NOT F) = T"
+| "rlfm (NOT (Lt a)) = rlfm (Ge a)"
+| "rlfm (NOT (Le a)) = rlfm (Gt a)"
+| "rlfm (NOT (Gt a)) = rlfm (Le a)"
+| "rlfm (NOT (Ge a)) = rlfm (Lt a)"
+| "rlfm (NOT (Eq a)) = rlfm (NEq a)"
+| "rlfm (NOT (NEq a)) = rlfm (Eq a)"
+| "rlfm p = p"
lemma rlfm_I:
assumes qfp: "qfree p"
shows "(Ifm bs (rlfm p) = Ifm bs p) \<and> isrlfm (rlfm p)"
using qfp
- by (induct p rule: rlfm.induct) (auto simp add: lt le gt ge eq neq conj disj conj_lin disj_lin)
+ by (induct p rule: rlfm.induct) (auto simp add: lt le gt ge eq neq conj_lin disj_lin)
(* Operations needed for Ferrante and Rackoff *)
lemma rminusinf_inf:
@@ -1555,29 +1556,29 @@
ultimately show ?thesis using z_def by auto
qed
-consts
- uset:: "fm \<Rightarrow> (num \<times> int) list"
- usubst :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
-recdef uset "measure size"
+fun uset :: "fm \<Rightarrow> (num \<times> int) list"
+where
"uset (And p q) = (uset p @ uset q)"
- "uset (Or p q) = (uset p @ uset q)"
- "uset (Eq (CN 0 c e)) = [(Neg e,c)]"
- "uset (NEq (CN 0 c e)) = [(Neg e,c)]"
- "uset (Lt (CN 0 c e)) = [(Neg e,c)]"
- "uset (Le (CN 0 c e)) = [(Neg e,c)]"
- "uset (Gt (CN 0 c e)) = [(Neg e,c)]"
- "uset (Ge (CN 0 c e)) = [(Neg e,c)]"
- "uset p = []"
-recdef usubst "measure size"
+| "uset (Or p q) = (uset p @ uset q)"
+| "uset (Eq (CN 0 c e)) = [(Neg e,c)]"
+| "uset (NEq (CN 0 c e)) = [(Neg e,c)]"
+| "uset (Lt (CN 0 c e)) = [(Neg e,c)]"
+| "uset (Le (CN 0 c e)) = [(Neg e,c)]"
+| "uset (Gt (CN 0 c e)) = [(Neg e,c)]"
+| "uset (Ge (CN 0 c e)) = [(Neg e,c)]"
+| "uset p = []"
+
+fun usubst :: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm"
+where
"usubst (And p q) = (\<lambda>(t,n). And (usubst p (t,n)) (usubst q (t,n)))"
- "usubst (Or p q) = (\<lambda>(t,n). Or (usubst p (t,n)) (usubst q (t,n)))"
- "usubst (Eq (CN 0 c e)) = (\<lambda>(t,n). Eq (Add (Mul c t) (Mul n e)))"
- "usubst (NEq (CN 0 c e)) = (\<lambda>(t,n). NEq (Add (Mul c t) (Mul n e)))"
- "usubst (Lt (CN 0 c e)) = (\<lambda>(t,n). Lt (Add (Mul c t) (Mul n e)))"
- "usubst (Le (CN 0 c e)) = (\<lambda>(t,n). Le (Add (Mul c t) (Mul n e)))"
- "usubst (Gt (CN 0 c e)) = (\<lambda>(t,n). Gt (Add (Mul c t) (Mul n e)))"
- "usubst (Ge (CN 0 c e)) = (\<lambda>(t,n). Ge (Add (Mul c t) (Mul n e)))"
- "usubst p = (\<lambda>(t, n). p)"
+| "usubst (Or p q) = (\<lambda>(t,n). Or (usubst p (t,n)) (usubst q (t,n)))"
+| "usubst (Eq (CN 0 c e)) = (\<lambda>(t,n). Eq (Add (Mul c t) (Mul n e)))"
+| "usubst (NEq (CN 0 c e)) = (\<lambda>(t,n). NEq (Add (Mul c t) (Mul n e)))"
+| "usubst (Lt (CN 0 c e)) = (\<lambda>(t,n). Lt (Add (Mul c t) (Mul n e)))"
+| "usubst (Le (CN 0 c e)) = (\<lambda>(t,n). Le (Add (Mul c t) (Mul n e)))"
+| "usubst (Gt (CN 0 c e)) = (\<lambda>(t,n). Gt (Add (Mul c t) (Mul n e)))"
+| "usubst (Ge (CN 0 c e)) = (\<lambda>(t,n). Ge (Add (Mul c t) (Mul n e)))"
+| "usubst p = (\<lambda>(t, n). p)"
lemma usubst_I:
assumes lp: "isrlfm p"
--- a/src/HOL/Decision_Procs/MIR.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Mon Oct 09 22:08:05 2017 +0200
@@ -84,23 +84,32 @@
(**** SHADOW SYNTAX AND SEMANTICS ****)
(*********************************************************************************)
-datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
- | Mul int num | Floor num| CF int num num
-
- (* A size for num to make inductive proofs simpler*)
-primrec num_size :: "num \<Rightarrow> nat" where
- "num_size (C c) = 1"
-| "num_size (Bound n) = 1"
-| "num_size (Neg a) = 1 + num_size a"
-| "num_size (Add a b) = 1 + num_size a + num_size b"
-| "num_size (Sub a b) = 3 + num_size a + num_size b"
-| "num_size (CN n c a) = 4 + num_size a "
-| "num_size (CF c a b) = 4 + num_size a + num_size b"
-| "num_size (Mul c a) = 1 + num_size a"
-| "num_size (Floor a) = 1 + num_size a"
+datatype (plugins del: size) num = C int | Bound nat | CN nat int num
+ | Neg num | Add num num | Sub num num
+ | Mul int num | Floor num | CF int num num
+
+instantiation num :: size
+begin
+
+primrec size_num :: "num \<Rightarrow> nat"
+where
+ "size_num (C c) = 1"
+| "size_num (Bound n) = 1"
+| "size_num (Neg a) = 1 + size_num a"
+| "size_num (Add a b) = 1 + size_num a + size_num b"
+| "size_num (Sub a b) = 3 + size_num a + size_num b"
+| "size_num (CN n c a) = 4 + size_num a "
+| "size_num (CF c a b) = 4 + size_num a + size_num b"
+| "size_num (Mul c a) = 1 + size_num a"
+| "size_num (Floor a) = 1 + size_num a"
+
+instance ..
+
+end
(* Semantics of numeral terms (num) *)
-primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where
+primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real"
+where
"Inum bs (C c) = (real_of_int c)"
| "Inum bs (Bound n) = bs!n"
| "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (Inum bs a)"
@@ -169,48 +178,64 @@
(* FORMULAE *)
-datatype fm =
- T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
- NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
-
-
- (* A size for fm *)
-fun fmsize :: "fm \<Rightarrow> nat" where
- "fmsize (NOT p) = 1 + fmsize p"
-| "fmsize (And p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
-| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
-| "fmsize (E p) = 1 + fmsize p"
-| "fmsize (A p) = 4+ fmsize p"
-| "fmsize (Dvd i t) = 2"
-| "fmsize (NDvd i t) = 2"
-| "fmsize p = 1"
- (* several lemmas about fmsize *)
-lemma fmsize_pos: "fmsize p > 0"
- by (induct p rule: fmsize.induct) simp_all
+datatype (plugins del: size) fm =
+ T | F | Lt num | Le num | Gt num | Ge num | Eq num | NEq num |
+ Dvd int num | NDvd int num |
+ NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
+
+instantiation fm :: size
+begin
+
+primrec size_fm :: "fm \<Rightarrow> nat"
+where
+ "size_fm (NOT p) = 1 + size_fm p"
+| "size_fm (And p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+| "size_fm (E p) = 1 + size_fm p"
+| "size_fm (A p) = 4 + size_fm p"
+| "size_fm (Dvd i t) = 2"
+| "size_fm (NDvd i t) = 2"
+| "size_fm T = 1"
+| "size_fm F = 1"
+| "size_fm (Lt _) = 1"
+| "size_fm (Le _) = 1"
+| "size_fm (Gt _) = 1"
+| "size_fm (Ge _) = 1"
+| "size_fm (Eq _) = 1"
+| "size_fm (NEq _) = 1"
+
+instance ..
+
+end
+
+lemma size_fm_pos [simp]: "size p > 0" for p :: fm
+ by (induct p) simp_all
(* Semantics of formulae (fm) *)
-primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where
- "Ifm bs T = True"
-| "Ifm bs F = False"
-| "Ifm bs (Lt a) = (Inum bs a < 0)"
-| "Ifm bs (Gt a) = (Inum bs a > 0)"
-| "Ifm bs (Le a) = (Inum bs a \<le> 0)"
-| "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
-| "Ifm bs (Eq a) = (Inum bs a = 0)"
-| "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
-| "Ifm bs (Dvd i b) = (real_of_int i rdvd Inum bs b)"
-| "Ifm bs (NDvd i b) = (\<not>(real_of_int i rdvd Inum bs b))"
-| "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
-| "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
-| "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
-| "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
-| "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
-| "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)"
-| "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)"
-
-function (sequential) prep :: "fm \<Rightarrow> fm" where
+primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool"
+where
+ "Ifm bs T \<longleftrightarrow> True"
+| "Ifm bs F \<longleftrightarrow> False"
+| "Ifm bs (Lt a) \<longleftrightarrow> Inum bs a < 0"
+| "Ifm bs (Gt a) \<longleftrightarrow> Inum bs a > 0"
+| "Ifm bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0"
+| "Ifm bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0"
+| "Ifm bs (Eq a) \<longleftrightarrow> Inum bs a = 0"
+| "Ifm bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
+| "Ifm bs (Dvd i b) \<longleftrightarrow> real_of_int i rdvd Inum bs b"
+| "Ifm bs (NDvd i b) \<longleftrightarrow> \<not> (real_of_int i rdvd Inum bs b)"
+| "Ifm bs (NOT p) \<longleftrightarrow> \<not> (Ifm bs p)"
+| "Ifm bs (And p q) \<longleftrightarrow> Ifm bs p \<and> Ifm bs q"
+| "Ifm bs (Or p q) \<longleftrightarrow> Ifm bs p \<or> Ifm bs q"
+| "Ifm bs (Imp p q) \<longleftrightarrow> (Ifm bs p \<longrightarrow> Ifm bs q)"
+| "Ifm bs (Iff p q) \<longleftrightarrow> (Ifm bs p \<longleftrightarrow> Ifm bs q)"
+| "Ifm bs (E p) \<longleftrightarrow> (\<exists>x. Ifm (x # bs) p)"
+| "Ifm bs (A p) \<longleftrightarrow> (\<forall>x. Ifm (x # bs) p)"
+
+fun prep :: "fm \<Rightarrow> fm"
+where
"prep (E T) = T"
| "prep (E F) = F"
| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
@@ -234,35 +259,35 @@
| "prep (Imp p q) = prep (Or (NOT p) q)"
| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
| "prep p = p"
- by pat_completeness simp_all
-termination by (relation "measure fmsize") (simp_all add: fmsize_pos)
lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
by (induct p rule: prep.induct) auto
(* Quantifier freeness *)
-fun qfree:: "fm \<Rightarrow> bool" where
+fun qfree:: "fm \<Rightarrow> bool"
+where
"qfree (E p) = False"
- | "qfree (A p) = False"
- | "qfree (NOT p) = qfree p"
- | "qfree (And p q) = (qfree p \<and> qfree q)"
- | "qfree (Or p q) = (qfree p \<and> qfree q)"
- | "qfree (Imp p q) = (qfree p \<and> qfree q)"
- | "qfree (Iff p q) = (qfree p \<and> qfree q)"
- | "qfree p = True"
+| "qfree (A p) = False"
+| "qfree (NOT p) = qfree p"
+| "qfree (And p q) = (qfree p \<and> qfree q)"
+| "qfree (Or p q) = (qfree p \<and> qfree q)"
+| "qfree (Imp p q) = (qfree p \<and> qfree q)"
+| "qfree (Iff p q) = (qfree p \<and> qfree q)"
+| "qfree p = True"
(* Boundedness and substitution *)
-primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *) where
+primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *)
+where
"numbound0 (C c) = True"
- | "numbound0 (Bound n) = (n>0)"
- | "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
- | "numbound0 (Neg a) = numbound0 a"
- | "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
- | "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)"
- | "numbound0 (Mul i a) = numbound0 a"
- | "numbound0 (Floor a) = numbound0 a"
- | "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)"
+| "numbound0 (Bound n) = (n>0)"
+| "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
+| "numbound0 (Neg a) = numbound0 a"
+| "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
+| "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)"
+| "numbound0 (Mul i a) = numbound0 a"
+| "numbound0 (Floor a) = numbound0 a"
+| "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)"
lemma numbound0_I:
assumes nb: "numbound0 a"
@@ -280,24 +305,25 @@
by (simp add: isint_def)
qed
-primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) where
+primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
+where
"bound0 T = True"
- | "bound0 F = True"
- | "bound0 (Lt a) = numbound0 a"
- | "bound0 (Le a) = numbound0 a"
- | "bound0 (Gt a) = numbound0 a"
- | "bound0 (Ge a) = numbound0 a"
- | "bound0 (Eq a) = numbound0 a"
- | "bound0 (NEq a) = numbound0 a"
- | "bound0 (Dvd i a) = numbound0 a"
- | "bound0 (NDvd i a) = numbound0 a"
- | "bound0 (NOT p) = bound0 p"
- | "bound0 (And p q) = (bound0 p \<and> bound0 q)"
- | "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
- | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
- | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
- | "bound0 (E p) = False"
- | "bound0 (A p) = False"
+| "bound0 F = True"
+| "bound0 (Lt a) = numbound0 a"
+| "bound0 (Le a) = numbound0 a"
+| "bound0 (Gt a) = numbound0 a"
+| "bound0 (Ge a) = numbound0 a"
+| "bound0 (Eq a) = numbound0 a"
+| "bound0 (NEq a) = numbound0 a"
+| "bound0 (Dvd i a) = numbound0 a"
+| "bound0 (NDvd i a) = numbound0 a"
+| "bound0 (NOT p) = bound0 p"
+| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
+| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
+| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
+| "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
+| "bound0 (E p) = False"
+| "bound0 (A p) = False"
lemma bound0_I:
assumes bp: "bound0 p"
@@ -305,44 +331,47 @@
using bp numbound0_I [where b="b" and bs="bs" and b'="b'"]
by (induct p) auto
-primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *) where
+primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *)
+where
"numsubst0 t (C c) = (C c)"
- | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
- | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
- | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
- | "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
- | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
- | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
- | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
- | "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
+| "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
+| "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
+| "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
+| "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
+| "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
+| "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
+| "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
+| "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
lemma numsubst0_I:
shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
by (induct t) simp_all
-primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *) where
+primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *)
+where
"subst0 t T = T"
- | "subst0 t F = F"
- | "subst0 t (Lt a) = Lt (numsubst0 t a)"
- | "subst0 t (Le a) = Le (numsubst0 t a)"
- | "subst0 t (Gt a) = Gt (numsubst0 t a)"
- | "subst0 t (Ge a) = Ge (numsubst0 t a)"
- | "subst0 t (Eq a) = Eq (numsubst0 t a)"
- | "subst0 t (NEq a) = NEq (numsubst0 t a)"
- | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
- | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
- | "subst0 t (NOT p) = NOT (subst0 t p)"
- | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
- | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
- | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
- | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
+| "subst0 t F = F"
+| "subst0 t (Lt a) = Lt (numsubst0 t a)"
+| "subst0 t (Le a) = Le (numsubst0 t a)"
+| "subst0 t (Gt a) = Gt (numsubst0 t a)"
+| "subst0 t (Ge a) = Ge (numsubst0 t a)"
+| "subst0 t (Eq a) = Eq (numsubst0 t a)"
+| "subst0 t (NEq a) = NEq (numsubst0 t a)"
+| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
+| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
+| "subst0 t (NOT p) = NOT (subst0 t p)"
+| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
+| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
+| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
+| "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
lemma subst0_I: assumes qfp: "qfree p"
shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p"
using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
by (induct p) simp_all
-fun decrnum:: "num \<Rightarrow> num" where
+fun decrnum:: "num \<Rightarrow> num"
+where
"decrnum (Bound n) = Bound (n - 1)"
| "decrnum (Neg a) = Neg (decrnum a)"
| "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
@@ -353,7 +382,8 @@
| "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
| "decrnum a = a"
-fun decr :: "fm \<Rightarrow> fm" where
+fun decr :: "fm \<Rightarrow> fm"
+where
"decr (Lt a) = Lt (decrnum a)"
| "decr (Le a) = Le (decrnum a)"
| "decr (Gt a) = Gt (decrnum a)"
@@ -380,7 +410,8 @@
lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
by (induct p) simp_all
-fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) where
+fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
+where
"isatom T = True"
| "isatom F = True"
| "isatom (Lt a) = True"
@@ -441,12 +472,14 @@
apply auto
done
-fun disjuncts :: "fm \<Rightarrow> fm list" where
+fun disjuncts :: "fm \<Rightarrow> fm list"
+where
"disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
| "disjuncts F = []"
| "disjuncts p = [p]"
-fun conjuncts :: "fm \<Rightarrow> fm list" where
+fun conjuncts :: "fm \<Rightarrow> fm list"
+where
"conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
| "conjuncts T = []"
| "conjuncts p = [p]"
@@ -511,7 +544,8 @@
(* Simplification *)
(* Algebraic simplifications for nums *)
-fun bnds:: "num \<Rightarrow> nat list" where
+fun bnds:: "num \<Rightarrow> nat list"
+where
"bnds (Bound n) = [n]"
| "bnds (CN n c a) = n#(bnds a)"
| "bnds (Neg a) = bnds a"
@@ -521,14 +555,17 @@
| "bnds (Floor a) = bnds a"
| "bnds (CF c a b) = (bnds a)@(bnds b)"
| "bnds a = []"
-fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool" where
+
+fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
+where
"lex_ns [] ms = True"
| "lex_ns ns [] = False"
| "lex_ns (n#ns) (m#ms) = (n<m \<or> ((n = m) \<and> lex_ns ns ms)) "
definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
"lex_bnd t s \<equiv> lex_ns (bnds t) (bnds s)"
-fun maxcoeff:: "num \<Rightarrow> int" where
+fun maxcoeff:: "num \<Rightarrow> int"
+where
"maxcoeff (C i) = \<bar>i\<bar>"
| "maxcoeff (CN n c t) = max \<bar>c\<bar> (maxcoeff t)"
| "maxcoeff (CF c t s) = max \<bar>c\<bar> (maxcoeff s)"
@@ -537,7 +574,8 @@
lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
by (induct t rule: maxcoeff.induct) auto
-fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int" where
+fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int"
+where
"numgcdh (C i) = (\<lambda>g. gcd i g)"
| "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
| "numgcdh (CF c s t) = (\<lambda>g. gcd c (numgcdh t g))"
@@ -546,7 +584,8 @@
definition numgcd :: "num \<Rightarrow> int"
where "numgcd t = numgcdh t (maxcoeff t)"
-fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num" where
+fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num"
+where
"reducecoeffh (C i) = (\<lambda> g. C (i div g))"
| "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
| "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g) s (reducecoeffh t g))"
@@ -558,7 +597,8 @@
(let g = numgcd t in
if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)"
-fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
+fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
+where
"dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
| "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
| "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
@@ -602,7 +642,8 @@
from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps)
qed (auto simp add: numgcd_def gp)
-fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
+fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool"
+where
"ismaxcoeff (C i) = (\<lambda> x. \<bar>i\<bar> \<le> x)"
| "ismaxcoeff (CN n c t) = (\<lambda>x. \<bar>c\<bar> \<le> x \<and> (ismaxcoeff t x))"
| "ismaxcoeff (CF c s t) = (\<lambda>x. \<bar>c\<bar> \<le> x \<and> (ismaxcoeff t x))"
@@ -711,7 +752,7 @@
using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
consts numadd:: "num \<times> num \<Rightarrow> num"
-recdef numadd "measure (\<lambda> (t,s). size t + size s)"
+recdef numadd "measure (\<lambda>(t, s). size t + size s)"
"numadd (CN n1 c1 r1,CN n2 c2 r2) =
(if n1=n2 then
(let c = c1 + c2
@@ -730,22 +771,14 @@
"numadd (C b1, C b2) = C (b1+b2)"
"numadd (a,b) = Add a b"
-lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
-apply (induct t s rule: numadd.induct, simp_all add: Let_def)
- apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
- apply (case_tac "n1 = n2", simp_all add: algebra_simps)
- apply (simp only: distrib_right[symmetric])
- apply simp
-apply (case_tac "lex_bnd t1 t2", simp_all)
- apply (case_tac "c1+c2 = 0")
- apply (case_tac "t1 = t2")
- apply (simp_all add: algebra_simps distrib_right[symmetric] of_int_mult[symmetric] of_int_add[symmetric]del: of_int_mult of_int_add distrib_right)
- done
-
-lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
- by (induct t s rule: numadd.induct) (auto simp add: Let_def)
-
-fun nummul:: "num \<Rightarrow> int \<Rightarrow> num" where
+lemma numadd [simp]: "Inum bs (numadd (t, s)) = Inum bs (Add t s)"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff)
+
+lemma numadd_nb [simp]: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd (t, s))"
+ by (induct t s rule: numadd.induct) (simp_all add: Let_def)
+
+fun nummul:: "num \<Rightarrow> int \<Rightarrow> num"
+where
"nummul (C j) = (\<lambda> i. C (i*j))"
| "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
| "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
@@ -785,7 +818,8 @@
finally show ?thesis .
qed
-fun split_int:: "num \<Rightarrow> num \<times> num" where
+fun split_int:: "num \<Rightarrow> num \<times> num"
+where
"split_int (C c) = (C 0, C c)"
| "split_int (CN n c b) =
(let (bv,bi) = split_int b
@@ -853,7 +887,8 @@
using split_int_nb[where t="t"]
by (cases "fst (split_int t)") (auto simp add: numfloor_def Let_def split_def)
-function simpnum:: "num \<Rightarrow> num" where
+fun simpnum:: "num \<Rightarrow> num"
+where
"simpnum (C j) = C j"
| "simpnum (Bound n) = CN n 1 (C 0)"
| "simpnum (Neg t) = numneg (simpnum t)"
@@ -863,8 +898,6 @@
| "simpnum (Floor t) = numfloor (simpnum t)"
| "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))"
| "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)"
-by pat_completeness auto
-termination by (relation "measure num_size") auto
lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
by (induct t rule: simpnum.induct) auto
@@ -872,7 +905,8 @@
lemma simpnum_numbound0[simp]: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
by (induct t rule: simpnum.induct) auto
-fun nozerocoeff:: "num \<Rightarrow> bool" where
+fun nozerocoeff:: "num \<Rightarrow> bool"
+where
"nozerocoeff (C c) = True"
| "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
| "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
@@ -1001,7 +1035,8 @@
ultimately show ?thesis by blast
qed
-fun not:: "fm \<Rightarrow> fm" where
+fun not:: "fm \<Rightarrow> fm"
+where
"not (NOT p) = p"
| "not T = F"
| "not F = T"
@@ -1059,12 +1094,13 @@
Iff p q)"
lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
- by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp) (cases "not p= q", auto simp add:not)
+ by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp) (cases "not p= q", auto)
lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
by (unfold iff_def,cases "p=q", auto)
-fun check_int:: "num \<Rightarrow> bool" where
+fun check_int:: "num \<Rightarrow> bool"
+where
"check_int (C i) = True"
| "check_int (Floor t) = True"
| "check_int (Mul i t) = check_int t"
@@ -1139,14 +1175,15 @@
ultimately show ?thesis by blast
qed
-function (sequential) simpfm :: "fm \<Rightarrow> fm" where
+fun simpfm :: "fm \<Rightarrow> fm"
+where
"simpfm (And p q) = conj (simpfm p) (simpfm q)"
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
| "simpfm (NOT p) = not (simpfm p)"
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F
- | _ \<Rightarrow> Lt (reducecoeff a'))"
+ | _ \<Rightarrow> Lt (reducecoeff a'))"
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0) then T else F | _ \<Rightarrow> Le (reducecoeff a'))"
| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0) then T else F | _ \<Rightarrow> Gt (reducecoeff a'))"
| "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<ge> 0) then T else F | _ \<Rightarrow> Ge (reducecoeff a'))"
@@ -1159,8 +1196,6 @@
else if (\<bar>i\<bar> = 1) \<and> check_int a then F
else let a' = simpnum a in case a' of C v \<Rightarrow> if (\<not>(i dvd v)) then T else F | _ \<Rightarrow> (let (d,t) = simpdvd i a' in NDvd d t))"
| "simpfm p = p"
-by pat_completeness auto
-termination by (relation "measure fmsize") auto
lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p"
proof(induct p rule: simpfm.induct)
@@ -1414,7 +1449,8 @@
with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast
qed
-function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" where
+fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+where
"qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
| "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
| "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
@@ -1423,8 +1459,6 @@
| "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
| "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
| "qelim p = (\<lambda> y. simpfm p)"
-by pat_completeness auto
-termination by (relation "measure fmsize") auto
lemma qelim_ci:
assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
@@ -1436,7 +1470,8 @@
text \<open>The \<open>\<int>\<close> Part\<close>
text\<open>Linearity for fm where Bound 0 ranges over \<open>\<int>\<close>\<close>
-function zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) where
+fun zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*)
+where
"zsplit0 (C c) = (0,C c)"
| "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
| "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
@@ -1450,8 +1485,6 @@
in (ia-ib, Sub a' b'))"
| "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))"
| "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))"
-by pat_completeness auto
-termination by (relation "measure num_size") auto
lemma zsplit0_I:
shows "\<And> n a. zsplit0 t = (n,a) \<Longrightarrow> (Inum ((real_of_int (x::int)) #bs) (CN 0 n a) = Inum (real_of_int x #bs) t) \<and> numbound0 a"
@@ -1533,23 +1566,21 @@
with na show ?case by simp
qed
-consts
- iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool" (* Linearity test for fm *)
- zlfm :: "fm \<Rightarrow> fm" (* Linearity transformation for fm *)
-recdef iszlfm "measure size"
+fun iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool" (* Linearity test for fm *)
+where
"iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
- "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
- "iszlfm (Eq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (Lt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (Le (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (Gt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (Ge (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (Dvd i (CN 0 c e)) =
+| "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
+| "iszlfm (Eq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (Lt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (Le (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (Gt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (Ge (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
+| "iszlfm (Dvd i (CN 0 c e)) =
(\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm (NDvd i (CN 0 c e))=
+| "iszlfm (NDvd i (CN 0 c e))=
(\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
- "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
+| "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p"
by (induct p rule: iszlfm.induct) auto
@@ -1569,61 +1600,62 @@
lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs"
using disj_def by (cases p,auto)
-recdef zlfm "measure fmsize"
+fun zlfm :: "fm \<Rightarrow> fm" (* Linearity transformation for fm *)
+where
"zlfm (And p q) = conj (zlfm p) (zlfm q)"
- "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
- "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
- "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
- "zlfm (Lt a) = (let (c,r) = zsplit0 a in
+| "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
+| "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
+| "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
+| "zlfm (Lt a) = (let (c,r) = zsplit0 a in
if c=0 then Lt r else
if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r)))
else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
- "zlfm (Le a) = (let (c,r) = zsplit0 a in
+| "zlfm (Le a) = (let (c,r) = zsplit0 a in
if c=0 then Le r else
if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r)))
else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
- "zlfm (Gt a) = (let (c,r) = zsplit0 a in
+| "zlfm (Gt a) = (let (c,r) = zsplit0 a in
if c=0 then Gt r else
if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r)))
else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
- "zlfm (Ge a) = (let (c,r) = zsplit0 a in
+| "zlfm (Ge a) = (let (c,r) = zsplit0 a in
if c=0 then Ge r else
if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r)))
else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
- "zlfm (Eq a) = (let (c,r) = zsplit0 a in
+| "zlfm (Eq a) = (let (c,r) = zsplit0 a in
if c=0 then Eq r else
if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r)))
else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
- "zlfm (NEq a) = (let (c,r) = zsplit0 a in
+| "zlfm (NEq a) = (let (c,r) = zsplit0 a in
if c=0 then NEq r else
if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r)))
else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
- "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a)
+| "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a)
else (let (c,r) = zsplit0 a in
if c=0 then Dvd \<bar>i\<bar> r else
if c>0 then And (Eq (Sub (Floor r) r)) (Dvd \<bar>i\<bar> (CN 0 c (Floor r)))
else And (Eq (Sub (Floor r) r)) (Dvd \<bar>i\<bar> (CN 0 (-c) (Neg (Floor r))))))"
- "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a)
+| "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a)
else (let (c,r) = zsplit0 a in
if c=0 then NDvd \<bar>i\<bar> r else
if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 c (Floor r)))
else Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 (-c) (Neg (Floor r))))))"
- "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
- "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
- "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
- "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
- "zlfm (NOT (NOT p)) = zlfm p"
- "zlfm (NOT T) = F"
- "zlfm (NOT F) = T"
- "zlfm (NOT (Lt a)) = zlfm (Ge a)"
- "zlfm (NOT (Le a)) = zlfm (Gt a)"
- "zlfm (NOT (Gt a)) = zlfm (Le a)"
- "zlfm (NOT (Ge a)) = zlfm (Lt a)"
- "zlfm (NOT (Eq a)) = zlfm (NEq a)"
- "zlfm (NOT (NEq a)) = zlfm (Eq a)"
- "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
- "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
- "zlfm p = p" (hints simp add: fmsize_pos)
+| "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
+| "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
+| "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
+| "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
+| "zlfm (NOT (NOT p)) = zlfm p"
+| "zlfm (NOT T) = F"
+| "zlfm (NOT F) = T"
+| "zlfm (NOT (Lt a)) = zlfm (Ge a)"
+| "zlfm (NOT (Le a)) = zlfm (Gt a)"
+| "zlfm (NOT (Gt a)) = zlfm (Le a)"
+| "zlfm (NOT (Ge a)) = zlfm (Lt a)"
+| "zlfm (NOT (Eq a)) = zlfm (NEq a)"
+| "zlfm (NOT (NEq a)) = zlfm (Eq a)"
+| "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
+| "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
+| "zlfm p = p"
lemma split_int_less_real:
"(real_of_int (a::int) < b) = (a < \<lfloor>b\<rfloor> \<or> (a = \<lfloor>b\<rfloor> \<and> real_of_int \<lfloor>b\<rfloor> < b))"
@@ -1944,7 +1976,8 @@
\<open>\<delta>\<close> Compute lcm \<open>d| Dvd d c*x+t \<in> p\<close>
\<open>d_\<delta>\<close> checks if a given l divides all the ds above\<close>
-fun minusinf:: "fm \<Rightarrow> fm" where
+fun minusinf:: "fm \<Rightarrow> fm"
+where
"minusinf (And p q) = conj (minusinf p) (minusinf q)"
| "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
| "minusinf (Eq (CN 0 c e)) = F"
@@ -1958,7 +1991,8 @@
lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
by (induct p rule: minusinf.induct, auto)
-fun plusinf:: "fm \<Rightarrow> fm" where
+fun plusinf:: "fm \<Rightarrow> fm"
+where
"plusinf (And p q) = conj (plusinf p) (plusinf q)"
| "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
| "plusinf (Eq (CN 0 c e)) = F"
@@ -1969,14 +2003,16 @@
| "plusinf (Ge (CN 0 c e)) = T"
| "plusinf p = p"
-fun \<delta> :: "fm \<Rightarrow> int" where
+fun \<delta> :: "fm \<Rightarrow> int"
+where
"\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
| "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
| "\<delta> (Dvd i (CN 0 c e)) = i"
| "\<delta> (NDvd i (CN 0 c e)) = i"
| "\<delta> p = 1"
-fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" where
+fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool"
+where
"d_\<delta> (And p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
| "d_\<delta> (Or p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
| "d_\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
@@ -2237,88 +2273,88 @@
from periodic_finite_ex[OF dpos th1] show ?thesis by blast
qed
-lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto
-
-consts
- a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coefficients of a formula *)
- d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
- \<zeta> :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
- \<beta> :: "fm \<Rightarrow> num list"
- \<alpha> :: "fm \<Rightarrow> num list"
-
-recdef a_\<beta> "measure size"
+lemma dvd1_eq1: "x > 0 \<Longrightarrow> is_unit x \<longleftrightarrow> x = 1" for x :: int
+ by simp
+
+fun a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coefficients of a formula *)
+where
"a_\<beta> (And p q) = (\<lambda> k. And (a_\<beta> p k) (a_\<beta> q k))"
- "a_\<beta> (Or p q) = (\<lambda> k. Or (a_\<beta> p k) (a_\<beta> q k))"
- "a_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (Le (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
- "a_\<beta> p = (\<lambda> k. p)"
-
-recdef d_\<beta> "measure size"
+| "a_\<beta> (Or p q) = (\<lambda> k. Or (a_\<beta> p k) (a_\<beta> q k))"
+| "a_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (Le (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
+| "a_\<beta> p = (\<lambda> k. p)"
+
+fun d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
+where
"d_\<beta> (And p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
- "d_\<beta> (Or p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
- "d_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (Le (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. c dvd k)"
- "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
- "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
- "d_\<beta> p = (\<lambda> k. True)"
-
-recdef \<zeta> "measure size"
+| "d_\<beta> (Or p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
+| "d_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (Le (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. c dvd k)"
+| "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
+| "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
+| "d_\<beta> p = (\<lambda> k. True)"
+
+fun \<zeta> :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
+where
"\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
- "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
- "\<zeta> (Eq (CN 0 c e)) = c"
- "\<zeta> (NEq (CN 0 c e)) = c"
- "\<zeta> (Lt (CN 0 c e)) = c"
- "\<zeta> (Le (CN 0 c e)) = c"
- "\<zeta> (Gt (CN 0 c e)) = c"
- "\<zeta> (Ge (CN 0 c e)) = c"
- "\<zeta> (Dvd i (CN 0 c e)) = c"
- "\<zeta> (NDvd i (CN 0 c e))= c"
- "\<zeta> p = 1"
-
-recdef \<beta> "measure size"
+| "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
+| "\<zeta> (Eq (CN 0 c e)) = c"
+| "\<zeta> (NEq (CN 0 c e)) = c"
+| "\<zeta> (Lt (CN 0 c e)) = c"
+| "\<zeta> (Le (CN 0 c e)) = c"
+| "\<zeta> (Gt (CN 0 c e)) = c"
+| "\<zeta> (Ge (CN 0 c e)) = c"
+| "\<zeta> (Dvd i (CN 0 c e)) = c"
+| "\<zeta> (NDvd i (CN 0 c e))= c"
+| "\<zeta> p = 1"
+
+fun \<beta> :: "fm \<Rightarrow> num list"
+where
"\<beta> (And p q) = (\<beta> p @ \<beta> q)"
- "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
- "\<beta> (Eq (CN 0 c e)) = [Sub (C (- 1)) e]"
- "\<beta> (NEq (CN 0 c e)) = [Neg e]"
- "\<beta> (Lt (CN 0 c e)) = []"
- "\<beta> (Le (CN 0 c e)) = []"
- "\<beta> (Gt (CN 0 c e)) = [Neg e]"
- "\<beta> (Ge (CN 0 c e)) = [Sub (C (- 1)) e]"
- "\<beta> p = []"
-
-recdef \<alpha> "measure size"
+| "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
+| "\<beta> (Eq (CN 0 c e)) = [Sub (C (- 1)) e]"
+| "\<beta> (NEq (CN 0 c e)) = [Neg e]"
+| "\<beta> (Lt (CN 0 c e)) = []"
+| "\<beta> (Le (CN 0 c e)) = []"
+| "\<beta> (Gt (CN 0 c e)) = [Neg e]"
+| "\<beta> (Ge (CN 0 c e)) = [Sub (C (- 1)) e]"
+| "\<beta> p = []"
+
+fun \<alpha> :: "fm \<Rightarrow> num list"
+where
"\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)"
- "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)"
- "\<alpha> (Eq (CN 0 c e)) = [Add (C (- 1)) e]"
- "\<alpha> (NEq (CN 0 c e)) = [e]"
- "\<alpha> (Lt (CN 0 c e)) = [e]"
- "\<alpha> (Le (CN 0 c e)) = [Add (C (- 1)) e]"
- "\<alpha> (Gt (CN 0 c e)) = []"
- "\<alpha> (Ge (CN 0 c e)) = []"
- "\<alpha> p = []"
-consts mirror :: "fm \<Rightarrow> fm"
-recdef mirror "measure size"
+| "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)"
+| "\<alpha> (Eq (CN 0 c e)) = [Add (C (- 1)) e]"
+| "\<alpha> (NEq (CN 0 c e)) = [e]"
+| "\<alpha> (Lt (CN 0 c e)) = [e]"
+| "\<alpha> (Le (CN 0 c e)) = [Add (C (- 1)) e]"
+| "\<alpha> (Gt (CN 0 c e)) = []"
+| "\<alpha> (Ge (CN 0 c e)) = []"
+| "\<alpha> p = []"
+
+fun mirror :: "fm \<Rightarrow> fm"
+where
"mirror (And p q) = And (mirror p) (mirror q)"
- "mirror (Or p q) = Or (mirror p) (mirror q)"
- "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))"
- "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
- "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))"
- "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))"
- "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))"
- "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))"
- "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
- "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
- "mirror p = p"
+| "mirror (Or p q) = Or (mirror p) (mirror q)"
+| "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))"
+| "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
+| "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))"
+| "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))"
+| "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))"
+| "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))"
+| "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
+| "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
+| "mirror p = p"
lemma mirror_\<alpha>_\<beta>:
assumes lp: "iszlfm p (a#bs)"
@@ -2770,51 +2806,49 @@
(* Reddy and Loveland *)
-consts
- \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
- \<sigma>_\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
- \<alpha>_\<rho> :: "fm \<Rightarrow> (num\<times>int) list"
- a_\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm"
-recdef \<rho> "measure size"
+fun \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
+ where
"\<rho> (And p q) = (\<rho> p @ \<rho> q)"
- "\<rho> (Or p q) = (\<rho> p @ \<rho> q)"
- "\<rho> (Eq (CN 0 c e)) = [(Sub (C (- 1)) e,c)]"
- "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
- "\<rho> (Lt (CN 0 c e)) = []"
- "\<rho> (Le (CN 0 c e)) = []"
- "\<rho> (Gt (CN 0 c e)) = [(Neg e, c)]"
- "\<rho> (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
- "\<rho> p = []"
-
-recdef \<sigma>_\<rho> "measure size"
+| "\<rho> (Or p q) = (\<rho> p @ \<rho> q)"
+| "\<rho> (Eq (CN 0 c e)) = [(Sub (C (- 1)) e,c)]"
+| "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
+| "\<rho> (Lt (CN 0 c e)) = []"
+| "\<rho> (Le (CN 0 c e)) = []"
+| "\<rho> (Gt (CN 0 c e)) = [(Neg e, c)]"
+| "\<rho> (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
+| "\<rho> p = []"
+
+fun \<sigma>_\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
+where
"\<sigma>_\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
- "\<sigma>_\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
- "\<sigma>_\<rho> (Eq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
+| "\<sigma>_\<rho> (Eq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e))
else (Eq (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e))
else (NEq (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (Lt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Lt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e))
else (Lt (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (Le (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Le (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e))
else (Le (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (Gt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Gt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e))
else (Gt (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (Ge (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Ge (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e))
else (Ge (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (Dvd i (CN 0 c e)) =(\<lambda> (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (Dvd i (CN 0 c e)) =(\<lambda> (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e))
else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> (NDvd i (CN 0 c e))=(\<lambda> (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e))
+| "\<sigma>_\<rho> (NDvd i (CN 0 c e))=(\<lambda> (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e))
else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
- "\<sigma>_\<rho> p = (\<lambda> (t,k). p)"
-
-recdef \<alpha>_\<rho> "measure size"
+| "\<sigma>_\<rho> p = (\<lambda> (t,k). p)"
+
+fun \<alpha>_\<rho> :: "fm \<Rightarrow> (num \<times> int) list"
+where
"\<alpha>_\<rho> (And p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
- "\<alpha>_\<rho> (Or p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
- "\<alpha>_\<rho> (Eq (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
- "\<alpha>_\<rho> (NEq (CN 0 c e)) = [(e,c)]"
- "\<alpha>_\<rho> (Lt (CN 0 c e)) = [(e,c)]"
- "\<alpha>_\<rho> (Le (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
- "\<alpha>_\<rho> p = []"
+| "\<alpha>_\<rho> (Or p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
+| "\<alpha>_\<rho> (Eq (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
+| "\<alpha>_\<rho> (NEq (CN 0 c e)) = [(e,c)]"
+| "\<alpha>_\<rho> (Lt (CN 0 c e)) = [(e,c)]"
+| "\<alpha>_\<rho> (Le (CN 0 c e)) = [(Add (C (- 1)) e,c)]"
+| "\<alpha>_\<rho> p = []"
(* Simulates normal substituion by modifying the formula see correctness theorem *)
@@ -3277,18 +3311,17 @@
text \<open>The \<open>\<real>\<close> part\<close>
text\<open>Linearity for fm where Bound 0 ranges over \<open>\<real>\<close>\<close>
-consts
- isrlfm :: "fm \<Rightarrow> bool" (* Linearity test for fm *)
-recdef isrlfm "measure size"
+fun isrlfm :: "fm \<Rightarrow> bool" (* Linearity test for fm *)
+where
"isrlfm (And p q) = (isrlfm p \<and> isrlfm q)"
- "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)"
- "isrlfm (Eq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm (Lt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm (Le (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm (Gt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm (Ge (CN 0 c e)) = (c>0 \<and> numbound0 e)"
- "isrlfm p = (isatom p \<and> (bound0 p))"
+| "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)"
+| "isrlfm (Eq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm (Lt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm (Le (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm (Gt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm (Ge (CN 0 c e)) = (c>0 \<and> numbound0 e)"
+| "isrlfm p = (isatom p \<and> (bound0 p))"
definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
"fp p n s j \<equiv> (if n > 0 then
@@ -3299,7 +3332,8 @@
(Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))"
(* splits the bounded from the unbounded part*)
-function (sequential) rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" where
+fun rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list"
+where
"rsplit0 (Bound 0) = [(T,1,C 0)]"
| "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b
in map (\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\<leftarrow>acs,b\<leftarrow>bcs])"
@@ -3314,8 +3348,6 @@
| "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
| "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)"
| "rsplit0 t = [(T,0,t)]"
-by pat_completeness auto
-termination by (relation "measure num_size") auto
lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
using conj_def by (cases p, auto)
@@ -3914,36 +3946,36 @@
by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l)
(case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all)
-consts rlfm :: "fm \<Rightarrow> fm"
-recdef rlfm "measure fmsize"
+fun rlfm :: "fm \<Rightarrow> fm"
+where
"rlfm (And p q) = conj (rlfm p) (rlfm q)"
- "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
- "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
- "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
- "rlfm (Lt a) = rsplit lt a"
- "rlfm (Le a) = rsplit le a"
- "rlfm (Gt a) = rsplit gt a"
- "rlfm (Ge a) = rsplit ge a"
- "rlfm (Eq a) = rsplit eq a"
- "rlfm (NEq a) = rsplit neq a"
- "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
- "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
- "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
- "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
- "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
- "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
- "rlfm (NOT (NOT p)) = rlfm p"
- "rlfm (NOT T) = F"
- "rlfm (NOT F) = T"
- "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
- "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
- "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
- "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
- "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
- "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
- "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
- "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
- "rlfm p = p" (hints simp add: fmsize_pos)
+| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
+| "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
+| "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
+| "rlfm (Lt a) = rsplit lt a"
+| "rlfm (Le a) = rsplit le a"
+| "rlfm (Gt a) = rsplit gt a"
+| "rlfm (Ge a) = rsplit ge a"
+| "rlfm (Eq a) = rsplit eq a"
+| "rlfm (NEq a) = rsplit neq a"
+| "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
+| "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
+| "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
+| "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
+| "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
+| "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
+| "rlfm (NOT (NOT p)) = rlfm p"
+| "rlfm (NOT T) = F"
+| "rlfm (NOT F) = T"
+| "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
+| "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
+| "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
+| "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
+| "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
+| "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
+| "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
+| "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
+| "rlfm p = p"
lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
by (induct p rule: isrlfm.induct, auto)
@@ -4377,30 +4409,29 @@
ultimately show ?thesis using z_def by auto
qed
-consts
- \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
- \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
-recdef \<Upsilon> "measure size"
+fun \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
+where
"\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)"
- "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)"
- "\<Upsilon> (Eq (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> (Lt (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> (Le (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> (Gt (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> (Ge (CN 0 c e)) = [(Neg e,c)]"
- "\<Upsilon> p = []"
-
-recdef \<upsilon> "measure size"
+| "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)"
+| "\<Upsilon> (Eq (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> (Lt (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> (Le (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> (Gt (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> (Ge (CN 0 c e)) = [(Neg e,c)]"
+| "\<Upsilon> p = []"
+
+fun \<upsilon> :: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm"
+where
"\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
- "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
- "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
- "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
- "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
- "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
- "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
- "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
- "\<upsilon> p = (\<lambda> (t,n). p)"
+| "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
+| "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
+| "\<upsilon> p = (\<lambda> (t,n). p)"
lemma \<upsilon>_I: assumes lp: "isrlfm p"
and np: "real_of_int n > 0" and nbt: "numbound0 t"
@@ -4765,7 +4796,8 @@
ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real_of_int i + u))" by blast
qed
-fun exsplitnum :: "num \<Rightarrow> num" where
+fun exsplitnum :: "num \<Rightarrow> num"
+where
"exsplitnum (C c) = (C c)"
| "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
| "exsplitnum (Bound n) = Bound (n+1)"
@@ -4778,7 +4810,8 @@
| "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
| "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
-fun exsplit :: "fm \<Rightarrow> fm" where
+fun exsplit :: "fm \<Rightarrow> fm"
+where
"exsplit (Lt a) = Lt (exsplitnum a)"
| "exsplit (Le a) = Le (exsplitnum a)"
| "exsplit (Gt a) = Gt (exsplitnum a)"
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Oct 09 22:08:05 2017 +0200
@@ -10,24 +10,29 @@
Dense_Linear_Order
DP_Library
"HOL-Library.Code_Target_Numeral"
- "HOL-Library.Old_Recdef"
begin
subsection \<open>Terms\<close>
-datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm
+datatype (plugins del: size) tm = CP poly | Bound nat | Add tm tm | Mul poly tm
| Neg tm | Sub tm tm | CNP nat poly tm
-text \<open>A size for poly to make inductive proofs simpler.\<close>
-primrec tmsize :: "tm \<Rightarrow> nat"
+instantiation tm :: size
+begin
+
+primrec size_tm :: "tm \<Rightarrow> nat"
where
- "tmsize (CP c) = polysize c"
-| "tmsize (Bound n) = 1"
-| "tmsize (Neg a) = 1 + tmsize a"
-| "tmsize (Add a b) = 1 + tmsize a + tmsize b"
-| "tmsize (Sub a b) = 3 + tmsize a + tmsize b"
-| "tmsize (Mul c a) = 1 + polysize c + tmsize a"
-| "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
+ "size_tm (CP c) = polysize c"
+| "size_tm (Bound n) = 1"
+| "size_tm (Neg a) = 1 + size_tm a"
+| "size_tm (Add a b) = 1 + size_tm a + size_tm b"
+| "size_tm (Sub a b) = 3 + size_tm a + size_tm b"
+| "size_tm (Mul c a) = 1 + polysize c + size_tm a"
+| "size_tm (CNP n c a) = 3 + polysize c + size_tm a "
+
+instance ..
+
+end
text \<open>Semantics of terms tm.\<close>
primrec Itm :: "'a::{field_char_0,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
@@ -258,42 +263,42 @@
text \<open>Simplification.\<close>
-consts tmadd:: "tm \<times> tm \<Rightarrow> tm"
-recdef tmadd "measure (\<lambda>(t,s). size t + size s)"
- "tmadd (CNP n1 c1 r1,CNP n2 c2 r2) =
+fun tmadd:: "tm \<Rightarrow> tm \<Rightarrow> tm"
+where
+ "tmadd (CNP n1 c1 r1) (CNP n2 c2 r2) =
(if n1 = n2 then
let c = c1 +\<^sub>p c2
- in if c = 0\<^sub>p then tmadd(r1,r2) else CNP n1 c (tmadd (r1, r2))
- else if n1 \<le> n2 then (CNP n1 c1 (tmadd (r1,CNP n2 c2 r2)))
- else (CNP n2 c2 (tmadd (CNP n1 c1 r1, r2))))"
- "tmadd (CNP n1 c1 r1, t) = CNP n1 c1 (tmadd (r1, t))"
- "tmadd (t, CNP n2 c2 r2) = CNP n2 c2 (tmadd (t, r2))"
- "tmadd (CP b1, CP b2) = CP (b1 +\<^sub>p b2)"
- "tmadd (a, b) = Add a b"
-
-lemma tmadd[simp]: "Itm vs bs (tmadd (t, s)) = Itm vs bs (Add t s)"
+ in if c = 0\<^sub>p then tmadd r1 r2 else CNP n1 c (tmadd r1 r2)
+ else if n1 \<le> n2 then (CNP n1 c1 (tmadd r1 (CNP n2 c2 r2)))
+ else (CNP n2 c2 (tmadd (CNP n1 c1 r1) r2)))"
+| "tmadd (CNP n1 c1 r1) t = CNP n1 c1 (tmadd r1 t)"
+| "tmadd t (CNP n2 c2 r2) = CNP n2 c2 (tmadd t r2)"
+| "tmadd (CP b1) (CP b2) = CP (b1 +\<^sub>p b2)"
+| "tmadd a b = Add a b"
+
+lemma tmadd [simp]: "Itm vs bs (tmadd t s) = Itm vs bs (Add t s)"
apply (induct t s rule: tmadd.induct)
apply (simp_all add: Let_def)
apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p")
apply (case_tac "n1 \<le> n2")
apply simp_all
apply (case_tac "n1 = n2")
- apply (simp_all add: field_simps)
- apply (simp only: distrib_left[symmetric])
- apply (auto simp del: polyadd simp add: polyadd[symmetric])
+ apply (simp_all add: algebra_simps)
+ apply (simp only: distrib_left [symmetric] polyadd [symmetric])
+ apply simp
done
-lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd (t, s))"
+lemma tmadd_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 s \<Longrightarrow> tmbound0 (tmadd t s)"
by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
-lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd (t, s))"
+lemma tmadd_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n s \<Longrightarrow> tmbound n (tmadd t s)"
by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
-lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd (t, s))"
+lemma tmadd_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n s \<Longrightarrow> tmboundslt n (tmadd t s)"
by (induct t s rule: tmadd.induct) (auto simp add: Let_def)
lemma tmadd_allpolys_npoly[simp]:
- "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t, s))"
+ "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd t s)"
by (induct t s rule: tmadd.induct) (simp_all add: Let_def polyadd_norm)
fun tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
@@ -323,7 +328,7 @@
where "tmneg t \<equiv> tmmul t (C (- 1,1))"
definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
- where "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s, tmneg t))"
+ where "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd s (tmneg t))"
lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
using tmneg_def[of t] by simp
@@ -367,12 +372,12 @@
"simptm (CP j) = CP (polynate j)"
| "simptm (Bound n) = CNP n (1)\<^sub>p (CP 0\<^sub>p)"
| "simptm (Neg t) = tmneg (simptm t)"
-| "simptm (Add t s) = tmadd (simptm t,simptm s)"
+| "simptm (Add t s) = tmadd (simptm t) (simptm s)"
| "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
| "simptm (Mul i t) =
(let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
| "simptm (CNP n c t) =
- (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))"
+ (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p)) (simptm t))"
lemma polynate_stupid:
assumes "SORT_CONSTRAINT('a::{field_char_0,field})"
@@ -497,24 +502,34 @@
subsection \<open>Formulae\<close>
-datatype fm = T| F| Le tm | Lt tm | Eq tm | NEq tm|
- NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
-
-
-text \<open>A size for fm.\<close>
-fun fmsize :: "fm \<Rightarrow> nat"
+datatype (plugins del: size) fm = T | F | Le tm | Lt tm | Eq tm | NEq tm |
+ NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
+
+instantiation fm :: size
+begin
+
+primrec size_fm :: "fm \<Rightarrow> nat"
where
- "fmsize (NOT p) = 1 + fmsize p"
-| "fmsize (And p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
-| "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
-| "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
-| "fmsize (E p) = 1 + fmsize p"
-| "fmsize (A p) = 4+ fmsize p"
-| "fmsize p = 1"
-
-lemma fmsize_pos[termination_simp]: "fmsize p > 0"
- by (induct p rule: fmsize.induct) simp_all
+ "size_fm (NOT p) = 1 + size_fm p"
+| "size_fm (And p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
+| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
+| "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)"
+| "size_fm (E p) = 1 + size_fm p"
+| "size_fm (A p) = 4 + size_fm p"
+| "size_fm T = 1"
+| "size_fm F = 1"
+| "size_fm (Le _) = 1"
+| "size_fm (Lt _) = 1"
+| "size_fm (Eq _) = 1"
+| "size_fm (NEq _) = 1"
+
+instance ..
+
+end
+
+lemma fmsize_pos [simp]: "size p > 0" for p :: fm
+ by (induct p) simp_all
text \<open>Semantics of formulae (fm).\<close>
primrec Ifm ::"'a::linordered_field list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
@@ -1507,30 +1522,30 @@
by blast
qed
-consts simpfm :: "fm \<Rightarrow> fm"
-recdef simpfm "measure fmsize"
+fun simpfm :: "fm \<Rightarrow> fm"
+where
"simpfm (Lt t) = simplt (simptm t)"
- "simpfm (Le t) = simple (simptm t)"
- "simpfm (Eq t) = simpeq(simptm t)"
- "simpfm (NEq t) = simpneq(simptm t)"
- "simpfm (And p q) = conj (simpfm p) (simpfm q)"
- "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
- "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
- "simpfm (Iff p q) =
+| "simpfm (Le t) = simple (simptm t)"
+| "simpfm (Eq t) = simpeq(simptm t)"
+| "simpfm (NEq t) = simpneq(simptm t)"
+| "simpfm (And p q) = conj (simpfm p) (simpfm q)"
+| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
+| "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
+| "simpfm (Iff p q) =
disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
- "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
- "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
- "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
- "simpfm (NOT (Iff p q)) =
+| "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
+| "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
+| "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
+| "simpfm (NOT (Iff p q)) =
disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
- "simpfm (NOT (Eq t)) = simpneq t"
- "simpfm (NOT (NEq t)) = simpeq t"
- "simpfm (NOT (Le t)) = simplt (Neg t)"
- "simpfm (NOT (Lt t)) = simple (Neg t)"
- "simpfm (NOT (NOT p)) = simpfm p"
- "simpfm (NOT T) = F"
- "simpfm (NOT F) = T"
- "simpfm p = p"
+| "simpfm (NOT (Eq t)) = simpneq t"
+| "simpfm (NOT (NEq t)) = simpeq t"
+| "simpfm (NOT (Le t)) = simplt (Neg t)"
+| "simpfm (NOT (Lt t)) = simple (Neg t)"
+| "simpfm (NOT (NOT p)) = simpfm p"
+| "simpfm (NOT T) = F"
+| "simpfm (NOT F) = T"
+| "simpfm p = p"
lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
by (induct p arbitrary: bs rule: simpfm.induct) auto
@@ -1589,39 +1604,38 @@
shows "qfree p \<Longrightarrow> islin (simpfm p)"
by (induct p rule: simpfm.induct) (simp_all add: conj_lin disj_lin)
-consts prep :: "fm \<Rightarrow> fm"
-recdef prep "measure fmsize"
+fun prep :: "fm \<Rightarrow> fm"
+where
"prep (E T) = T"
- "prep (E F) = F"
- "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
- "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
- "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
- "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
- "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
- "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
- "prep (E p) = E (prep p)"
- "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
- "prep (A p) = prep (NOT (E (NOT p)))"
- "prep (NOT (NOT p)) = prep p"
- "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
- "prep (NOT (A p)) = prep (E (NOT p))"
- "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
- "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
- "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
- "prep (NOT p) = not (prep p)"
- "prep (Or p q) = disj (prep p) (prep q)"
- "prep (And p q) = conj (prep p) (prep q)"
- "prep (Imp p q) = prep (Or (NOT p) q)"
- "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
- "prep p = p"
- (hints simp add: fmsize_pos)
+| "prep (E F) = F"
+| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
+| "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
+| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
+| "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
+| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+| "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+| "prep (E p) = E (prep p)"
+| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
+| "prep (A p) = prep (NOT (E (NOT p)))"
+| "prep (NOT (NOT p)) = prep p"
+| "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
+| "prep (NOT (A p)) = prep (E (NOT p))"
+| "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
+| "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
+| "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
+| "prep (NOT p) = not (prep p)"
+| "prep (Or p q) = disj (prep p) (prep q)"
+| "prep (And p q) = conj (prep p) (prep q)"
+| "prep (Imp p q) = prep (Or (NOT p) q)"
+| "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+| "prep p = p"
lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
by (induct p arbitrary: bs rule: prep.induct) auto
text \<open>Generic quantifier elimination.\<close>
-function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
@@ -1631,8 +1645,6 @@
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
| "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))"
| "qelim p = (\<lambda>y. simpfm p)"
- by pat_completeness simp_all
-termination by (relation "measure fmsize") auto
lemma qelim:
assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
@@ -2667,7 +2679,7 @@
also have "\<dots> \<longleftrightarrow> 2 * ?d * (?a * (-?s / (2*?d)) + ?r) = 0"
using cd(2) mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * (2*?d / (2*?d)) + 2 * ?d * ?r= 0"
- by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left)
+ by (simp add: field_simps distrib_left [of "2*?d"])
also have "\<dots> \<longleftrightarrow> - (?a * ?s) + 2*?d*?r = 0"
using cd(2) by simp
finally show ?thesis
@@ -2684,7 +2696,7 @@
also have "\<dots> \<longleftrightarrow> 2 * ?c * (?a * (-?t / (2 * ?c)) + ?r) = 0"
using cd(1) mult_cancel_left[of "2 * ?c" "(?a * (-?t / (2 * ?c)) + ?r)" 0] by simp
also have "\<dots> \<longleftrightarrow> (?a * -?t)* (2 * ?c) / (2 * ?c) + 2 * ?c * ?r= 0"
- by (simp add: field_simps distrib_left[of "2 * ?c"] del: distrib_left)
+ by (simp add: field_simps distrib_left [of "2 * ?c"])
also have "\<dots> \<longleftrightarrow> - (?a * ?t) + 2 * ?c * ?r = 0"
using cd(1) by simp
finally show ?thesis using cd
--- a/src/HOL/Divides.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Divides.thy Mon Oct 09 22:08:05 2017 +0200
@@ -9,530 +9,17 @@
imports Parity
begin
-subsection \<open>Quotient and remainder in integral domains with additional properties\<close>
-
-class semiring_div = semidom_modulo +
- assumes div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
- and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b"
-begin
-
-lemma div_mult_self2 [simp]:
- assumes "b \<noteq> 0"
- shows "(a + b * c) div b = c + a div b"
- using assms div_mult_self1 [of b a c] by (simp add: mult.commute)
-
-lemma div_mult_self3 [simp]:
- assumes "b \<noteq> 0"
- shows "(c * b + a) div b = c + a div b"
- using assms by (simp add: add.commute)
-
-lemma div_mult_self4 [simp]:
- assumes "b \<noteq> 0"
- shows "(b * c + a) div b = c + a div b"
- using assms by (simp add: add.commute)
-
-lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
-proof (cases "b = 0")
- case True then show ?thesis by simp
-next
- case False
- have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b"
- by (simp add: div_mult_mod_eq)
- also from False div_mult_self1 [of b a c] have
- "\<dots> = (c + a div b) * b + (a + c * b) mod b"
- by (simp add: algebra_simps)
- finally have "a = a div b * b + (a + c * b) mod b"
- by (simp add: add.commute [of a] add.assoc distrib_right)
- then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b"
- by (simp add: div_mult_mod_eq)
- then show ?thesis by simp
-qed
-
-lemma mod_mult_self2 [simp]:
- "(a + b * c) mod b = a mod b"
- by (simp add: mult.commute [of b])
-
-lemma mod_mult_self3 [simp]:
- "(c * b + a) mod b = a mod b"
- by (simp add: add.commute)
-
-lemma mod_mult_self4 [simp]:
- "(b * c + a) mod b = a mod b"
- by (simp add: add.commute)
-
-lemma mod_mult_self1_is_0 [simp]:
- "b * a mod b = 0"
- using mod_mult_self2 [of 0 b a] by simp
-
-lemma mod_mult_self2_is_0 [simp]:
- "a * b mod b = 0"
- using mod_mult_self1 [of 0 a b] by simp
-
-lemma div_add_self1:
- assumes "b \<noteq> 0"
- shows "(b + a) div b = a div b + 1"
- using assms div_mult_self1 [of b a 1] by (simp add: add.commute)
-
-lemma div_add_self2:
- assumes "b \<noteq> 0"
- shows "(a + b) div b = a div b + 1"
- using assms div_add_self1 [of b a] by (simp add: add.commute)
-
-lemma mod_add_self1 [simp]:
- "(b + a) mod b = a mod b"
- using mod_mult_self1 [of a 1 b] by (simp add: add.commute)
-
-lemma mod_add_self2 [simp]:
- "(a + b) mod b = a mod b"
- using mod_mult_self1 [of a 1 b] by simp
-
-lemma mod_div_trivial [simp]:
- "a mod b div b = 0"
-proof (cases "b = 0")
- assume "b = 0"
- thus ?thesis by simp
-next
- assume "b \<noteq> 0"
- hence "a div b + a mod b div b = (a mod b + a div b * b) div b"
- by (rule div_mult_self1 [symmetric])
- also have "\<dots> = a div b"
- by (simp only: mod_div_mult_eq)
- also have "\<dots> = a div b + 0"
- by simp
- finally show ?thesis
- by (rule add_left_imp_eq)
-qed
-
-lemma mod_mod_trivial [simp]:
- "a mod b mod b = a mod b"
-proof -
- have "a mod b mod b = (a mod b + a div b * b) mod b"
- by (simp only: mod_mult_self1)
- also have "\<dots> = a mod b"
- by (simp only: mod_div_mult_eq)
- finally show ?thesis .
-qed
-
-lemma mod_mod_cancel:
- assumes "c dvd b"
- shows "a mod b mod c = a mod c"
-proof -
- from \<open>c dvd b\<close> obtain k where "b = c * k"
- by (rule dvdE)
- have "a mod b mod c = a mod (c * k) mod c"
- by (simp only: \<open>b = c * k\<close>)
- also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
- by (simp only: mod_mult_self1)
- also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
- by (simp only: ac_simps)
- also have "\<dots> = a mod c"
- by (simp only: div_mult_mod_eq)
- finally show ?thesis .
-qed
-
-lemma div_mult_mult2 [simp]:
- "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b"
- by (drule div_mult_mult1) (simp add: mult.commute)
-
-lemma div_mult_mult1_if [simp]:
- "(c * a) div (c * b) = (if c = 0 then 0 else a div b)"
- by simp_all
-
-lemma mod_mult_mult1:
- "(c * a) mod (c * b) = c * (a mod b)"
-proof (cases "c = 0")
- case True then show ?thesis by simp
-next
- case False
- from div_mult_mod_eq
- have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" .
- with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b)
- = c * a + c * (a mod b)" by (simp add: algebra_simps)
- with div_mult_mod_eq show ?thesis by simp
-qed
-
-lemma mod_mult_mult2:
- "(a * c) mod (b * c) = (a mod b) * c"
- using mod_mult_mult1 [of c a b] by (simp add: mult.commute)
-
-lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)"
- by (fact mod_mult_mult2 [symmetric])
-
-lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)"
- by (fact mod_mult_mult1 [symmetric])
-
-lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)"
- unfolding dvd_def by (auto simp add: mod_mult_mult1)
-
-lemma div_plus_div_distrib_dvd_left:
- "c dvd a \<Longrightarrow> (a + b) div c = a div c + b div c"
- by (cases "c = 0") (auto elim: dvdE)
-
-lemma div_plus_div_distrib_dvd_right:
- "c dvd b \<Longrightarrow> (a + b) div c = a div c + b div c"
- using div_plus_div_distrib_dvd_left [of c b a]
- by (simp add: ac_simps)
-
-named_theorems mod_simps
-
-text \<open>Addition respects modular equivalence.\<close>
-
-lemma mod_add_left_eq [mod_simps]:
- "(a mod c + b) mod c = (a + b) mod c"
-proof -
- have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a mod c + b + a div c * c) mod c"
- by (simp only: ac_simps)
- also have "\<dots> = (a mod c + b) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis
- by (rule sym)
-qed
-
-lemma mod_add_right_eq [mod_simps]:
- "(a + b mod c) mod c = (a + b) mod c"
- using mod_add_left_eq [of b c a] by (simp add: ac_simps)
-
-lemma mod_add_eq:
- "(a mod c + b mod c) mod c = (a + b) mod c"
- by (simp add: mod_add_left_eq mod_add_right_eq)
-
-lemma mod_sum_eq [mod_simps]:
- "(\<Sum>i\<in>A. f i mod a) mod a = sum f A mod a"
-proof (induct A rule: infinite_finite_induct)
- case (insert i A)
- then have "(\<Sum>i\<in>insert i A. f i mod a) mod a
- = (f i mod a + (\<Sum>i\<in>A. f i mod a)) mod a"
- by simp
- also have "\<dots> = (f i + (\<Sum>i\<in>A. f i mod a) mod a) mod a"
- by (simp add: mod_simps)
- also have "\<dots> = (f i + (\<Sum>i\<in>A. f i) mod a) mod a"
- by (simp add: insert.hyps)
- finally show ?case
- by (simp add: insert.hyps mod_simps)
-qed simp_all
-
-lemma mod_add_cong:
- assumes "a mod c = a' mod c"
- assumes "b mod c = b' mod c"
- shows "(a + b) mod c = (a' + b') mod c"
-proof -
- have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
- unfolding assms ..
- then show ?thesis
- by (simp add: mod_add_eq)
-qed
-
-text \<open>Multiplication respects modular equivalence.\<close>
-
-lemma mod_mult_left_eq [mod_simps]:
- "((a mod c) * b) mod c = (a * b) mod c"
-proof -
- have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
- by (simp only: algebra_simps)
- also have "\<dots> = (a mod c * b) mod c"
- by (rule mod_mult_self1)
- finally show ?thesis
- by (rule sym)
-qed
-
-lemma mod_mult_right_eq [mod_simps]:
- "(a * (b mod c)) mod c = (a * b) mod c"
- using mod_mult_left_eq [of b c a] by (simp add: ac_simps)
-
-lemma mod_mult_eq:
- "((a mod c) * (b mod c)) mod c = (a * b) mod c"
- by (simp add: mod_mult_left_eq mod_mult_right_eq)
-
-lemma mod_prod_eq [mod_simps]:
- "(\<Prod>i\<in>A. f i mod a) mod a = prod f A mod a"
-proof (induct A rule: infinite_finite_induct)
- case (insert i A)
- then have "(\<Prod>i\<in>insert i A. f i mod a) mod a
- = (f i mod a * (\<Prod>i\<in>A. f i mod a)) mod a"
- by simp
- also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i mod a) mod a)) mod a"
- by (simp add: mod_simps)
- also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i) mod a)) mod a"
- by (simp add: insert.hyps)
- finally show ?case
- by (simp add: insert.hyps mod_simps)
-qed simp_all
-
-lemma mod_mult_cong:
- assumes "a mod c = a' mod c"
- assumes "b mod c = b' mod c"
- shows "(a * b) mod c = (a' * b') mod c"
-proof -
- have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
- unfolding assms ..
- then show ?thesis
- by (simp add: mod_mult_eq)
-qed
-
-text \<open>Exponentiation respects modular equivalence.\<close>
-
-lemma power_mod [mod_simps]:
- "((a mod b) ^ n) mod b = (a ^ n) mod b"
-proof (induct n)
- case 0
- then show ?case by simp
-next
- case (Suc n)
- have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b"
- by (simp add: mod_mult_right_eq)
- with Suc show ?case
- by (simp add: mod_mult_left_eq mod_mult_right_eq)
-qed
-
-end
-
-class ring_div = comm_ring_1 + semiring_div
-begin
-
-subclass idom_divide ..
-
-lemma div_minus_minus [simp]: "(- a) div (- b) = a div b"
- using div_mult_mult1 [of "- 1" a b] by simp
-
-lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)"
- using mod_mult_mult1 [of "- 1" a b] by simp
-
-lemma div_minus_right: "a div (- b) = (- a) div b"
- using div_minus_minus [of "- a" b] by simp
-
-lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)"
- using mod_minus_minus [of "- a" b] by simp
-
-lemma div_minus1_right [simp]: "a div (- 1) = - a"
- using div_minus_right [of a 1] by simp
-
-lemma mod_minus1_right [simp]: "a mod (- 1) = 0"
- using mod_minus_right [of a 1] by simp
-
-text \<open>Negation respects modular equivalence.\<close>
-
-lemma mod_minus_eq [mod_simps]:
- "(- (a mod b)) mod b = (- a) mod b"
-proof -
- have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
- by (simp only: div_mult_mod_eq)
- also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b"
- by (simp add: ac_simps)
- also have "\<dots> = (- (a mod b)) mod b"
- by (rule mod_mult_self1)
- finally show ?thesis
- by (rule sym)
-qed
-
-lemma mod_minus_cong:
- assumes "a mod b = a' mod b"
- shows "(- a) mod b = (- a') mod b"
-proof -
- have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
- unfolding assms ..
- then show ?thesis
- by (simp add: mod_minus_eq)
-qed
-
-text \<open>Subtraction respects modular equivalence.\<close>
-
-lemma mod_diff_left_eq [mod_simps]:
- "(a mod c - b) mod c = (a - b) mod c"
- using mod_add_cong [of a c "a mod c" "- b" "- b"]
- by simp
-
-lemma mod_diff_right_eq [mod_simps]:
- "(a - b mod c) mod c = (a - b) mod c"
- using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
- by simp
-
-lemma mod_diff_eq:
- "(a mod c - b mod c) mod c = (a - b) mod c"
- using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
- by simp
-
-lemma mod_diff_cong:
- assumes "a mod c = a' mod c"
- assumes "b mod c = b' mod c"
- shows "(a - b) mod c = (a' - b') mod c"
- using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"]
- by simp
-
-lemma minus_mod_self2 [simp]:
- "(a - b) mod b = a mod b"
- using mod_diff_right_eq [of a b b]
- by (simp add: mod_diff_right_eq)
-
-lemma minus_mod_self1 [simp]:
- "(b - a) mod b = - a mod b"
- using mod_add_self2 [of "- a" b] by simp
-
-end
-
-
-subsection \<open>Euclidean (semi)rings with cancel rules\<close>
-
-class euclidean_semiring_cancel = euclidean_semiring + semiring_div
-
-class euclidean_ring_cancel = euclidean_ring + ring_div
-
-context unique_euclidean_semiring
-begin
-
-subclass euclidean_semiring_cancel
-proof
- show "(a + c * b) div b = c + a div b" if "b \<noteq> 0" for a b c
- proof (cases a b rule: divmod_cases)
- case by0
- with \<open>b \<noteq> 0\<close> show ?thesis
- by simp
- next
- case (divides q)
- then show ?thesis
- by (simp add: ac_simps)
- next
- case (remainder q r)
- then show ?thesis
- by (auto intro: div_eqI simp add: algebra_simps)
- qed
-next
- show"(c * a) div (c * b) = a div b" if "c \<noteq> 0" for a b c
- proof (cases a b rule: divmod_cases)
- case by0
- then show ?thesis
- by simp
- next
- case (divides q)
- with \<open>c \<noteq> 0\<close> show ?thesis
- by (simp add: mult.left_commute [of c])
- next
- case (remainder q r)
- from \<open>b \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have "b * c \<noteq> 0"
- by simp
- from remainder \<open>c \<noteq> 0\<close>
- have "uniqueness_constraint (r * c) (b * c)"
- and "euclidean_size (r * c) < euclidean_size (b * c)"
- by (simp_all add: uniqueness_constraint_mono_mult uniqueness_constraint_mod size_mono_mult)
- with remainder show ?thesis
- by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps)
- (use \<open>b * c \<noteq> 0\<close> in simp)
- qed
-qed
-
-end
-
-context unique_euclidean_ring
-begin
-
-subclass euclidean_ring_cancel ..
-
-end
-
-
-subsection \<open>Parity\<close>
-
-class semiring_div_parity = semiring_div + comm_semiring_1_cancel + numeral +
- assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
- assumes one_mod_two_eq_one [simp]: "1 mod 2 = 1"
- assumes zero_not_eq_two: "0 \<noteq> 2"
-begin
-
-lemma parity_cases [case_names even odd]:
- assumes "a mod 2 = 0 \<Longrightarrow> P"
- assumes "a mod 2 = 1 \<Longrightarrow> P"
- shows P
- using assms parity by blast
-
-lemma one_div_two_eq_zero [simp]:
- "1 div 2 = 0"
-proof (cases "2 = 0")
- case True then show ?thesis by simp
-next
- case False
- from div_mult_mod_eq have "1 div 2 * 2 + 1 mod 2 = 1" .
- with one_mod_two_eq_one have "1 div 2 * 2 + 1 = 1" by simp
- then have "1 div 2 * 2 = 0" by (simp add: ac_simps add_left_imp_eq del: mult_eq_0_iff)
- then have "1 div 2 = 0 \<or> 2 = 0" by simp
- with False show ?thesis by auto
-qed
-
-lemma not_mod_2_eq_0_eq_1 [simp]:
- "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1"
- by (cases a rule: parity_cases) simp_all
-
-lemma not_mod_2_eq_1_eq_0 [simp]:
- "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0"
- by (cases a rule: parity_cases) simp_all
-
-subclass semiring_parity
-proof (unfold_locales, unfold dvd_eq_mod_eq_0 not_mod_2_eq_0_eq_1)
- show "1 mod 2 = 1"
- by (fact one_mod_two_eq_one)
-next
- fix a b
- assume "a mod 2 = 1"
- moreover assume "b mod 2 = 1"
- ultimately show "(a + b) mod 2 = 0"
- using mod_add_eq [of a 2 b] by simp
-next
- fix a b
- assume "(a * b) mod 2 = 0"
- then have "(a mod 2) * (b mod 2) mod 2 = 0"
- by (simp add: mod_mult_eq)
- then have "(a mod 2) * (b mod 2) = 0"
- by (cases "a mod 2 = 0") simp_all
- then show "a mod 2 = 0 \<or> b mod 2 = 0"
- by (rule divisors_zero)
-next
- fix a
- assume "a mod 2 = 1"
- then have "a = a div 2 * 2 + 1"
- using div_mult_mod_eq [of a 2] by simp
- then show "\<exists>b. a = b + 1" ..
-qed
-
-lemma even_iff_mod_2_eq_zero:
- "even a \<longleftrightarrow> a mod 2 = 0"
- by (fact dvd_eq_mod_eq_0)
-
-lemma odd_iff_mod_2_eq_one:
- "odd a \<longleftrightarrow> a mod 2 = 1"
- by (simp add: even_iff_mod_2_eq_zero)
-
-lemma even_succ_div_two [simp]:
- "even a \<Longrightarrow> (a + 1) div 2 = a div 2"
- by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero)
-
-lemma odd_succ_div_two [simp]:
- "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1"
- by (auto elim!: oddE simp add: zero_not_eq_two [symmetric] add.assoc)
-
-lemma even_two_times_div_two:
- "even a \<Longrightarrow> 2 * (a div 2) = a"
- by (fact dvd_mult_div_cancel)
-
-lemma odd_two_times_div_two_succ [simp]:
- "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a"
- using mult_div_mod_eq [of 2 a] by (simp add: even_iff_mod_2_eq_zero)
-
-end
-
-
subsection \<open>Numeral division with a pragmatic type class\<close>
text \<open>
The following type class contains everything necessary to formulate
a division algorithm in ring structures with numerals, restricted
- to its positive segments. This is its primary motiviation, and it
+ to its positive segments. This is its primary motivation, and it
could surely be formulated using a more fine-grained, more algebraic
and less technical class hierarchy.
\<close>
-class semiring_numeral_div = semiring_div + comm_semiring_1_cancel + linordered_semidom +
+class unique_euclidean_semiring_numeral = semiring_parity + linordered_semidom +
assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0"
and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a"
and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0"
@@ -551,32 +38,8 @@
\<comment> \<open>These are conceptually definitions but force generated code
to be monomorphic wrt. particular instances of this class which
yields a significant speedup.\<close>
-
begin
-subclass semiring_div_parity
-proof
- fix a
- show "a mod 2 = 0 \<or> a mod 2 = 1"
- proof (rule ccontr)
- assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
- then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
- have "0 < 2" by simp
- with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
- with \<open>a mod 2 \<noteq> 0\<close> have "0 < a mod 2" by simp
- with discrete have "1 \<le> a mod 2" by simp
- with \<open>a mod 2 \<noteq> 1\<close> have "1 < a mod 2" by simp
- with discrete have "2 \<le> a mod 2" by simp
- with \<open>a mod 2 < 2\<close> show False by simp
- qed
-next
- show "1 mod 2 = 1"
- by (rule mod_less) simp_all
-next
- show "0 \<noteq> 2"
- by simp
-qed
-
lemma divmod_digit_1:
assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)"
shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P")
@@ -588,7 +51,7 @@
then have [simp]: "1 \<le> a div b" by (simp add: discrete)
with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound)
define w where "w = a div b mod 2"
- with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+ then have w_exhaust: "w = 0 \<or> w = 1" by auto
have mod_w: "a mod (2 * b) = a mod b + b * w"
by (simp add: w_def mod_mult2_eq ac_simps)
from assms w_exhaust have "w = 1"
@@ -607,7 +70,7 @@
and "a mod (2 * b) = a mod b" (is "?Q")
proof -
define w where "w = a div b mod 2"
- with parity have w_exhaust: "w = 0 \<or> w = 1" by auto
+ then have w_exhaust: "w = 0 \<or> w = 1" by auto
have mod_w: "a mod (2 * b) = a mod b + b * w"
by (simp add: w_def mod_mult2_eq ac_simps)
moreover have "b \<le> a mod b + b"
@@ -807,463 +270,12 @@
end
-
-subsection \<open>Division on @{typ nat}\<close>
-
-context
-begin
-
-text \<open>
- We define @{const divide} and @{const modulo} on @{typ nat} by means
- of a characteristic relation with two input arguments
- @{term "m::nat"}, @{term "n::nat"} and two output arguments
- @{term "q::nat"}(uotient) and @{term "r::nat"}(emainder).
-\<close>
-
-inductive eucl_rel_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool"
- where eucl_rel_nat_by0: "eucl_rel_nat m 0 (0, m)"
- | eucl_rel_natI: "r < n \<Longrightarrow> m = q * n + r \<Longrightarrow> eucl_rel_nat m n (q, r)"
-
-text \<open>@{const eucl_rel_nat} is total:\<close>
-
-qualified lemma eucl_rel_nat_ex:
- obtains q r where "eucl_rel_nat m n (q, r)"
-proof (cases "n = 0")
- case True
- with that eucl_rel_nat_by0 show thesis
- by blast
-next
- case False
- have "\<exists>q r. m = q * n + r \<and> r < n"
- proof (induct m)
- case 0 with \<open>n \<noteq> 0\<close>
- have "(0::nat) = 0 * n + 0 \<and> 0 < n" by simp
- then show ?case by blast
- next
- case (Suc m) then obtain q' r'
- where m: "m = q' * n + r'" and n: "r' < n" by auto
- then show ?case proof (cases "Suc r' < n")
- case True
- from m n have "Suc m = q' * n + Suc r'" by simp
- with True show ?thesis by blast
- next
- case False then have "n \<le> Suc r'"
- by (simp add: not_less)
- moreover from n have "Suc r' \<le> n"
- by (simp add: Suc_le_eq)
- ultimately have "n = Suc r'" by auto
- with m have "Suc m = Suc q' * n + 0" by simp
- with \<open>n \<noteq> 0\<close> show ?thesis by blast
- qed
- qed
- with that \<open>n \<noteq> 0\<close> eucl_rel_natI show thesis
- by blast
-qed
-
-text \<open>@{const eucl_rel_nat} is injective:\<close>
-
-qualified lemma eucl_rel_nat_unique_div:
- assumes "eucl_rel_nat m n (q, r)"
- and "eucl_rel_nat m n (q', r')"
- shows "q = q'"
-proof (cases "n = 0")
- case True with assms show ?thesis
- by (auto elim: eucl_rel_nat.cases)
-next
- case False
- have *: "q' \<le> q" if "q' * n + r' = q * n + r" "r < n" for q r q' r' :: nat
- proof (rule ccontr)
- assume "\<not> q' \<le> q"
- then have "q < q'"
- by (simp add: not_le)
- with that show False
- by (auto simp add: less_iff_Suc_add algebra_simps)
- qed
- from \<open>n \<noteq> 0\<close> assms show ?thesis
- by (auto intro: order_antisym elim: eucl_rel_nat.cases dest: * sym split: if_splits)
-qed
-
-qualified lemma eucl_rel_nat_unique_mod:
- assumes "eucl_rel_nat m n (q, r)"
- and "eucl_rel_nat m n (q', r')"
- shows "r = r'"
-proof -
- from assms have "q' = q"
- by (auto intro: eucl_rel_nat_unique_div)
- with assms show ?thesis
- by (auto elim!: eucl_rel_nat.cases)
-qed
-
-text \<open>
- We instantiate divisibility on the natural numbers by
- means of @{const eucl_rel_nat}:
-\<close>
-
-qualified definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
- "divmod_nat m n = (THE qr. eucl_rel_nat m n qr)"
-
-qualified lemma eucl_rel_nat_divmod_nat:
- "eucl_rel_nat m n (divmod_nat m n)"
-proof -
- from eucl_rel_nat_ex
- obtain q r where rel: "eucl_rel_nat m n (q, r)" .
- then show ?thesis
- by (auto simp add: divmod_nat_def intro: theI
- elim: eucl_rel_nat_unique_div eucl_rel_nat_unique_mod)
-qed
-
-qualified lemma divmod_nat_unique:
- "divmod_nat m n = (q, r)" if "eucl_rel_nat m n (q, r)"
- using that
- by (auto simp add: divmod_nat_def intro: eucl_rel_nat_divmod_nat elim: eucl_rel_nat_unique_div eucl_rel_nat_unique_mod)
-
-qualified lemma divmod_nat_zero:
- "divmod_nat m 0 = (0, m)"
- by (rule divmod_nat_unique) (fact eucl_rel_nat_by0)
-
-qualified lemma divmod_nat_zero_left:
- "divmod_nat 0 n = (0, 0)"
- by (rule divmod_nat_unique)
- (cases n, auto intro: eucl_rel_nat_by0 eucl_rel_natI)
-
-qualified lemma divmod_nat_base:
- "m < n \<Longrightarrow> divmod_nat m n = (0, m)"
- by (rule divmod_nat_unique)
- (cases n, auto intro: eucl_rel_nat_by0 eucl_rel_natI)
-
-qualified lemma divmod_nat_step:
- assumes "0 < n" and "n \<le> m"
- shows "divmod_nat m n =
- (Suc (fst (divmod_nat (m - n) n)), snd (divmod_nat (m - n) n))"
-proof (rule divmod_nat_unique)
- have "eucl_rel_nat (m - n) n (divmod_nat (m - n) n)"
- by (fact eucl_rel_nat_divmod_nat)
- then show "eucl_rel_nat m n (Suc
- (fst (divmod_nat (m - n) n)), snd (divmod_nat (m - n) n))"
- using assms
- by (auto split: if_splits intro: eucl_rel_natI elim!: eucl_rel_nat.cases simp add: algebra_simps)
-qed
-
-end
-
-instantiation nat :: "{semidom_modulo, normalization_semidom}"
-begin
-
-definition normalize_nat :: "nat \<Rightarrow> nat"
- where [simp]: "normalize = (id :: nat \<Rightarrow> nat)"
-
-definition unit_factor_nat :: "nat \<Rightarrow> nat"
- where "unit_factor n = (if n = 0 then 0 else 1 :: nat)"
-
-lemma unit_factor_simps [simp]:
- "unit_factor 0 = (0::nat)"
- "unit_factor (Suc n) = 1"
- by (simp_all add: unit_factor_nat_def)
-
-definition divide_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
- where div_nat_def: "m div n = fst (Divides.divmod_nat m n)"
-
-definition modulo_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
- where mod_nat_def: "m mod n = snd (Divides.divmod_nat m n)"
-
-lemma fst_divmod_nat [simp]:
- "fst (Divides.divmod_nat m n) = m div n"
- by (simp add: div_nat_def)
-
-lemma snd_divmod_nat [simp]:
- "snd (Divides.divmod_nat m n) = m mod n"
- by (simp add: mod_nat_def)
-
-lemma divmod_nat_div_mod:
- "Divides.divmod_nat m n = (m div n, m mod n)"
- by (simp add: prod_eq_iff)
-
-lemma div_nat_unique:
- assumes "eucl_rel_nat m n (q, r)"
- shows "m div n = q"
- using assms
- by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
-
-lemma mod_nat_unique:
- assumes "eucl_rel_nat m n (q, r)"
- shows "m mod n = r"
- using assms
- by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
-
-lemma eucl_rel_nat: "eucl_rel_nat m n (m div n, m mod n)"
- using Divides.eucl_rel_nat_divmod_nat
- by (simp add: divmod_nat_div_mod)
-
-text \<open>The ''recursion'' equations for @{const divide} and @{const modulo}\<close>
-
-lemma div_less [simp]:
- fixes m n :: nat
- assumes "m < n"
- shows "m div n = 0"
- using assms Divides.divmod_nat_base by (simp add: prod_eq_iff)
-
-lemma le_div_geq:
- fixes m n :: nat
- assumes "0 < n" and "n \<le> m"
- shows "m div n = Suc ((m - n) div n)"
- using assms Divides.divmod_nat_step by (simp add: prod_eq_iff)
-
-lemma mod_less [simp]:
- fixes m n :: nat
- assumes "m < n"
- shows "m mod n = m"
- using assms Divides.divmod_nat_base by (simp add: prod_eq_iff)
-
-lemma le_mod_geq:
- fixes m n :: nat
- assumes "n \<le> m"
- shows "m mod n = (m - n) mod n"
- using assms Divides.divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff)
-
-lemma mod_less_divisor [simp]:
- fixes m n :: nat
- assumes "n > 0"
- shows "m mod n < n"
- using assms eucl_rel_nat [of m n]
- by (auto elim: eucl_rel_nat.cases)
-
-lemma mod_le_divisor [simp]:
- fixes m n :: nat
- assumes "n > 0"
- shows "m mod n \<le> n"
- using assms eucl_rel_nat [of m n]
- by (auto elim: eucl_rel_nat.cases)
-
-instance proof
- fix m n :: nat
- show "m div n * n + m mod n = m"
- using eucl_rel_nat [of m n]
- by (auto elim: eucl_rel_nat.cases)
-next
- fix n :: nat show "n div 0 = 0"
- by (simp add: div_nat_def Divides.divmod_nat_zero)
-next
- fix m n :: nat
- assume "n \<noteq> 0"
- then show "m * n div n = m"
- by (auto intro!: eucl_rel_natI div_nat_unique [of _ _ _ 0])
-qed (simp_all add: unit_factor_nat_def)
-
-end
-
-instance nat :: semiring_div
-proof
- fix m n q :: nat
- assume "n \<noteq> 0"
- then show "(q + m * n) div n = m + q div n"
- by (induct m) (simp_all add: le_div_geq)
-next
- fix m n q :: nat
- assume "m \<noteq> 0"
- show "(m * n) div (m * q) = n div q"
- proof (cases "q = 0")
- case True
- then show ?thesis
- by simp
- next
- case False
- show ?thesis
- proof (rule div_nat_unique [of _ _ _ "m * (n mod q)"])
- show "eucl_rel_nat (m * n) (m * q) (n div q, m * (n mod q))"
- by (rule eucl_rel_natI)
- (use \<open>m \<noteq> 0\<close> \<open>q \<noteq> 0\<close> div_mult_mod_eq [of n q] in \<open>auto simp add: algebra_simps distrib_left [symmetric]\<close>)
- qed
- qed
-qed
-
-lemma div_by_Suc_0 [simp]:
- "m div Suc 0 = m"
- using div_by_1 [of m] by simp
-
-lemma mod_by_Suc_0 [simp]:
- "m mod Suc 0 = 0"
- using mod_by_1 [of m] by simp
-
-lemma mod_greater_zero_iff_not_dvd:
- fixes m n :: nat
- shows "m mod n > 0 \<longleftrightarrow> \<not> n dvd m"
- by (simp add: dvd_eq_mod_eq_0)
-
-instantiation nat :: unique_euclidean_semiring
-begin
-
-definition [simp]:
- "euclidean_size_nat = (id :: nat \<Rightarrow> nat)"
-
-definition [simp]:
- "uniqueness_constraint_nat = (top :: nat \<Rightarrow> nat \<Rightarrow> bool)"
-
-instance
- by standard (use mult_le_mono2 [of 1] in \<open>simp_all add: unit_factor_nat_def mod_greater_zero_iff_not_dvd\<close>)
-
-end
-
-text \<open>Simproc for cancelling @{const divide} and @{const modulo}\<close>
-
-lemma (in semiring_modulo) cancel_div_mod_rules:
- "((a div b) * b + a mod b) + c = a + c"
- "(b * (a div b) + a mod b) + c = a + c"
- by (simp_all add: div_mult_mod_eq mult_div_mod_eq)
-
-ML_file "~~/src/Provers/Arith/cancel_div_mod.ML"
-
-ML \<open>
-structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
-(
- val div_name = @{const_name divide};
- val mod_name = @{const_name modulo};
- val mk_binop = HOLogic.mk_binop;
- val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
- val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
- fun mk_sum [] = HOLogic.zero
- | mk_sum [t] = t
- | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
- fun dest_sum tm =
- if HOLogic.is_zero tm then []
- else
- (case try HOLogic.dest_Suc tm of
- SOME t => HOLogic.Suc_zero :: dest_sum t
- | NONE =>
- (case try dest_plus tm of
- SOME (t, u) => dest_sum t @ dest_sum u
- | NONE => [tm]));
-
- val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
-
- val prove_eq_sums = Arith_Data.prove_conv2 all_tac
- (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps})
-)
-\<close>
-
-simproc_setup cancel_div_mod_nat ("(m::nat) + n") =
- \<open>K Cancel_Div_Mod_Nat.proc\<close>
-
-lemma divmod_nat_if [code]:
- "Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
- let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))"
- by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
-
-lemma mod_Suc_eq [mod_simps]:
- "Suc (m mod n) mod n = Suc m mod n"
-proof -
- have "(m mod n + 1) mod n = (m + 1) mod n"
- by (simp only: mod_simps)
- then show ?thesis
- by simp
-qed
-
-lemma mod_Suc_Suc_eq [mod_simps]:
- "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n"
-proof -
- have "(m mod n + 2) mod n = (m + 2) mod n"
- by (simp only: mod_simps)
- then show ?thesis
- by simp
-qed
+hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq
-subsubsection \<open>Quotient\<close>
-
-lemma div_geq: "0 < n \<Longrightarrow> \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)"
-by (simp add: le_div_geq linorder_not_less)
-
-lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))"
-by (simp add: div_geq)
-
-lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)"
-by simp
-
-lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)"
-by simp
-
-lemma div_positive:
- fixes m n :: nat
- assumes "n > 0"
- assumes "m \<ge> n"
- shows "m div n > 0"
-proof -
- from \<open>m \<ge> n\<close> obtain q where "m = n + q"
- by (auto simp add: le_iff_add)
- with \<open>n > 0\<close> show ?thesis by (simp add: div_add_self1)
-qed
-
-lemma div_eq_0_iff: "(a div b::nat) = 0 \<longleftrightarrow> a < b \<or> b = 0"
- by auto (metis div_positive less_numeral_extra(3) not_less)
-
-
-subsubsection \<open>Remainder\<close>
-
-lemma mod_Suc_le_divisor [simp]:
- "m mod Suc n \<le> n"
- using mod_less_divisor [of "Suc n" m] by arith
-
-lemma mod_less_eq_dividend [simp]:
- fixes m n :: nat
- shows "m mod n \<le> m"
-proof (rule add_leD2)
- from div_mult_mod_eq have "m div n * n + m mod n = m" .
- then show "m div n * n + m mod n \<le> m" by auto
-qed
-
-lemma mod_geq: "\<not> m < (n::nat) \<Longrightarrow> m mod n = (m - n) mod n"
-by (simp add: le_mod_geq linorder_not_less)
-
-lemma mod_if: "m mod (n::nat) = (if m < n then m else (m - n) mod n)"
-by (simp add: le_mod_geq)
+subsection \<open>More on division\<close>
-
-subsubsection \<open>Quotient and Remainder\<close>
-
-lemma div_mult1_eq:
- "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)"
- by (cases "c = 0")
- (auto simp add: algebra_simps distrib_left [symmetric]
- intro!: div_nat_unique [of _ _ _ "(a * (b mod c)) mod c"] eucl_rel_natI)
-
-lemma eucl_rel_nat_add1_eq:
- "eucl_rel_nat a c (aq, ar) \<Longrightarrow> eucl_rel_nat b c (bq, br)
- \<Longrightarrow> eucl_rel_nat (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)"
- by (auto simp add: split_ifs algebra_simps elim!: eucl_rel_nat.cases intro: eucl_rel_nat_by0 eucl_rel_natI)
-
-(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
-lemma div_add1_eq:
- "(a + b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)"
-by (blast intro: eucl_rel_nat_add1_eq [THEN div_nat_unique] eucl_rel_nat)
-
-lemma eucl_rel_nat_mult2_eq:
- assumes "eucl_rel_nat a b (q, r)"
- shows "eucl_rel_nat a (b * c) (q div c, b *(q mod c) + r)"
-proof (cases "c = 0")
- case True
- with assms show ?thesis
- by (auto intro: eucl_rel_nat_by0 elim!: eucl_rel_nat.cases simp add: ac_simps)
-next
- case False
- { assume "r < b"
- with False have "b * (q mod c) + r < b * c"
- apply (cut_tac m = q and n = c in mod_less_divisor)
- apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto)
- apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst)
- apply (simp add: add_mult_distrib2)
- done
- then have "r + b * (q mod c) < b * c"
- by (simp add: ac_simps)
- } with assms False show ?thesis
- by (auto simp add: algebra_simps add_mult_distrib2 [symmetric] elim!: eucl_rel_nat.cases intro: eucl_rel_nat.intros)
-qed
-
-lemma div_mult2_eq: "a div (b * c) = (a div b) div (c::nat)"
-by (force simp add: eucl_rel_nat [THEN eucl_rel_nat_mult2_eq, THEN div_nat_unique])
-
-lemma mod_mult2_eq: "a mod (b * c) = b * (a div b mod c) + a mod (b::nat)"
-by (auto simp add: mult.commute eucl_rel_nat [THEN eucl_rel_nat_mult2_eq, THEN mod_nat_unique])
-
-instantiation nat :: semiring_numeral_div
+instantiation nat :: unique_euclidean_semiring_numeral
begin
definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"
@@ -1276,420 +288,12 @@
in if r \<ge> numeral l then (2 * q + 1, r - numeral l)
else (2 * q, r))"
-instance
- by standard (auto intro: div_positive simp add: divmod'_nat_def divmod_step_nat_def mod_mult2_eq div_mult2_eq)
+instance by standard
+ (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq)
end
declare divmod_algorithm_code [where ?'a = nat, code]
-
-
-subsubsection \<open>Further Facts about Quotient and Remainder\<close>
-
-lemma div_le_mono:
- fixes m n k :: nat
- assumes "m \<le> n"
- shows "m div k \<le> n div k"
-proof -
- from assms obtain q where "n = m + q"
- by (auto simp add: le_iff_add)
- then show ?thesis
- by (simp add: div_add1_eq [of m q k])
-qed
-
-(* Antimonotonicity of div in second argument *)
-lemma div_le_mono2: "!!m::nat. [| 0<m; m\<le>n |] ==> (k div n) \<le> (k div m)"
-apply (subgoal_tac "0<n")
- prefer 2 apply simp
-apply (induct_tac k rule: nat_less_induct)
-apply (rename_tac "k")
-apply (case_tac "k<n", simp)
-apply (subgoal_tac "~ (k<m) ")
- prefer 2 apply simp
-apply (simp add: div_geq)
-apply (subgoal_tac "(k-n) div n \<le> (k-m) div n")
- prefer 2
- apply (blast intro: div_le_mono diff_le_mono2)
-apply (rule le_trans, simp)
-apply (simp)
-done
-
-lemma div_le_dividend [simp]: "m div n \<le> (m::nat)"
-apply (case_tac "n=0", simp)
-apply (subgoal_tac "m div n \<le> m div 1", simp)
-apply (rule div_le_mono2)
-apply (simp_all (no_asm_simp))
-done
-
-(* Similar for "less than" *)
-lemma div_less_dividend [simp]:
- "\<lbrakk>(1::nat) < n; 0 < m\<rbrakk> \<Longrightarrow> m div n < m"
-apply (induct m rule: nat_less_induct)
-apply (rename_tac "m")
-apply (case_tac "m<n", simp)
-apply (subgoal_tac "0<n")
- prefer 2 apply simp
-apply (simp add: div_geq)
-apply (case_tac "n<m")
- apply (subgoal_tac "(m-n) div n < (m-n) ")
- apply (rule impI less_trans_Suc)+
-apply assumption
- apply (simp_all)
-done
-
-text\<open>A fact for the mutilated chess board\<close>
-lemma mod_Suc: "Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))"
-apply (case_tac "n=0", simp)
-apply (induct "m" rule: nat_less_induct)
-apply (case_tac "Suc (na) <n")
-(* case Suc(na) < n *)
-apply (frule lessI [THEN less_trans], simp add: less_not_refl3)
-(* case n \<le> Suc(na) *)
-apply (simp add: linorder_not_less le_Suc_eq mod_geq)
-apply (auto simp add: Suc_diff_le le_mod_geq)
-done
-
-lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)"
-by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
-
-lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1]
-
-(*Loses information, namely we also have r<d provided d is nonzero*)
-lemma mod_eqD:
- fixes m d r q :: nat
- assumes "m mod d = r"
- shows "\<exists>q. m = r + q * d"
-proof -
- from div_mult_mod_eq obtain q where "q * d + m mod d = m" by blast
- with assms have "m = r + q * d" by simp
- then show ?thesis ..
-qed
-
-lemma split_div:
- "P(n div k :: nat) =
- ((k = 0 \<longrightarrow> P 0) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P i)))"
- (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))")
-proof
- assume P: ?P
- show ?Q
- proof (cases)
- assume "k = 0"
- with P show ?Q by simp
- next
- assume not0: "k \<noteq> 0"
- thus ?Q
- proof (simp, intro allI impI)
- fix i j
- assume n: "n = k*i + j" and j: "j < k"
- show "P i"
- proof (cases)
- assume "i = 0"
- with n j P show "P i" by simp
- next
- assume "i \<noteq> 0"
- with not0 n j P show "P i" by(simp add:ac_simps)
- qed
- qed
- qed
-next
- assume Q: ?Q
- show ?P
- proof (cases)
- assume "k = 0"
- with Q show ?P by simp
- next
- assume not0: "k \<noteq> 0"
- with Q have R: ?R by simp
- from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"]
- show ?P by simp
- qed
-qed
-
-lemma split_div_lemma:
- assumes "0 < n"
- shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m::nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume ?rhs
- with minus_mod_eq_mult_div [symmetric] have nq: "n * q = m - (m mod n)" by simp
- then have A: "n * q \<le> m" by simp
- have "n - (m mod n) > 0" using mod_less_divisor assms by auto
- then have "m < m + (n - (m mod n))" by simp
- then have "m < n + (m - (m mod n))" by simp
- with nq have "m < n + n * q" by simp
- then have B: "m < n * Suc q" by simp
- from A B show ?lhs ..
-next
- assume P: ?lhs
- then have "eucl_rel_nat m n (q, m - n * q)"
- by (auto intro: eucl_rel_natI simp add: ac_simps)
- then have "m div n = q"
- by (rule div_nat_unique)
- then show ?rhs by simp
-qed
-
-theorem split_div':
- "P ((m::nat) div n) = ((n = 0 \<and> P 0) \<or>
- (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))"
- apply (cases "0 < n")
- apply (simp only: add: split_div_lemma)
- apply simp_all
- done
-
-lemma split_mod:
- "P(n mod k :: nat) =
- ((k = 0 \<longrightarrow> P n) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P j)))"
- (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))")
-proof
- assume P: ?P
- show ?Q
- proof (cases)
- assume "k = 0"
- with P show ?Q by simp
- next
- assume not0: "k \<noteq> 0"
- thus ?Q
- proof (simp, intro allI impI)
- fix i j
- assume "n = k*i + j" "j < k"
- thus "P j" using not0 P by (simp add: ac_simps)
- qed
- qed
-next
- assume Q: ?Q
- show ?P
- proof (cases)
- assume "k = 0"
- with Q show ?P by simp
- next
- assume not0: "k \<noteq> 0"
- with Q have R: ?R by simp
- from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"]
- show ?P by simp
- qed
-qed
-
-lemma div_eq_dividend_iff: "a \<noteq> 0 \<Longrightarrow> (a :: nat) div b = a \<longleftrightarrow> b = 1"
- apply rule
- apply (cases "b = 0")
- apply simp_all
- apply (metis (full_types) One_nat_def Suc_lessI div_less_dividend less_not_refl3)
- done
-
-lemma (in field_char_0) of_nat_div:
- "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)"
-proof -
- have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)"
- unfolding of_nat_add by (cases "n = 0") simp_all
- then show ?thesis
- by simp
-qed
-
-
-subsubsection \<open>An ``induction'' law for modulus arithmetic.\<close>
-
-lemma mod_induct_0:
- assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)"
- and base: "P i" and i: "i<p"
- shows "P 0"
-proof (rule ccontr)
- assume contra: "\<not>(P 0)"
- from i have p: "0<p" by simp
- have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k")
- proof
- fix k
- show "?A k"
- proof (induct k)
- show "?A 0" by simp \<comment> "by contradiction"
- next
- fix n
- assume ih: "?A n"
- show "?A (Suc n)"
- proof (clarsimp)
- assume y: "P (p - Suc n)"
- have n: "Suc n < p"
- proof (rule ccontr)
- assume "\<not>(Suc n < p)"
- hence "p - Suc n = 0"
- by simp
- with y contra show "False"
- by simp
- qed
- hence n2: "Suc (p - Suc n) = p-n" by arith
- from p have "p - Suc n < p" by arith
- with y step have z: "P ((Suc (p - Suc n)) mod p)"
- by blast
- show "False"
- proof (cases "n=0")
- case True
- with z n2 contra show ?thesis by simp
- next
- case False
- with p have "p-n < p" by arith
- with z n2 False ih show ?thesis by simp
- qed
- qed
- qed
- qed
- moreover
- from i obtain k where "0<k \<and> i+k=p"
- by (blast dest: less_imp_add_positive)
- hence "0<k \<and> i=p-k" by auto
- moreover
- note base
- ultimately
- show "False" by blast
-qed
-
-lemma mod_induct:
- assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)"
- and base: "P i" and i: "i<p" and j: "j<p"
- shows "P j"
-proof -
- have "\<forall>j<p. P j"
- proof
- fix j
- show "j<p \<longrightarrow> P j" (is "?A j")
- proof (induct j)
- from step base i show "?A 0"
- by (auto elim: mod_induct_0)
- next
- fix k
- assume ih: "?A k"
- show "?A (Suc k)"
- proof
- assume suc: "Suc k < p"
- hence k: "k<p" by simp
- with ih have "P k" ..
- with step k have "P (Suc k mod p)"
- by blast
- moreover
- from suc have "Suc k mod p = Suc k"
- by simp
- ultimately
- show "P (Suc k)" by simp
- qed
- qed
- qed
- with j show ?thesis by blast
-qed
-
-lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)"
- by (simp add: numeral_2_eq_2 le_div_geq)
-
-lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2"
- by (simp add: numeral_2_eq_2 le_mod_geq)
-
-lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
-by (simp add: mult_2 [symmetric])
-
-lemma mod2_gr_0 [simp]: "0 < (m::nat) mod 2 \<longleftrightarrow> m mod 2 = 1"
-proof -
- { fix n :: nat have "(n::nat) < 2 \<Longrightarrow> n = 0 \<or> n = 1" by (cases n) simp_all }
- moreover have "m mod 2 < 2" by simp
- ultimately have "m mod 2 = 0 \<or> m mod 2 = 1" .
- then show ?thesis by auto
-qed
-
-text\<open>These lemmas collapse some needless occurrences of Suc:
- at least three Sucs, since two and fewer are rewritten back to Suc again!
- We already have some rules to simplify operands smaller than 3.\<close>
-
-lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
-by (simp add: Suc3_eq_add_3)
-
-lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
-by (simp add: Suc3_eq_add_3)
-
-lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v
-lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v
-
-lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1"
-apply (induct "m")
-apply (simp_all add: mod_Suc)
-done
-
-declare Suc_times_mod_eq [of "numeral w", simp] for w
-
-lemma Suc_div_le_mono [simp]: "n div k \<le> (Suc n) div k"
-by (simp add: div_le_mono)
-
-lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2"
-by (cases n) simp_all
-
-lemma div_2_gt_zero [simp]: assumes A: "(1::nat) < n" shows "0 < n div 2"
-proof -
- from A have B: "0 < n - 1" and C: "n - 1 + 1 = n" by simp_all
- from Suc_n_div_2_gt_zero [OF B] C show ?thesis by simp
-qed
-
-lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n"
-proof -
- have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp
- also have "... = Suc m mod n" by (rule mod_mult_self3)
- finally show ?thesis .
-qed
-
-lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n"
-apply (subst mod_Suc [of m])
-apply (subst mod_Suc [of "m mod n"], simp)
-done
-
-lemma mod_2_not_eq_zero_eq_one_nat:
- fixes n :: nat
- shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1"
- by (fact not_mod_2_eq_0_eq_1)
-
-lemma even_Suc_div_two [simp]:
- "even n \<Longrightarrow> Suc n div 2 = n div 2"
- using even_succ_div_two [of n] by simp
-
-lemma odd_Suc_div_two [simp]:
- "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)"
- using odd_succ_div_two [of n] by simp
-
-lemma odd_two_times_div_two_nat [simp]:
- assumes "odd n"
- shows "2 * (n div 2) = n - (1 :: nat)"
-proof -
- from assms have "2 * (n div 2) + 1 = n"
- by (rule odd_two_times_div_two_succ)
- then have "Suc (2 * (n div 2)) - 1 = n - 1"
- by simp
- then show ?thesis
- by simp
-qed
-
-lemma parity_induct [case_names zero even odd]:
- assumes zero: "P 0"
- assumes even: "\<And>n. P n \<Longrightarrow> P (2 * n)"
- assumes odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))"
- shows "P n"
-proof (induct n rule: less_induct)
- case (less n)
- show "P n"
- proof (cases "n = 0")
- case True with zero show ?thesis by simp
- next
- case False
- with less have hyp: "P (n div 2)" by simp
- show ?thesis
- proof (cases "even n")
- case True
- with hyp even [of "n div 2"] show ?thesis
- by simp
- next
- case False
- with hyp odd [of "n div 2"] show ?thesis
- by simp
- qed
- qed
-qed
lemma Suc_0_div_numeral [simp]:
fixes k l :: num
@@ -1701,11 +305,26 @@
shows "Suc 0 mod numeral k = snd (divmod Num.One k)"
by (simp_all add: snd_divmod)
+definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat"
+ where "divmod_nat m n = (m div n, m mod n)"
-subsection \<open>Division on @{typ int}\<close>
+lemma fst_divmod_nat [simp]:
+ "fst (divmod_nat m n) = m div n"
+ by (simp add: divmod_nat_def)
-context
-begin
+lemma snd_divmod_nat [simp]:
+ "snd (divmod_nat m n) = m mod n"
+ by (simp add: divmod_nat_def)
+
+lemma divmod_nat_if [code]:
+ "Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else
+ let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))"
+ by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq)
+
+lemma [code]:
+ "m div n = fst (divmod_nat m n)"
+ "m mod n = snd (divmod_nat m n)"
+ by simp_all
inductive eucl_rel_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool"
where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)"
@@ -1752,31 +371,6 @@
apply (blast intro: unique_quotient)
done
-end
-
-instantiation int :: "{idom_modulo, normalization_semidom}"
-begin
-
-definition normalize_int :: "int \<Rightarrow> int"
- where [simp]: "normalize = (abs :: int \<Rightarrow> int)"
-
-definition unit_factor_int :: "int \<Rightarrow> int"
- where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)"
-
-definition divide_int :: "int \<Rightarrow> int \<Rightarrow> int"
- where "k div l = (if l = 0 \<or> k = 0 then 0
- else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0
- then int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
- else
- if l dvd k then - int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
- else - int (Suc (nat \<bar>k\<bar> div nat \<bar>l\<bar>)))"
-
-definition modulo_int :: "int \<Rightarrow> int \<Rightarrow> int"
- where "k mod l = (if l = 0 then k else if l dvd k then 0
- else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0
- then sgn l * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)
- else sgn l * (\<bar>l\<bar> - int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)))"
-
lemma eucl_rel_int:
"eucl_rel_int k l (k div l, k mod l)"
proof (cases k rule: int_cases3)
@@ -1808,42 +402,21 @@
using unique_quotient [of k l] unique_remainder [of k l]
by auto
-instance proof
- fix k l :: int
- show "k div l * l + k mod l = k"
- using eucl_rel_int [of k l]
- unfolding eucl_rel_int_iff by (simp add: ac_simps)
-next
- fix k :: int show "k div 0 = 0"
- by (rule div_int_unique, simp add: eucl_rel_int_iff)
-next
- fix k l :: int
- assume "l \<noteq> 0"
- then show "k * l div l = k"
- by (auto simp add: eucl_rel_int_iff ac_simps intro: div_int_unique [of _ _ _ 0])
-qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff')
-
-end
-
-lemma is_unit_int:
- "is_unit (k::int) \<longleftrightarrow> k = 1 \<or> k = - 1"
- by auto
-
-lemma zdiv_int:
- "int (a div b) = int a div int b"
- by (simp add: divide_int_def)
-
-lemma zmod_int:
- "int (a mod b) = int a mod int b"
- by (simp add: modulo_int_def int_dvd_iff)
-
lemma div_abs_eq_div_nat:
"\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)"
by (simp add: divide_int_def)
lemma mod_abs_eq_div_nat:
"\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)"
- by (simp add: modulo_int_def dvd_int_unfold_dvd_nat)
+ by (simp add: modulo_int_def)
+
+lemma zdiv_int:
+ "int (a div b) = int a div int b"
+ by (simp add: divide_int_def sgn_1_pos)
+
+lemma zmod_int:
+ "int (a mod b) = int a mod int b"
+ by (simp add: modulo_int_def sgn_1_pos)
lemma div_sgn_abs_cancel:
fixes k l v :: int
@@ -1868,7 +441,7 @@
next
case False
with assms have "(sgn k * \<bar>k\<bar>) div (sgn l * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"
- by (simp add: div_sgn_abs_cancel)
+ using div_sgn_abs_cancel [of l k l] by simp
then show ?thesis
by (simp add: sgn_mult_abs)
qed
@@ -1877,12 +450,14 @@
fixes k l :: int
assumes "l dvd k"
shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)"
-proof (cases "k = 0")
+proof (cases "k = 0 \<or> l = 0")
case True
then show ?thesis
- by simp
+ by auto
next
case False
+ then have "k \<noteq> 0" and "l \<noteq> 0"
+ by auto
show ?thesis
proof (cases "sgn l = sgn k")
case True
@@ -1890,9 +465,12 @@
by (simp add: div_eq_sgn_abs)
next
case False
- with \<open>k \<noteq> 0\<close> assms show ?thesis
+ with \<open>k \<noteq> 0\<close> \<open>l \<noteq> 0\<close>
+ have "sgn l * sgn k = - 1"
+ by (simp add: sgn_if split: if_splits)
+ with assms show ?thesis
unfolding divide_int_def [of k l]
- by (auto simp add: zdiv_int)
+ by (auto simp add: zdiv_int ac_simps)
qed
qed
@@ -1904,81 +482,14 @@
using assms
by (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int)
-lemma sgn_mod:
- fixes k l :: int
- assumes "l \<noteq> 0" "\<not> l dvd k"
- shows "sgn (k mod l) = sgn l"
-proof -
- from \<open>\<not> l dvd k\<close>
- have "k mod l \<noteq> 0"
- by (simp add: dvd_eq_mod_eq_0)
- show ?thesis
- using \<open>l \<noteq> 0\<close> \<open>\<not> l dvd k\<close>
- unfolding modulo_int_def [of k l]
- by (auto simp add: sgn_1_pos sgn_1_neg mod_greater_zero_iff_not_dvd nat_dvd_iff not_less
- zless_nat_eq_int_zless [symmetric] elim: nonpos_int_cases)
-qed
-
-lemma abs_mod_less:
- fixes k l :: int
- assumes "l \<noteq> 0"
- shows "\<bar>k mod l\<bar> < \<bar>l\<bar>"
- using assms unfolding modulo_int_def [of k l]
- by (auto simp add: not_less int_dvd_iff mod_greater_zero_iff_not_dvd elim: pos_int_cases neg_int_cases nonneg_int_cases nonpos_int_cases)
-
-instance int :: ring_div
-proof
- fix k l s :: int
- assume "l \<noteq> 0"
- then have "eucl_rel_int (k + s * l) l (s + k div l, k mod l)"
- using eucl_rel_int [of k l]
- unfolding eucl_rel_int_iff by (auto simp: algebra_simps)
- then show "(k + s * l) div l = s + k div l"
- by (rule div_int_unique)
-next
- fix k l s :: int
- assume "s \<noteq> 0"
- have "\<And>q r. eucl_rel_int k l (q, r)
- \<Longrightarrow> eucl_rel_int (s * k) (s * l) (q, s * r)"
- unfolding eucl_rel_int_iff
- by (rule linorder_cases [of 0 l])
- (use \<open>s \<noteq> 0\<close> in \<open>auto simp: algebra_simps
- mult_less_0_iff zero_less_mult_iff mult_strict_right_mono
- mult_strict_right_mono_neg zero_le_mult_iff mult_le_0_iff\<close>)
- then have "eucl_rel_int (s * k) (s * l) (k div l, s * (k mod l))"
- using eucl_rel_int [of k l] .
- then show "(s * k) div (s * l) = k div l"
- by (rule div_int_unique)
-qed
-
-ML \<open>
-structure Cancel_Div_Mod_Int = Cancel_Div_Mod
-(
- val div_name = @{const_name divide};
- val mod_name = @{const_name modulo};
- val mk_binop = HOLogic.mk_binop;
- val mk_sum = Arith_Data.mk_sum HOLogic.intT;
- val dest_sum = Arith_Data.dest_sum;
-
- val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
-
- val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
- @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps})
-)
-\<close>
-
-simproc_setup cancel_div_mod_int ("(k::int) + l") =
- \<open>K Cancel_Div_Mod_Int.proc\<close>
-
-
text\<open>Basic laws about division and remainder\<close>
lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b"
using eucl_rel_int [of a b]
by (auto simp add: eucl_rel_int_iff prod_eq_iff)
-lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1]
- and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2]
+lemmas pos_mod_sign = pos_mod_conj [THEN conjunct1]
+ and pos_mod_bound = pos_mod_conj [THEN conjunct2]
lemma neg_mod_conj: "b < (0::int) \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b"
using eucl_rel_int [of a b]
@@ -2005,6 +516,10 @@
apply (auto simp add: eucl_rel_int_iff)
done
+lemma div_positive_int:
+ "k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int
+ using that by (simp add: divide_int_def div_positive)
+
(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*)
lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a"
@@ -2057,14 +572,14 @@
by (simp add: mod_eq_0_iff_dvd)
lemma zdiv_zminus2_eq_if:
- "b \<noteq> (0::int)
+ "b \<noteq> (0::int)
==> a div (-b) =
(if a mod b = 0 then - (a div b) else - (a div b) - 1)"
-by (simp add: zdiv_zminus1_eq_if div_minus_right)
+ by (auto simp add: zdiv_zminus1_eq_if div_minus_right)
lemma zmod_zminus2_eq_if:
- "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)"
-by (simp add: zmod_zminus1_eq_if mod_minus_right)
+ "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)"
+ by (auto simp add: zmod_zminus1_eq_if mod_minus_right)
subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close>
@@ -2154,31 +669,13 @@
subsubsection \<open>More Algebraic Laws for div and mod\<close>
-text\<open>proving (a*b) div c = a * (b div c) + a * (b mod c)\<close>
-
-lemma zmult1_lemma:
- "[| eucl_rel_int b c (q, r) |]
- ==> eucl_rel_int (a * b) c (a*q + a*r div c, a*r mod c)"
-by (auto simp add: split_ifs eucl_rel_int_iff linorder_neq_iff distrib_left ac_simps)
-
lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
-apply (case_tac "c = 0", simp)
-apply (blast intro: eucl_rel_int [THEN zmult1_lemma, THEN div_int_unique])
-done
-
-text\<open>proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c)\<close>
-
-lemma zadd1_lemma:
- "[| eucl_rel_int a c (aq, ar); eucl_rel_int b c (bq, br) |]
- ==> eucl_rel_int (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)"
-by (force simp add: split_ifs eucl_rel_int_iff linorder_neq_iff distrib_left)
+ by (fact div_mult1_eq)
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
lemma zdiv_zadd1_eq:
"(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
-apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF eucl_rel_int eucl_rel_int] div_int_unique)
-done
+ by (fact div_add1_eq)
lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
@@ -2421,21 +918,6 @@
by simp
qed
-instantiation int :: unique_euclidean_ring
-begin
-
-definition [simp]:
- "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)"
-
-definition [simp]:
- "uniqueness_constraint_int (k :: int) l \<longleftrightarrow> unit_factor k = unit_factor l"
-
-instance
- by standard
- (use mult_le_mono2 [of 1] in \<open>auto simp add: abs_mult nat_mult_distrib sgn_mod zdiv_eq_0_iff sgn_1_pos sgn_mult split: abs_split\<close>)
-
-end
-
subsubsection \<open>Quotients of Signs\<close>
@@ -2507,7 +989,7 @@
subsubsection \<open>Computation of Division and Remainder\<close>
-instantiation int :: semiring_numeral_div
+instantiation int :: unique_euclidean_semiring_numeral
begin
definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"
@@ -2554,11 +1036,17 @@
lemma minus_numeral_mod_numeral [simp]:
"- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)"
-proof -
- have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
- using that by (simp only: snd_divmod modulo_int_def) auto
+proof (cases "snd (divmod m n) = (0::int)")
+ case True
then show ?thesis
- by (auto simp add: split_def Let_def adjust_div_def divides_aux_def modulo_int_def)
+ by (simp add: mod_eq_0_iff_dvd divides_aux_def)
+next
+ case False
+ then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
+ by (simp only: snd_divmod modulo_int_def) auto
+ then show ?thesis
+ by (simp add: divides_aux_def adjust_div_def)
+ (simp add: divides_aux_def modulo_int_def)
qed
lemma numeral_div_minus_numeral [simp]:
@@ -2572,11 +1060,17 @@
lemma numeral_mod_minus_numeral [simp]:
"numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)"
-proof -
- have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
- using that by (simp only: snd_divmod modulo_int_def) auto
+proof (cases "snd (divmod m n) = (0::int)")
+ case True
then show ?thesis
- by (auto simp add: split_def Let_def adjust_div_def divides_aux_def modulo_int_def)
+ by (simp add: mod_eq_0_iff_dvd divides_aux_def)
+next
+ case False
+ then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
+ by (simp only: snd_divmod modulo_int_def) auto
+ then show ?thesis
+ by (simp add: divides_aux_def adjust_div_def)
+ (simp add: divides_aux_def modulo_int_def)
qed
lemma minus_one_div_numeral [simp]:
@@ -2600,6 +1094,24 @@
subsubsection \<open>Further properties\<close>
+lemma div_int_pos_iff:
+ "k div l \<ge> 0 \<longleftrightarrow> k = 0 \<or> l = 0 \<or> k \<ge> 0 \<and> l \<ge> 0
+ \<or> k < 0 \<and> l < 0"
+ for k l :: int
+ apply (cases "k = 0 \<or> l = 0")
+ apply (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff)
+ apply (rule ccontr)
+ apply (simp add: neg_imp_zdiv_nonneg_iff)
+ done
+
+lemma mod_int_pos_iff:
+ "k mod l \<ge> 0 \<longleftrightarrow> l dvd k \<or> l = 0 \<and> k \<ge> 0 \<or> l > 0"
+ for k l :: int
+ apply (cases "l > 0")
+ apply (simp_all add: dvd_eq_mod_eq_0)
+ apply (use neg_mod_conj [of l k] in \<open>auto simp add: le_less not_less\<close>)
+ done
+
text \<open>Simplify expresions in which div and mod combine numerical constants\<close>
lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q"
@@ -2645,65 +1157,13 @@
apply (simp add: nat_eq_iff zmod_int)
done
-text \<open>transfer setup\<close>
-
-lemma transfer_nat_int_functions:
- "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
- "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
- by (auto simp add: nat_div_distrib nat_mod_distrib)
-
-lemma transfer_nat_int_function_closures:
- "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
- "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
- apply (cases "y = 0")
- apply (auto simp add: pos_imp_zdiv_nonneg_iff)
- apply (cases "y = 0")
- apply auto
-done
-
-declare transfer_morphism_nat_int [transfer add return:
- transfer_nat_int_functions
- transfer_nat_int_function_closures
-]
-
-lemma transfer_int_nat_functions:
- "(int x) div (int y) = int (x div y)"
- "(int x) mod (int y) = int (x mod y)"
- by (auto simp add: zdiv_int zmod_int)
-
-lemma transfer_int_nat_function_closures:
- "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
- "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
- by (simp_all only: is_nat_def transfer_nat_int_function_closures)
-
-declare transfer_morphism_int_nat [transfer add return:
- transfer_int_nat_functions
- transfer_int_nat_function_closures
-]
-
text\<open>Suggested by Matthias Daum\<close>
lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)"
apply (subgoal_tac "nat x div nat k < nat x")
apply (simp add: nat_div_distrib [symmetric])
-apply (rule Divides.div_less_dividend, simp_all)
+apply (rule div_less_dividend, simp_all)
done
-lemma (in ring_div) mod_eq_dvd_iff:
- "a mod c = b mod c \<longleftrightarrow> c dvd a - b" (is "?P \<longleftrightarrow> ?Q")
-proof
- assume ?P
- then have "(a mod c - b mod c) mod c = 0"
- by simp
- then show ?Q
- by (simp add: dvd_eq_mod_eq_0 mod_simps)
-next
- assume ?Q
- then obtain d where d: "a - b = c * d" ..
- then have "a = c * d + b"
- by (simp add: algebra_simps)
- then show ?P by simp
-qed
-
lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x"
shows "\<exists>q. x = y + n * q"
proof-
@@ -2734,32 +1194,33 @@
thus ?lhs by simp
qed
+
subsubsection \<open>Dedicated simproc for calculation\<close>
text \<open>
There is space for improvement here: the calculation itself
- could be carried outside the logic, and a generic simproc
+ could be carried out outside the logic, and a generic simproc
(simplifier setup) for generic calculation would be helpful.
\<close>
simproc_setup numeral_divmod
- ("0 div 0 :: 'a :: semiring_numeral_div" | "0 mod 0 :: 'a :: semiring_numeral_div" |
- "0 div 1 :: 'a :: semiring_numeral_div" | "0 mod 1 :: 'a :: semiring_numeral_div" |
+ ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
+ "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
"0 div - 1 :: int" | "0 mod - 1 :: int" |
- "0 div numeral b :: 'a :: semiring_numeral_div" | "0 mod numeral b :: 'a :: semiring_numeral_div" |
+ "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
"0 div - numeral b :: int" | "0 mod - numeral b :: int" |
- "1 div 0 :: 'a :: semiring_numeral_div" | "1 mod 0 :: 'a :: semiring_numeral_div" |
- "1 div 1 :: 'a :: semiring_numeral_div" | "1 mod 1 :: 'a :: semiring_numeral_div" |
+ "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
+ "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
"1 div - 1 :: int" | "1 mod - 1 :: int" |
- "1 div numeral b :: 'a :: semiring_numeral_div" | "1 mod numeral b :: 'a :: semiring_numeral_div" |
+ "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
"1 div - numeral b :: int" |"1 mod - numeral b :: int" |
"- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |
"- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |
"- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |
- "numeral a div 0 :: 'a :: semiring_numeral_div" | "numeral a mod 0 :: 'a :: semiring_numeral_div" |
- "numeral a div 1 :: 'a :: semiring_numeral_div" | "numeral a mod 1 :: 'a :: semiring_numeral_div" |
+ "numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
+ "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" |
"numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |
- "numeral a div numeral b :: 'a :: semiring_numeral_div" | "numeral a mod numeral b :: 'a :: semiring_numeral_div" |
+ "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" |
"numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |
"- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |
"- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |
@@ -2819,9 +1280,58 @@
code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
lemma dvd_eq_mod_eq_0_numeral:
- "numeral x dvd (numeral y :: 'a) \<longleftrightarrow> numeral y mod numeral x = (0 :: 'a::semiring_div)"
+ "numeral x dvd (numeral y :: 'a) \<longleftrightarrow> numeral y mod numeral x = (0 :: 'a::semidom_modulo)"
by (fact dvd_eq_mod_eq_0)
declare minus_div_mult_eq_mod [symmetric, nitpick_unfold]
+
+subsubsection \<open>Lemmas of doubtful value\<close>
+
+lemma mod_mult_self3':
+ "Suc (k * n + m) mod n = Suc m mod n"
+ by (fact Suc_mod_mult_self3)
+
+lemma mod_Suc_eq_Suc_mod:
+ "Suc m mod n = Suc (m mod n) mod n"
+ by (simp add: mod_simps)
+
+lemma div_geq:
+ "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat
+ by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>)
+
+lemma mod_geq:
+ "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat
+ by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>)
+
+lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)"
+ by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def)
+
+lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1]
+
+(*Loses information, namely we also have r<d provided d is nonzero*)
+lemma mod_eqD:
+ fixes m d r q :: nat
+ assumes "m mod d = r"
+ shows "\<exists>q. m = r + q * d"
+proof -
+ from div_mult_mod_eq obtain q where "q * d + m mod d = m" by blast
+ with assms have "m = r + q * d" by simp
+ then show ?thesis ..
+qed
+
+lemmas even_times_iff = even_mult_iff -- \<open>FIXME duplicate\<close>
+
+lemma mod_2_not_eq_zero_eq_one_nat:
+ fixes n :: nat
+ shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1"
+ by (fact not_mod_2_eq_0_eq_1)
+
+lemma even_int_iff [simp]: "even (int n) \<longleftrightarrow> even n"
+ by (fact even_of_nat)
+
+lemma is_unit_int:
+ "is_unit (k::int) \<longleftrightarrow> k = 1 \<or> k = - 1"
+ by auto
+
end
--- a/src/HOL/Enum.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Enum.thy Mon Oct 09 22:08:05 2017 +0200
@@ -683,7 +683,7 @@
instance finite_2 :: complete_linorder ..
-instantiation finite_2 :: "{field, idom_abs_sgn}" begin
+instantiation finite_2 :: "{field, idom_abs_sgn, idom_modulo}" begin
definition [simp]: "0 = a\<^sub>1"
definition [simp]: "1 = a\<^sub>2"
definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>1 | _ \<Rightarrow> a\<^sub>2)"
@@ -692,12 +692,15 @@
definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
definition "inverse = (\<lambda>x :: finite_2. x)"
definition "divide = (op * :: finite_2 \<Rightarrow> _)"
+definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
definition "abs = (\<lambda>x :: finite_2. x)"
definition "sgn = (\<lambda>x :: finite_2. x)"
instance
by standard
- (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def times_finite_2_def
- inverse_finite_2_def divide_finite_2_def abs_finite_2_def sgn_finite_2_def
+ (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def
+ times_finite_2_def
+ inverse_finite_2_def divide_finite_2_def modulo_finite_2_def
+ abs_finite_2_def sgn_finite_2_def
split: finite_2.splits)
end
@@ -709,14 +712,15 @@
"x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> y = a\<^sub>1"
by (auto simp add: dvd_def times_finite_2_def split: finite_2.splits)
-instantiation finite_2 :: "{ring_div, normalization_semidom}" begin
+instantiation finite_2 :: "{normalization_semidom, unique_euclidean_semiring}" begin
definition [simp]: "normalize = (id :: finite_2 \<Rightarrow> _)"
definition [simp]: "unit_factor = (id :: finite_2 \<Rightarrow> _)"
-definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
+definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | a\<^sub>2 \<Rightarrow> 1)"
+definition [simp]: "uniqueness_constraint = (\<top> :: finite_2 \<Rightarrow> _)"
instance
by standard
- (simp_all add: dvd_finite_2_unfold times_finite_2_def
- divide_finite_2_def modulo_finite_2_def split: finite_2.splits)
+ (auto simp add: divide_finite_2_def times_finite_2_def dvd_finite_2_unfold
+ split: finite_2.splits)
end
@@ -826,7 +830,7 @@
instance finite_3 :: complete_linorder ..
-instantiation finite_3 :: "{field, idom_abs_sgn}" begin
+instantiation finite_3 :: "{field, idom_abs_sgn, idom_modulo}" begin
definition [simp]: "0 = a\<^sub>1"
definition [simp]: "1 = a\<^sub>2"
definition
@@ -839,12 +843,15 @@
definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \<Rightarrow> a\<^sub>2 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>3 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
definition "inverse = (\<lambda>x :: finite_3. x)"
definition "x div y = x * inverse (y :: finite_3)"
+definition "x mod y = (case y of a\<^sub>1 \<Rightarrow> x | _ \<Rightarrow> a\<^sub>1)"
definition "abs = (\<lambda>x. case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
definition "sgn = (\<lambda>x :: finite_3. x)"
instance
by standard
- (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def times_finite_3_def
- inverse_finite_3_def divide_finite_3_def abs_finite_3_def sgn_finite_3_def
+ (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def
+ times_finite_3_def
+ inverse_finite_3_def divide_finite_3_def modulo_finite_3_def
+ abs_finite_3_def sgn_finite_3_def
less_finite_3_def
split: finite_3.splits)
end
@@ -857,20 +864,21 @@
"x dvd y \<longleftrightarrow> x = a\<^sub>2 \<or> x = a\<^sub>3 \<or> y = a\<^sub>1"
by (cases x) (auto simp add: dvd_def times_finite_3_def split: finite_3.splits)
-instantiation finite_3 :: "{ring_div, normalization_semidom}" begin
-definition "normalize x = (case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
+instantiation finite_3 :: "{normalization_semidom, unique_euclidean_semiring}" begin
+definition [simp]: "normalize x = (case x of a\<^sub>3 \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> x)"
definition [simp]: "unit_factor = (id :: finite_3 \<Rightarrow> _)"
-definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>1) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
-instance
- by standard
- (auto simp add: finite_3_not_eq_unfold plus_finite_3_def
- dvd_finite_3_unfold times_finite_3_def inverse_finite_3_def
- normalize_finite_3_def divide_finite_3_def modulo_finite_3_def
- split: finite_3.splits)
+definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \<Rightarrow> 0 | _ \<Rightarrow> 1)"
+definition [simp]: "uniqueness_constraint = (\<top> :: finite_3 \<Rightarrow> _)"
+instance proof
+ fix x :: finite_3
+ assume "x \<noteq> 0"
+ then show "is_unit (unit_factor x)"
+ by (cases x) (simp_all add: dvd_finite_3_unfold)
+qed (auto simp add: divide_finite_3_def times_finite_3_def
+ dvd_finite_3_unfold inverse_finite_3_def plus_finite_3_def
+ split: finite_3.splits)
end
-
-
hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3
datatype (plugins only: code "quickcheck" extraction) finite_4 =
--- a/src/HOL/Euclidean_Division.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Euclidean_Division.thy Mon Oct 09 22:08:05 2017 +0200
@@ -3,91 +3,15 @@
Author: Florian Haftmann, TU Muenchen
*)
-section \<open>Uniquely determined division in euclidean (semi)rings\<close>
+section \<open>Division in euclidean (semi)rings\<close>
theory Euclidean_Division
- imports Nat_Transfer
-begin
-
-subsection \<open>Quotient and remainder in integral domains\<close>
-
-class semidom_modulo = algebraic_semidom + semiring_modulo
+ imports Int Lattices_Big
begin
-lemma mod_0 [simp]: "0 mod a = 0"
- using div_mult_mod_eq [of 0 a] by simp
-
-lemma mod_by_0 [simp]: "a mod 0 = a"
- using div_mult_mod_eq [of a 0] by simp
-
-lemma mod_by_1 [simp]:
- "a mod 1 = 0"
-proof -
- from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp
- then have "a + a mod 1 = a + 0" by simp
- then show ?thesis by (rule add_left_imp_eq)
-qed
-
-lemma mod_self [simp]:
- "a mod a = 0"
- using div_mult_mod_eq [of a a] by simp
-
-lemma dvd_imp_mod_0 [simp]:
- assumes "a dvd b"
- shows "b mod a = 0"
- using assms minus_div_mult_eq_mod [of b a] by simp
-
-lemma mod_0_imp_dvd:
- assumes "a mod b = 0"
- shows "b dvd a"
-proof -
- have "b dvd ((a div b) * b)" by simp
- also have "(a div b) * b = a"
- using div_mult_mod_eq [of a b] by (simp add: assms)
- finally show ?thesis .
-qed
-
-lemma mod_eq_0_iff_dvd:
- "a mod b = 0 \<longleftrightarrow> b dvd a"
- by (auto intro: mod_0_imp_dvd)
-
-lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]:
- "a dvd b \<longleftrightarrow> b mod a = 0"
- by (simp add: mod_eq_0_iff_dvd)
-
-lemma dvd_mod_iff:
- assumes "c dvd b"
- shows "c dvd a mod b \<longleftrightarrow> c dvd a"
-proof -
- from assms have "(c dvd a mod b) \<longleftrightarrow> (c dvd ((a div b) * b + a mod b))"
- by (simp add: dvd_add_right_iff)
- also have "(a div b) * b + a mod b = a"
- using div_mult_mod_eq [of a b] by simp
- finally show ?thesis .
-qed
-
-lemma dvd_mod_imp_dvd:
- assumes "c dvd a mod b" and "c dvd b"
- shows "c dvd a"
- using assms dvd_mod_iff [of c b a] by simp
-
-end
-
-class idom_modulo = idom + semidom_modulo
-begin
-
-subclass idom_divide ..
-
-lemma div_diff [simp]:
- "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> (a - b) div c = a div c - b div c"
- using div_add [of _ _ "- b"] by (simp add: dvd_neg_div)
-
-end
-
-
subsection \<open>Euclidean (semi)rings with explicit division and remainder\<close>
-class euclidean_semiring = semidom_modulo + normalization_semidom +
+class euclidean_semiring = semidom_modulo +
fixes euclidean_size :: "'a \<Rightarrow> nat"
assumes size_0 [simp]: "euclidean_size 0 = 0"
assumes mod_size_less:
@@ -99,22 +23,6 @@
lemma size_mult_mono': "b \<noteq> 0 \<Longrightarrow> euclidean_size a \<le> euclidean_size (b * a)"
by (subst mult.commute) (rule size_mult_mono)
-lemma euclidean_size_normalize [simp]:
- "euclidean_size (normalize a) = euclidean_size a"
-proof (cases "a = 0")
- case True
- then show ?thesis
- by simp
-next
- case [simp]: False
- have "euclidean_size (normalize a) \<le> euclidean_size (normalize a * unit_factor a)"
- by (rule size_mult_mono) simp
- moreover have "euclidean_size a \<le> euclidean_size (a * (1 div unit_factor a))"
- by (rule size_mult_mono) simp
- ultimately show ?thesis
- by simp
-qed
-
lemma dvd_euclidean_size_eq_imp_dvd:
assumes "a \<noteq> 0" and "euclidean_size a = euclidean_size b"
and "b dvd a"
@@ -192,11 +100,397 @@
by (auto intro!: euclidean_size_times_nonunit)
qed
+lemma unit_imp_mod_eq_0:
+ "a mod b = 0" if "is_unit b"
+ using that by (simp add: mod_eq_0_iff_dvd unit_imp_dvd)
+
end
class euclidean_ring = idom_modulo + euclidean_semiring
+subsection \<open>Euclidean (semi)rings with cancel rules\<close>
+
+class euclidean_semiring_cancel = euclidean_semiring +
+ assumes div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b"
+ and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b"
+begin
+
+lemma div_mult_self2 [simp]:
+ assumes "b \<noteq> 0"
+ shows "(a + b * c) div b = c + a div b"
+ using assms div_mult_self1 [of b a c] by (simp add: mult.commute)
+
+lemma div_mult_self3 [simp]:
+ assumes "b \<noteq> 0"
+ shows "(c * b + a) div b = c + a div b"
+ using assms by (simp add: add.commute)
+
+lemma div_mult_self4 [simp]:
+ assumes "b \<noteq> 0"
+ shows "(b * c + a) div b = c + a div b"
+ using assms by (simp add: add.commute)
+
+lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
+proof (cases "b = 0")
+ case True then show ?thesis by simp
+next
+ case False
+ have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b"
+ by (simp add: div_mult_mod_eq)
+ also from False div_mult_self1 [of b a c] have
+ "\<dots> = (c + a div b) * b + (a + c * b) mod b"
+ by (simp add: algebra_simps)
+ finally have "a = a div b * b + (a + c * b) mod b"
+ by (simp add: add.commute [of a] add.assoc distrib_right)
+ then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b"
+ by (simp add: div_mult_mod_eq)
+ then show ?thesis by simp
+qed
+
+lemma mod_mult_self2 [simp]:
+ "(a + b * c) mod b = a mod b"
+ by (simp add: mult.commute [of b])
+
+lemma mod_mult_self3 [simp]:
+ "(c * b + a) mod b = a mod b"
+ by (simp add: add.commute)
+
+lemma mod_mult_self4 [simp]:
+ "(b * c + a) mod b = a mod b"
+ by (simp add: add.commute)
+
+lemma mod_mult_self1_is_0 [simp]:
+ "b * a mod b = 0"
+ using mod_mult_self2 [of 0 b a] by simp
+
+lemma mod_mult_self2_is_0 [simp]:
+ "a * b mod b = 0"
+ using mod_mult_self1 [of 0 a b] by simp
+
+lemma div_add_self1:
+ assumes "b \<noteq> 0"
+ shows "(b + a) div b = a div b + 1"
+ using assms div_mult_self1 [of b a 1] by (simp add: add.commute)
+
+lemma div_add_self2:
+ assumes "b \<noteq> 0"
+ shows "(a + b) div b = a div b + 1"
+ using assms div_add_self1 [of b a] by (simp add: add.commute)
+
+lemma mod_add_self1 [simp]:
+ "(b + a) mod b = a mod b"
+ using mod_mult_self1 [of a 1 b] by (simp add: add.commute)
+
+lemma mod_add_self2 [simp]:
+ "(a + b) mod b = a mod b"
+ using mod_mult_self1 [of a 1 b] by simp
+
+lemma mod_div_trivial [simp]:
+ "a mod b div b = 0"
+proof (cases "b = 0")
+ assume "b = 0"
+ thus ?thesis by simp
+next
+ assume "b \<noteq> 0"
+ hence "a div b + a mod b div b = (a mod b + a div b * b) div b"
+ by (rule div_mult_self1 [symmetric])
+ also have "\<dots> = a div b"
+ by (simp only: mod_div_mult_eq)
+ also have "\<dots> = a div b + 0"
+ by simp
+ finally show ?thesis
+ by (rule add_left_imp_eq)
+qed
+
+lemma mod_mod_trivial [simp]:
+ "a mod b mod b = a mod b"
+proof -
+ have "a mod b mod b = (a mod b + a div b * b) mod b"
+ by (simp only: mod_mult_self1)
+ also have "\<dots> = a mod b"
+ by (simp only: mod_div_mult_eq)
+ finally show ?thesis .
+qed
+
+lemma mod_mod_cancel:
+ assumes "c dvd b"
+ shows "a mod b mod c = a mod c"
+proof -
+ from \<open>c dvd b\<close> obtain k where "b = c * k"
+ by (rule dvdE)
+ have "a mod b mod c = a mod (c * k) mod c"
+ by (simp only: \<open>b = c * k\<close>)
+ also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c"
+ by (simp only: mod_mult_self1)
+ also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c"
+ by (simp only: ac_simps)
+ also have "\<dots> = a mod c"
+ by (simp only: div_mult_mod_eq)
+ finally show ?thesis .
+qed
+
+lemma div_mult_mult2 [simp]:
+ "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b"
+ by (drule div_mult_mult1) (simp add: mult.commute)
+
+lemma div_mult_mult1_if [simp]:
+ "(c * a) div (c * b) = (if c = 0 then 0 else a div b)"
+ by simp_all
+
+lemma mod_mult_mult1:
+ "(c * a) mod (c * b) = c * (a mod b)"
+proof (cases "c = 0")
+ case True then show ?thesis by simp
+next
+ case False
+ from div_mult_mod_eq
+ have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" .
+ with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b)
+ = c * a + c * (a mod b)" by (simp add: algebra_simps)
+ with div_mult_mod_eq show ?thesis by simp
+qed
+
+lemma mod_mult_mult2:
+ "(a * c) mod (b * c) = (a mod b) * c"
+ using mod_mult_mult1 [of c a b] by (simp add: mult.commute)
+
+lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)"
+ by (fact mod_mult_mult2 [symmetric])
+
+lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)"
+ by (fact mod_mult_mult1 [symmetric])
+
+lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)"
+ unfolding dvd_def by (auto simp add: mod_mult_mult1)
+
+lemma div_plus_div_distrib_dvd_left:
+ "c dvd a \<Longrightarrow> (a + b) div c = a div c + b div c"
+ by (cases "c = 0") (auto elim: dvdE)
+
+lemma div_plus_div_distrib_dvd_right:
+ "c dvd b \<Longrightarrow> (a + b) div c = a div c + b div c"
+ using div_plus_div_distrib_dvd_left [of c b a]
+ by (simp add: ac_simps)
+
+named_theorems mod_simps
+
+text \<open>Addition respects modular equivalence.\<close>
+
+lemma mod_add_left_eq [mod_simps]:
+ "(a mod c + b) mod c = (a + b) mod c"
+proof -
+ have "(a + b) mod c = (a div c * c + a mod c + b) mod c"
+ by (simp only: div_mult_mod_eq)
+ also have "\<dots> = (a mod c + b + a div c * c) mod c"
+ by (simp only: ac_simps)
+ also have "\<dots> = (a mod c + b) mod c"
+ by (rule mod_mult_self1)
+ finally show ?thesis
+ by (rule sym)
+qed
+
+lemma mod_add_right_eq [mod_simps]:
+ "(a + b mod c) mod c = (a + b) mod c"
+ using mod_add_left_eq [of b c a] by (simp add: ac_simps)
+
+lemma mod_add_eq:
+ "(a mod c + b mod c) mod c = (a + b) mod c"
+ by (simp add: mod_add_left_eq mod_add_right_eq)
+
+lemma mod_sum_eq [mod_simps]:
+ "(\<Sum>i\<in>A. f i mod a) mod a = sum f A mod a"
+proof (induct A rule: infinite_finite_induct)
+ case (insert i A)
+ then have "(\<Sum>i\<in>insert i A. f i mod a) mod a
+ = (f i mod a + (\<Sum>i\<in>A. f i mod a)) mod a"
+ by simp
+ also have "\<dots> = (f i + (\<Sum>i\<in>A. f i mod a) mod a) mod a"
+ by (simp add: mod_simps)
+ also have "\<dots> = (f i + (\<Sum>i\<in>A. f i) mod a) mod a"
+ by (simp add: insert.hyps)
+ finally show ?case
+ by (simp add: insert.hyps mod_simps)
+qed simp_all
+
+lemma mod_add_cong:
+ assumes "a mod c = a' mod c"
+ assumes "b mod c = b' mod c"
+ shows "(a + b) mod c = (a' + b') mod c"
+proof -
+ have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c"
+ unfolding assms ..
+ then show ?thesis
+ by (simp add: mod_add_eq)
+qed
+
+text \<open>Multiplication respects modular equivalence.\<close>
+
+lemma mod_mult_left_eq [mod_simps]:
+ "((a mod c) * b) mod c = (a * b) mod c"
+proof -
+ have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c"
+ by (simp only: div_mult_mod_eq)
+ also have "\<dots> = (a mod c * b + a div c * b * c) mod c"
+ by (simp only: algebra_simps)
+ also have "\<dots> = (a mod c * b) mod c"
+ by (rule mod_mult_self1)
+ finally show ?thesis
+ by (rule sym)
+qed
+
+lemma mod_mult_right_eq [mod_simps]:
+ "(a * (b mod c)) mod c = (a * b) mod c"
+ using mod_mult_left_eq [of b c a] by (simp add: ac_simps)
+
+lemma mod_mult_eq:
+ "((a mod c) * (b mod c)) mod c = (a * b) mod c"
+ by (simp add: mod_mult_left_eq mod_mult_right_eq)
+
+lemma mod_prod_eq [mod_simps]:
+ "(\<Prod>i\<in>A. f i mod a) mod a = prod f A mod a"
+proof (induct A rule: infinite_finite_induct)
+ case (insert i A)
+ then have "(\<Prod>i\<in>insert i A. f i mod a) mod a
+ = (f i mod a * (\<Prod>i\<in>A. f i mod a)) mod a"
+ by simp
+ also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i mod a) mod a)) mod a"
+ by (simp add: mod_simps)
+ also have "\<dots> = (f i * ((\<Prod>i\<in>A. f i) mod a)) mod a"
+ by (simp add: insert.hyps)
+ finally show ?case
+ by (simp add: insert.hyps mod_simps)
+qed simp_all
+
+lemma mod_mult_cong:
+ assumes "a mod c = a' mod c"
+ assumes "b mod c = b' mod c"
+ shows "(a * b) mod c = (a' * b') mod c"
+proof -
+ have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c"
+ unfolding assms ..
+ then show ?thesis
+ by (simp add: mod_mult_eq)
+qed
+
+text \<open>Exponentiation respects modular equivalence.\<close>
+
+lemma power_mod [mod_simps]:
+ "((a mod b) ^ n) mod b = (a ^ n) mod b"
+proof (induct n)
+ case 0
+ then show ?case by simp
+next
+ case (Suc n)
+ have "(a mod b) ^ Suc n mod b = (a mod b) * ((a mod b) ^ n mod b) mod b"
+ by (simp add: mod_mult_right_eq)
+ with Suc show ?case
+ by (simp add: mod_mult_left_eq mod_mult_right_eq)
+qed
+
+end
+
+
+class euclidean_ring_cancel = euclidean_ring + euclidean_semiring_cancel
+begin
+
+subclass idom_divide ..
+
+lemma div_minus_minus [simp]: "(- a) div (- b) = a div b"
+ using div_mult_mult1 [of "- 1" a b] by simp
+
+lemma mod_minus_minus [simp]: "(- a) mod (- b) = - (a mod b)"
+ using mod_mult_mult1 [of "- 1" a b] by simp
+
+lemma div_minus_right: "a div (- b) = (- a) div b"
+ using div_minus_minus [of "- a" b] by simp
+
+lemma mod_minus_right: "a mod (- b) = - ((- a) mod b)"
+ using mod_minus_minus [of "- a" b] by simp
+
+lemma div_minus1_right [simp]: "a div (- 1) = - a"
+ using div_minus_right [of a 1] by simp
+
+lemma mod_minus1_right [simp]: "a mod (- 1) = 0"
+ using mod_minus_right [of a 1] by simp
+
+text \<open>Negation respects modular equivalence.\<close>
+
+lemma mod_minus_eq [mod_simps]:
+ "(- (a mod b)) mod b = (- a) mod b"
+proof -
+ have "(- a) mod b = (- (a div b * b + a mod b)) mod b"
+ by (simp only: div_mult_mod_eq)
+ also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b"
+ by (simp add: ac_simps)
+ also have "\<dots> = (- (a mod b)) mod b"
+ by (rule mod_mult_self1)
+ finally show ?thesis
+ by (rule sym)
+qed
+
+lemma mod_minus_cong:
+ assumes "a mod b = a' mod b"
+ shows "(- a) mod b = (- a') mod b"
+proof -
+ have "(- (a mod b)) mod b = (- (a' mod b)) mod b"
+ unfolding assms ..
+ then show ?thesis
+ by (simp add: mod_minus_eq)
+qed
+
+text \<open>Subtraction respects modular equivalence.\<close>
+
+lemma mod_diff_left_eq [mod_simps]:
+ "(a mod c - b) mod c = (a - b) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- b"]
+ by simp
+
+lemma mod_diff_right_eq [mod_simps]:
+ "(a - b mod c) mod c = (a - b) mod c"
+ using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
+ by simp
+
+lemma mod_diff_eq:
+ "(a mod c - b mod c) mod c = (a - b) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b]
+ by simp
+
+lemma mod_diff_cong:
+ assumes "a mod c = a' mod c"
+ assumes "b mod c = b' mod c"
+ shows "(a - b) mod c = (a' - b') mod c"
+ using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"]
+ by simp
+
+lemma minus_mod_self2 [simp]:
+ "(a - b) mod b = a mod b"
+ using mod_diff_right_eq [of a b b]
+ by (simp add: mod_diff_right_eq)
+
+lemma minus_mod_self1 [simp]:
+ "(b - a) mod b = - a mod b"
+ using mod_add_self2 [of "- a" b] by simp
+
+lemma mod_eq_dvd_iff:
+ "a mod c = b mod c \<longleftrightarrow> c dvd a - b" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?P
+ then have "(a mod c - b mod c) mod c = 0"
+ by simp
+ then show ?Q
+ by (simp add: dvd_eq_mod_eq_0 mod_simps)
+next
+ assume ?Q
+ then obtain d where d: "a - b = c * d" ..
+ then have "a = c * d + b"
+ by (simp add: algebra_simps)
+ then show ?P by simp
+qed
+
+end
+
+
subsection \<open>Uniquely determined division\<close>
class unique_euclidean_semiring = euclidean_semiring +
@@ -221,9 +515,10 @@
and "a div b = q"
and "a mod b = 0"
and "a = q * b"
- | (remainder) q r where "b \<noteq> 0" and "r \<noteq> 0"
+ | (remainder) q r where "b \<noteq> 0"
and "uniqueness_constraint r b"
and "euclidean_size r < euclidean_size b"
+ and "r \<noteq> 0"
and "a div b = q"
and "a mod b = r"
and "a = q * b + r"
@@ -280,8 +575,931 @@
by simp
qed
+subclass euclidean_semiring_cancel
+proof
+ show "(a + c * b) div b = c + a div b" if "b \<noteq> 0" for a b c
+ proof (cases a b rule: divmod_cases)
+ case by0
+ with \<open>b \<noteq> 0\<close> show ?thesis
+ by simp
+ next
+ case (divides q)
+ then show ?thesis
+ by (simp add: ac_simps)
+ next
+ case (remainder q r)
+ then show ?thesis
+ by (auto intro: div_eqI simp add: algebra_simps)
+ qed
+next
+ show"(c * a) div (c * b) = a div b" if "c \<noteq> 0" for a b c
+ proof (cases a b rule: divmod_cases)
+ case by0
+ then show ?thesis
+ by simp
+ next
+ case (divides q)
+ with \<open>c \<noteq> 0\<close> show ?thesis
+ by (simp add: mult.left_commute [of c])
+ next
+ case (remainder q r)
+ from \<open>b \<noteq> 0\<close> \<open>c \<noteq> 0\<close> have "b * c \<noteq> 0"
+ by simp
+ from remainder \<open>c \<noteq> 0\<close>
+ have "uniqueness_constraint (r * c) (b * c)"
+ and "euclidean_size (r * c) < euclidean_size (b * c)"
+ by (simp_all add: uniqueness_constraint_mono_mult uniqueness_constraint_mod size_mono_mult)
+ with remainder show ?thesis
+ by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps)
+ (use \<open>b * c \<noteq> 0\<close> in simp)
+ qed
+qed
+
+lemma div_mult1_eq:
+ "(a * b) div c = a * (b div c) + a * (b mod c) div c"
+proof (cases "a * (b mod c)" c rule: divmod_cases)
+ case (divides q)
+ have "a * b = a * (b div c * c + b mod c)"
+ by (simp add: div_mult_mod_eq)
+ also have "\<dots> = (a * (b div c) + q) * c"
+ using divides by (simp add: algebra_simps)
+ finally have "(a * b) div c = \<dots> div c"
+ by simp
+ with divides show ?thesis
+ by simp
+next
+ case (remainder q r)
+ from remainder(1-3) show ?thesis
+ proof (rule div_eqI)
+ have "a * b = a * (b div c * c + b mod c)"
+ by (simp add: div_mult_mod_eq)
+ also have "\<dots> = a * c * (b div c) + q * c + r"
+ using remainder by (simp add: algebra_simps)
+ finally show "(a * (b div c) + a * (b mod c) div c) * c + r = a * b"
+ using remainder(5-7) by (simp add: algebra_simps)
+ qed
+next
+ case by0
+ then show ?thesis
+ by simp
+qed
+
+lemma div_add1_eq:
+ "(a + b) div c = a div c + b div c + (a mod c + b mod c) div c"
+proof (cases "a mod c + b mod c" c rule: divmod_cases)
+ case (divides q)
+ have "a + b = (a div c * c + a mod c) + (b div c * c + b mod c)"
+ using mod_mult_div_eq [of a c] mod_mult_div_eq [of b c] by (simp add: ac_simps)
+ also have "\<dots> = (a div c + b div c) * c + (a mod c + b mod c)"
+ by (simp add: algebra_simps)
+ also have "\<dots> = (a div c + b div c + q) * c"
+ using divides by (simp add: algebra_simps)
+ finally have "(a + b) div c = (a div c + b div c + q) * c div c"
+ by simp
+ with divides show ?thesis
+ by simp
+next
+ case (remainder q r)
+ from remainder(1-3) show ?thesis
+ proof (rule div_eqI)
+ have "(a div c + b div c + q) * c + r + (a mod c + b mod c) =
+ (a div c * c + a mod c) + (b div c * c + b mod c) + q * c + r"
+ by (simp add: algebra_simps)
+ also have "\<dots> = a + b + (a mod c + b mod c)"
+ by (simp add: div_mult_mod_eq remainder) (simp add: ac_simps)
+ finally show "(a div c + b div c + (a mod c + b mod c) div c) * c + r = a + b"
+ using remainder by simp
+ qed
+next
+ case by0
+ then show ?thesis
+ by simp
+qed
+
end
class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring
+begin
+
+subclass euclidean_ring_cancel ..
+
+end
+
+
+subsection \<open>Euclidean division on @{typ nat}\<close>
+
+instantiation nat :: normalization_semidom
+begin
+
+definition normalize_nat :: "nat \<Rightarrow> nat"
+ where [simp]: "normalize = (id :: nat \<Rightarrow> nat)"
+
+definition unit_factor_nat :: "nat \<Rightarrow> nat"
+ where "unit_factor n = (if n = 0 then 0 else 1 :: nat)"
+
+lemma unit_factor_simps [simp]:
+ "unit_factor 0 = (0::nat)"
+ "unit_factor (Suc n) = 1"
+ by (simp_all add: unit_factor_nat_def)
+
+definition divide_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+ where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \<le> m})"
+
+instance
+ by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI)
+
+end
+
+instantiation nat :: unique_euclidean_semiring
+begin
+
+definition euclidean_size_nat :: "nat \<Rightarrow> nat"
+ where [simp]: "euclidean_size_nat = id"
+
+definition uniqueness_constraint_nat :: "nat \<Rightarrow> nat \<Rightarrow> bool"
+ where [simp]: "uniqueness_constraint_nat = \<top>"
+
+definition modulo_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+ where "m mod n = m - (m div n * (n::nat))"
+
+instance proof
+ fix m n :: nat
+ have ex: "\<exists>k. k * n \<le> l" for l :: nat
+ by (rule exI [of _ 0]) simp
+ have fin: "finite {k. k * n \<le> l}" if "n > 0" for l
+ proof -
+ from that have "{k. k * n \<le> l} \<subseteq> {k. k \<le> l}"
+ by (cases n) auto
+ then show ?thesis
+ by (rule finite_subset) simp
+ qed
+ have mult_div_unfold: "n * (m div n) = Max {l. l \<le> m \<and> n dvd l}"
+ proof (cases "n = 0")
+ case True
+ moreover have "{l. l = 0 \<and> l \<le> m} = {0::nat}"
+ by auto
+ ultimately show ?thesis
+ by simp
+ next
+ case False
+ with ex [of m] fin have "n * Max {k. k * n \<le> m} = Max (times n ` {k. k * n \<le> m})"
+ by (auto simp add: nat_mult_max_right intro: hom_Max_commute)
+ also have "times n ` {k. k * n \<le> m} = {l. l \<le> m \<and> n dvd l}"
+ by (auto simp add: ac_simps elim!: dvdE)
+ finally show ?thesis
+ using False by (simp add: divide_nat_def ac_simps)
+ qed
+ have less_eq: "m div n * n \<le> m"
+ by (auto simp add: mult_div_unfold ac_simps intro: Max.boundedI)
+ then show "m div n * n + m mod n = m"
+ by (simp add: modulo_nat_def)
+ assume "n \<noteq> 0"
+ show "euclidean_size (m mod n) < euclidean_size n"
+ proof -
+ have "m < Suc (m div n) * n"
+ proof (rule ccontr)
+ assume "\<not> m < Suc (m div n) * n"
+ then have "Suc (m div n) * n \<le> m"
+ by (simp add: not_less)
+ moreover from \<open>n \<noteq> 0\<close> have "Max {k. k * n \<le> m} < Suc (m div n)"
+ by (simp add: divide_nat_def)
+ with \<open>n \<noteq> 0\<close> ex fin have "\<And>k. k * n \<le> m \<Longrightarrow> k < Suc (m div n)"
+ by auto
+ ultimately have "Suc (m div n) < Suc (m div n)"
+ by blast
+ then show False
+ by simp
+ qed
+ with \<open>n \<noteq> 0\<close> show ?thesis
+ by (simp add: modulo_nat_def)
+ qed
+ show "euclidean_size m \<le> euclidean_size (m * n)"
+ using \<open>n \<noteq> 0\<close> by (cases n) simp_all
+ fix q r :: nat
+ show "(q * n + r) div n = q" if "euclidean_size r < euclidean_size n"
+ proof -
+ from that have "r < n"
+ by simp
+ have "k \<le> q" if "k * n \<le> q * n + r" for k
+ proof (rule ccontr)
+ assume "\<not> k \<le> q"
+ then have "q < k"
+ by simp
+ then obtain l where "k = Suc (q + l)"
+ by (auto simp add: less_iff_Suc_add)
+ with \<open>r < n\<close> that show False
+ by (simp add: algebra_simps)
+ qed
+ with \<open>n \<noteq> 0\<close> ex fin show ?thesis
+ by (auto simp add: divide_nat_def Max_eq_iff)
+ qed
+qed simp_all
+
+end
+
+text \<open>Tool support\<close>
+
+ML \<open>
+structure Cancel_Div_Mod_Nat = Cancel_Div_Mod
+(
+ val div_name = @{const_name divide};
+ val mod_name = @{const_name modulo};
+ val mk_binop = HOLogic.mk_binop;
+ val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
+ val mk_sum = Arith_Data.mk_sum;
+ fun dest_sum tm =
+ if HOLogic.is_zero tm then []
+ else
+ (case try HOLogic.dest_Suc tm of
+ SOME t => HOLogic.Suc_zero :: dest_sum t
+ | NONE =>
+ (case try dest_plus tm of
+ SOME (t, u) => dest_sum t @ dest_sum u
+ | NONE => [tm]));
+
+ val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
+
+ val prove_eq_sums = Arith_Data.prove_conv2 all_tac
+ (Arith_Data.simp_all_tac @{thms add_0_left add_0_right ac_simps})
+)
+\<close>
+
+simproc_setup cancel_div_mod_nat ("(m::nat) + n") =
+ \<open>K Cancel_Div_Mod_Nat.proc\<close>
+
+lemma div_nat_eqI:
+ "m div n = q" if "n * q \<le> m" and "m < n * Suc q" for m n q :: nat
+ by (rule div_eqI [of _ "m - n * q"]) (use that in \<open>simp_all add: algebra_simps\<close>)
+
+lemma mod_nat_eqI:
+ "m mod n = r" if "r < n" and "r \<le> m" and "n dvd m - r" for m n r :: nat
+ by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \<open>simp_all add: algebra_simps\<close>)
+
+lemma div_mult_self_is_m [simp]:
+ "m * n div n = m" if "n > 0" for m n :: nat
+ using that by simp
+
+lemma div_mult_self1_is_m [simp]:
+ "n * m div n = m" if "n > 0" for m n :: nat
+ using that by simp
+
+lemma mod_less_divisor [simp]:
+ "m mod n < n" if "n > 0" for m n :: nat
+ using mod_size_less [of n m] that by simp
+
+lemma mod_le_divisor [simp]:
+ "m mod n \<le> n" if "n > 0" for m n :: nat
+ using that by (auto simp add: le_less)
+
+lemma div_times_less_eq_dividend [simp]:
+ "m div n * n \<le> m" for m n :: nat
+ by (simp add: minus_mod_eq_div_mult [symmetric])
+
+lemma times_div_less_eq_dividend [simp]:
+ "n * (m div n) \<le> m" for m n :: nat
+ using div_times_less_eq_dividend [of m n]
+ by (simp add: ac_simps)
+
+lemma dividend_less_div_times:
+ "m < n + (m div n) * n" if "0 < n" for m n :: nat
+proof -
+ from that have "m mod n < n"
+ by simp
+ then show ?thesis
+ by (simp add: minus_mod_eq_div_mult [symmetric])
+qed
+
+lemma dividend_less_times_div:
+ "m < n + n * (m div n)" if "0 < n" for m n :: nat
+ using dividend_less_div_times [of n m] that
+ by (simp add: ac_simps)
+
+lemma mod_Suc_le_divisor [simp]:
+ "m mod Suc n \<le> n"
+ using mod_less_divisor [of "Suc n" m] by arith
+
+lemma mod_less_eq_dividend [simp]:
+ "m mod n \<le> m" for m n :: nat
+proof (rule add_leD2)
+ from div_mult_mod_eq have "m div n * n + m mod n = m" .
+ then show "m div n * n + m mod n \<le> m" by auto
+qed
+
+lemma
+ div_less [simp]: "m div n = 0"
+ and mod_less [simp]: "m mod n = m"
+ if "m < n" for m n :: nat
+ using that by (auto intro: div_eqI mod_eqI)
+
+lemma le_div_geq:
+ "m div n = Suc ((m - n) div n)" if "0 < n" and "n \<le> m" for m n :: nat
+proof -
+ from \<open>n \<le> m\<close> obtain q where "m = n + q"
+ by (auto simp add: le_iff_add)
+ with \<open>0 < n\<close> show ?thesis
+ by (simp add: div_add_self1)
+qed
+
+lemma le_mod_geq:
+ "m mod n = (m - n) mod n" if "n \<le> m" for m n :: nat
+proof -
+ from \<open>n \<le> m\<close> obtain q where "m = n + q"
+ by (auto simp add: le_iff_add)
+ then show ?thesis
+ by simp
+qed
+
+lemma div_if:
+ "m div n = (if m < n \<or> n = 0 then 0 else Suc ((m - n) div n))"
+ by (simp add: le_div_geq)
+
+lemma mod_if:
+ "m mod n = (if m < n then m else (m - n) mod n)" for m n :: nat
+ by (simp add: le_mod_geq)
+
+lemma div_eq_0_iff:
+ "m div n = 0 \<longleftrightarrow> m < n \<or> n = 0" for m n :: nat
+ by (simp add: div_if)
+
+lemma div_greater_zero_iff:
+ "m div n > 0 \<longleftrightarrow> n \<le> m \<and> n > 0" for m n :: nat
+ using div_eq_0_iff [of m n] by auto
+
+lemma mod_greater_zero_iff_not_dvd:
+ "m mod n > 0 \<longleftrightarrow> \<not> n dvd m" for m n :: nat
+ by (simp add: dvd_eq_mod_eq_0)
+
+lemma div_by_Suc_0 [simp]:
+ "m div Suc 0 = m"
+ using div_by_1 [of m] by simp
+
+lemma mod_by_Suc_0 [simp]:
+ "m mod Suc 0 = 0"
+ using mod_by_1 [of m] by simp
+
+lemma div2_Suc_Suc [simp]:
+ "Suc (Suc m) div 2 = Suc (m div 2)"
+ by (simp add: numeral_2_eq_2 le_div_geq)
+
+lemma Suc_n_div_2_gt_zero [simp]:
+ "0 < Suc n div 2" if "n > 0" for n :: nat
+ using that by (cases n) simp_all
+
+lemma div_2_gt_zero [simp]:
+ "0 < n div 2" if "Suc 0 < n" for n :: nat
+ using that Suc_n_div_2_gt_zero [of "n - 1"] by simp
+
+lemma mod2_Suc_Suc [simp]:
+ "Suc (Suc m) mod 2 = m mod 2"
+ by (simp add: numeral_2_eq_2 le_mod_geq)
+
+lemma add_self_div_2 [simp]:
+ "(m + m) div 2 = m" for m :: nat
+ by (simp add: mult_2 [symmetric])
+
+lemma add_self_mod_2 [simp]:
+ "(m + m) mod 2 = 0" for m :: nat
+ by (simp add: mult_2 [symmetric])
+
+lemma mod2_gr_0 [simp]:
+ "0 < m mod 2 \<longleftrightarrow> m mod 2 = 1" for m :: nat
+proof -
+ have "m mod 2 < 2"
+ by (rule mod_less_divisor) simp
+ then have "m mod 2 = 0 \<or> m mod 2 = 1"
+ by arith
+ then show ?thesis
+ by auto
+qed
+
+lemma mod_Suc_eq [mod_simps]:
+ "Suc (m mod n) mod n = Suc m mod n"
+proof -
+ have "(m mod n + 1) mod n = (m + 1) mod n"
+ by (simp only: mod_simps)
+ then show ?thesis
+ by simp
+qed
+
+lemma mod_Suc_Suc_eq [mod_simps]:
+ "Suc (Suc (m mod n)) mod n = Suc (Suc m) mod n"
+proof -
+ have "(m mod n + 2) mod n = (m + 2) mod n"
+ by (simp only: mod_simps)
+ then show ?thesis
+ by simp
+qed
+
+lemma
+ Suc_mod_mult_self1 [simp]: "Suc (m + k * n) mod n = Suc m mod n"
+ and Suc_mod_mult_self2 [simp]: "Suc (m + n * k) mod n = Suc m mod n"
+ and Suc_mod_mult_self3 [simp]: "Suc (k * n + m) mod n = Suc m mod n"
+ and Suc_mod_mult_self4 [simp]: "Suc (n * k + m) mod n = Suc m mod n"
+ by (subst mod_Suc_eq [symmetric], simp add: mod_simps)+
+
+context
+ fixes m n q :: nat
+begin
+
+private lemma eucl_rel_mult2:
+ "m mod n + n * (m div n mod q) < n * q"
+ if "n > 0" and "q > 0"
+proof -
+ from \<open>n > 0\<close> have "m mod n < n"
+ by (rule mod_less_divisor)
+ from \<open>q > 0\<close> have "m div n mod q < q"
+ by (rule mod_less_divisor)
+ then obtain s where "q = Suc (m div n mod q + s)"
+ by (blast dest: less_imp_Suc_add)
+ moreover have "m mod n + n * (m div n mod q) < n * Suc (m div n mod q + s)"
+ using \<open>m mod n < n\<close> by (simp add: add_mult_distrib2)
+ ultimately show ?thesis
+ by simp
+qed
+
+lemma div_mult2_eq:
+ "m div (n * q) = (m div n) div q"
+proof (cases "n = 0 \<or> q = 0")
+ case True
+ then show ?thesis
+ by auto
+next
+ case False
+ with eucl_rel_mult2 show ?thesis
+ by (auto intro: div_eqI [of _ "n * (m div n mod q) + m mod n"]
+ simp add: algebra_simps add_mult_distrib2 [symmetric])
+qed
+
+lemma mod_mult2_eq:
+ "m mod (n * q) = n * (m div n mod q) + m mod n"
+proof (cases "n = 0 \<or> q = 0")
+ case True
+ then show ?thesis
+ by auto
+next
+ case False
+ with eucl_rel_mult2 show ?thesis
+ by (auto intro: mod_eqI [of _ _ "(m div n) div q"]
+ simp add: algebra_simps add_mult_distrib2 [symmetric])
+qed
end
+
+lemma div_le_mono:
+ "m div k \<le> n div k" if "m \<le> n" for m n k :: nat
+proof -
+ from that obtain q where "n = m + q"
+ by (auto simp add: le_iff_add)
+ then show ?thesis
+ by (simp add: div_add1_eq [of m q k])
+qed
+
+text \<open>Antimonotonicity of @{const divide} in second argument\<close>
+
+lemma div_le_mono2:
+ "k div n \<le> k div m" if "0 < m" and "m \<le> n" for m n k :: nat
+using that proof (induct k arbitrary: m rule: less_induct)
+ case (less k)
+ show ?case
+ proof (cases "n \<le> k")
+ case False
+ then show ?thesis
+ by simp
+ next
+ case True
+ have "(k - n) div n \<le> (k - m) div n"
+ using less.prems
+ by (blast intro: div_le_mono diff_le_mono2)
+ also have "\<dots> \<le> (k - m) div m"
+ using \<open>n \<le> k\<close> less.prems less.hyps [of "k - m" m]
+ by simp
+ finally show ?thesis
+ using \<open>n \<le> k\<close> less.prems
+ by (simp add: le_div_geq)
+ qed
+qed
+
+lemma div_le_dividend [simp]:
+ "m div n \<le> m" for m n :: nat
+ using div_le_mono2 [of 1 n m] by (cases "n = 0") simp_all
+
+lemma div_less_dividend [simp]:
+ "m div n < m" if "1 < n" and "0 < m" for m n :: nat
+using that proof (induct m rule: less_induct)
+ case (less m)
+ show ?case
+ proof (cases "n < m")
+ case False
+ with less show ?thesis
+ by (cases "n = m") simp_all
+ next
+ case True
+ then show ?thesis
+ using less.hyps [of "m - n"] less.prems
+ by (simp add: le_div_geq)
+ qed
+qed
+
+lemma div_eq_dividend_iff:
+ "m div n = m \<longleftrightarrow> n = 1" if "m > 0" for m n :: nat
+proof
+ assume "n = 1"
+ then show "m div n = m"
+ by simp
+next
+ assume P: "m div n = m"
+ show "n = 1"
+ proof (rule ccontr)
+ have "n \<noteq> 0"
+ by (rule ccontr) (use that P in auto)
+ moreover assume "n \<noteq> 1"
+ ultimately have "n > 1"
+ by simp
+ with that have "m div n < m"
+ by simp
+ with P show False
+ by simp
+ qed
+qed
+
+lemma less_mult_imp_div_less:
+ "m div n < i" if "m < i * n" for m n i :: nat
+proof -
+ from that have "i * n > 0"
+ by (cases "i * n = 0") simp_all
+ then have "i > 0" and "n > 0"
+ by simp_all
+ have "m div n * n \<le> m"
+ by simp
+ then have "m div n * n < i * n"
+ using that by (rule le_less_trans)
+ with \<open>n > 0\<close> show ?thesis
+ by simp
+qed
+
+text \<open>A fact for the mutilated chess board\<close>
+
+lemma mod_Suc:
+ "Suc m mod n = (if Suc (m mod n) = n then 0 else Suc (m mod n))" (is "_ = ?rhs")
+proof (cases "n = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ have "Suc m mod n = Suc (m mod n) mod n"
+ by (simp add: mod_simps)
+ also have "\<dots> = ?rhs"
+ using False by (auto intro!: mod_nat_eqI intro: neq_le_trans simp add: Suc_le_eq)
+ finally show ?thesis .
+qed
+
+lemma Suc_times_mod_eq:
+ "Suc (m * n) mod m = 1" if "Suc 0 < m"
+ using that by (simp add: mod_Suc)
+
+lemma Suc_times_numeral_mod_eq [simp]:
+ "Suc (numeral k * n) mod numeral k = 1" if "numeral k \<noteq> (1::nat)"
+ by (rule Suc_times_mod_eq) (use that in simp)
+
+lemma Suc_div_le_mono [simp]:
+ "m div n \<le> Suc m div n"
+ by (simp add: div_le_mono)
+
+text \<open>These lemmas collapse some needless occurrences of Suc:
+ at least three Sucs, since two and fewer are rewritten back to Suc again!
+ We already have some rules to simplify operands smaller than 3.\<close>
+
+lemma div_Suc_eq_div_add3 [simp]:
+ "m div Suc (Suc (Suc n)) = m div (3 + n)"
+ by (simp add: Suc3_eq_add_3)
+
+lemma mod_Suc_eq_mod_add3 [simp]:
+ "m mod Suc (Suc (Suc n)) = m mod (3 + n)"
+ by (simp add: Suc3_eq_add_3)
+
+lemma Suc_div_eq_add3_div:
+ "Suc (Suc (Suc m)) div n = (3 + m) div n"
+ by (simp add: Suc3_eq_add_3)
+
+lemma Suc_mod_eq_add3_mod:
+ "Suc (Suc (Suc m)) mod n = (3 + m) mod n"
+ by (simp add: Suc3_eq_add_3)
+
+lemmas Suc_div_eq_add3_div_numeral [simp] =
+ Suc_div_eq_add3_div [of _ "numeral v"] for v
+
+lemmas Suc_mod_eq_add3_mod_numeral [simp] =
+ Suc_mod_eq_add3_mod [of _ "numeral v"] for v
+
+lemma (in field_char_0) of_nat_div:
+ "of_nat (m div n) = ((of_nat m - of_nat (m mod n)) / of_nat n)"
+proof -
+ have "of_nat (m div n) = ((of_nat (m div n * n + m mod n) - of_nat (m mod n)) / of_nat n :: 'a)"
+ unfolding of_nat_add by (cases "n = 0") simp_all
+ then show ?thesis
+ by simp
+qed
+
+text \<open>An ``induction'' law for modulus arithmetic.\<close>
+
+lemma mod_induct [consumes 3, case_names step]:
+ "P m" if "P n" and "n < p" and "m < p"
+ and step: "\<And>n. n < p \<Longrightarrow> P n \<Longrightarrow> P (Suc n mod p)"
+using \<open>m < p\<close> proof (induct m)
+ case 0
+ show ?case
+ proof (rule ccontr)
+ assume "\<not> P 0"
+ from \<open>n < p\<close> have "0 < p"
+ by simp
+ from \<open>n < p\<close> obtain m where "0 < m" and "p = n + m"
+ by (blast dest: less_imp_add_positive)
+ with \<open>P n\<close> have "P (p - m)"
+ by simp
+ moreover have "\<not> P (p - m)"
+ using \<open>0 < m\<close> proof (induct m)
+ case 0
+ then show ?case
+ by simp
+ next
+ case (Suc m)
+ show ?case
+ proof
+ assume P: "P (p - Suc m)"
+ with \<open>\<not> P 0\<close> have "Suc m < p"
+ by (auto intro: ccontr)
+ then have "Suc (p - Suc m) = p - m"
+ by arith
+ moreover from \<open>0 < p\<close> have "p - Suc m < p"
+ by arith
+ with P step have "P ((Suc (p - Suc m)) mod p)"
+ by blast
+ ultimately show False
+ using \<open>\<not> P 0\<close> Suc.hyps by (cases "m = 0") simp_all
+ qed
+ qed
+ ultimately show False
+ by blast
+ qed
+next
+ case (Suc m)
+ then have "m < p" and mod: "Suc m mod p = Suc m"
+ by simp_all
+ from \<open>m < p\<close> have "P m"
+ by (rule Suc.hyps)
+ with \<open>m < p\<close> have "P (Suc m mod p)"
+ by (rule step)
+ with mod show ?case
+ by simp
+qed
+
+lemma split_div:
+ "P (m div n) \<longleftrightarrow> (n = 0 \<longrightarrow> P 0) \<and> (n \<noteq> 0 \<longrightarrow>
+ (\<forall>i j. j < n \<longrightarrow> m = n * i + j \<longrightarrow> P i))"
+ (is "?P = ?Q") for m n :: nat
+proof (cases "n = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ show ?thesis
+ proof
+ assume ?P
+ with False show ?Q
+ by auto
+ next
+ assume ?Q
+ with False have *: "\<And>i j. j < n \<Longrightarrow> m = n * i + j \<Longrightarrow> P i"
+ by simp
+ with False show ?P
+ by (auto intro: * [of "m mod n"])
+ qed
+qed
+
+lemma split_div':
+ "P (m div n) \<longleftrightarrow> n = 0 \<and> P 0 \<or> (\<exists>q. (n * q \<le> m \<and> m < n * Suc q) \<and> P q)"
+proof (cases "n = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ then have "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> m div n = q" for q
+ by (auto intro: div_nat_eqI dividend_less_times_div)
+ then show ?thesis
+ by auto
+qed
+
+lemma split_mod:
+ "P (m mod n) \<longleftrightarrow> (n = 0 \<longrightarrow> P m) \<and> (n \<noteq> 0 \<longrightarrow>
+ (\<forall>i j. j < n \<longrightarrow> m = n * i + j \<longrightarrow> P j))"
+ (is "?P \<longleftrightarrow> ?Q") for m n :: nat
+proof (cases "n = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ show ?thesis
+ proof
+ assume ?P
+ with False show ?Q
+ by auto
+ next
+ assume ?Q
+ with False have *: "\<And>i j. j < n \<Longrightarrow> m = n * i + j \<Longrightarrow> P j"
+ by simp
+ with False show ?P
+ by (auto intro: * [of _ "m div n"])
+ qed
+qed
+
+
+subsection \<open>Euclidean division on @{typ int}\<close>
+
+instantiation int :: normalization_semidom
+begin
+
+definition normalize_int :: "int \<Rightarrow> int"
+ where [simp]: "normalize = (abs :: int \<Rightarrow> int)"
+
+definition unit_factor_int :: "int \<Rightarrow> int"
+ where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)"
+
+definition divide_int :: "int \<Rightarrow> int \<Rightarrow> int"
+ where "k div l = (if l = 0 then 0
+ else if sgn k = sgn l
+ then int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)
+ else - int (nat \<bar>k\<bar> div nat \<bar>l\<bar> + of_bool (\<not> l dvd k)))"
+
+lemma divide_int_unfold:
+ "(sgn k * int m) div (sgn l * int n) =
+ (if sgn l = 0 \<or> sgn k = 0 \<or> n = 0 then 0
+ else if sgn k = sgn l
+ then int (m div n)
+ else - int (m div n + of_bool (\<not> n dvd m)))"
+ by (auto simp add: divide_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult
+ nat_mult_distrib dvd_int_iff)
+
+instance proof
+ fix k :: int show "k div 0 = 0"
+ by (simp add: divide_int_def)
+next
+ fix k l :: int
+ assume "l \<noteq> 0"
+ obtain n m and s t where k: "k = sgn s * int n" and l: "l = sgn t * int m"
+ by (blast intro: int_sgnE elim: that)
+ then have "k * l = sgn (s * t) * int (n * m)"
+ by (simp add: ac_simps sgn_mult)
+ with k l \<open>l \<noteq> 0\<close> show "k * l div l = k"
+ by (simp only: divide_int_unfold)
+ (auto simp add: algebra_simps sgn_mult sgn_1_pos sgn_0_0)
+qed (auto simp add: sgn_mult mult_sgn_abs abs_eq_iff')
+
+end
+
+instantiation int :: unique_euclidean_ring
+begin
+
+definition euclidean_size_int :: "int \<Rightarrow> nat"
+ where [simp]: "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)"
+
+definition uniqueness_constraint_int :: "int \<Rightarrow> int \<Rightarrow> bool"
+ where [simp]: "uniqueness_constraint_int k l \<longleftrightarrow> unit_factor k = unit_factor l"
+
+definition modulo_int :: "int \<Rightarrow> int \<Rightarrow> int"
+ where "k mod l = (if l = 0 then k
+ else if sgn k = sgn l
+ then sgn l * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)
+ else sgn l * (\<bar>l\<bar> * of_bool (\<not> l dvd k) - int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)))"
+
+lemma modulo_int_unfold:
+ "(sgn k * int m) mod (sgn l * int n) =
+ (if sgn l = 0 \<or> sgn k = 0 \<or> n = 0 then sgn k * int m
+ else if sgn k = sgn l
+ then sgn l * int (m mod n)
+ else sgn l * (int (n * of_bool (\<not> n dvd m)) - int (m mod n)))"
+ by (auto simp add: modulo_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult
+ nat_mult_distrib dvd_int_iff)
+
+lemma abs_mod_less:
+ "\<bar>k mod l\<bar> < \<bar>l\<bar>" if "l \<noteq> 0" for k l :: int
+proof -
+ obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
+ by (blast intro: int_sgnE elim: that)
+ with that show ?thesis
+ by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg
+ abs_mult mod_greater_zero_iff_not_dvd)
+qed
+
+lemma sgn_mod:
+ "sgn (k mod l) = sgn l" if "l \<noteq> 0" "\<not> l dvd k" for k l :: int
+proof -
+ obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
+ by (blast intro: int_sgnE elim: that)
+ with that show ?thesis
+ by (simp add: modulo_int_unfold sgn_0_0 sgn_1_pos sgn_1_neg
+ sgn_mult mod_eq_0_iff_dvd int_dvd_iff)
+qed
+
+instance proof
+ fix k l :: int
+ obtain n m and s t where "k = sgn s * int n" and "l = sgn t * int m"
+ by (blast intro: int_sgnE elim: that)
+ then show "k div l * l + k mod l = k"
+ by (auto simp add: divide_int_unfold modulo_int_unfold algebra_simps dest!: sgn_not_eq_imp)
+ (simp_all add: of_nat_mult [symmetric] of_nat_add [symmetric]
+ distrib_left [symmetric] minus_mult_right
+ del: of_nat_mult minus_mult_right [symmetric])
+next
+ fix l q r :: int
+ obtain n m and s t
+ where l: "l = sgn s * int n" and q: "q = sgn t * int m"
+ by (blast intro: int_sgnE elim: that)
+ assume \<open>l \<noteq> 0\<close>
+ with l have "s \<noteq> 0" and "n > 0"
+ by (simp_all add: sgn_0_0)
+ assume "uniqueness_constraint r l"
+ moreover have "r = sgn r * \<bar>r\<bar>"
+ by (simp add: sgn_mult_abs)
+ moreover define u where "u = nat \<bar>r\<bar>"
+ ultimately have "r = sgn l * int u"
+ by simp
+ with l \<open>n > 0\<close> have r: "r = sgn s * int u"
+ by (simp add: sgn_mult)
+ assume "euclidean_size r < euclidean_size l"
+ with l r \<open>s \<noteq> 0\<close> have "u < n"
+ by (simp add: abs_mult)
+ show "(q * l + r) div l = q"
+ proof (cases "q = 0 \<or> r = 0")
+ case True
+ then show ?thesis
+ proof
+ assume "q = 0"
+ then show ?thesis
+ using l r \<open>u < n\<close> by (simp add: divide_int_unfold)
+ next
+ assume "r = 0"
+ from \<open>r = 0\<close> have *: "q * l + r = sgn (t * s) * int (n * m)"
+ using q l by (simp add: ac_simps sgn_mult)
+ from \<open>s \<noteq> 0\<close> \<open>n > 0\<close> show ?thesis
+ by (simp only: *, simp only: q l divide_int_unfold)
+ (auto simp add: sgn_mult sgn_0_0 sgn_1_pos)
+ qed
+ next
+ case False
+ with q r have "t \<noteq> 0" and "m > 0" and "s \<noteq> 0" and "u > 0"
+ by (simp_all add: sgn_0_0)
+ moreover from \<open>0 < m\<close> \<open>u < n\<close> have "u \<le> m * n"
+ using mult_le_less_imp_less [of 1 m u n] by simp
+ ultimately have *: "q * l + r = sgn (s * t)
+ * int (if t < 0 then m * n - u else m * n + u)"
+ using l q r
+ by (simp add: sgn_mult algebra_simps of_nat_diff)
+ have "(m * n - u) div n = m - 1" if "u > 0"
+ using \<open>0 < m\<close> \<open>u < n\<close> that
+ by (auto intro: div_nat_eqI simp add: algebra_simps)
+ moreover have "n dvd m * n - u \<longleftrightarrow> n dvd u"
+ using \<open>u \<le> m * n\<close> dvd_diffD1 [of n "m * n" u]
+ by auto
+ ultimately show ?thesis
+ using \<open>s \<noteq> 0\<close> \<open>m > 0\<close> \<open>u > 0\<close> \<open>u < n\<close> \<open>u \<le> m * n\<close>
+ by (simp only: *, simp only: l q divide_int_unfold)
+ (auto simp add: sgn_mult sgn_0_0 sgn_1_pos algebra_simps dest: dvd_imp_le)
+ qed
+qed (use mult_le_mono2 [of 1] in \<open>auto simp add: abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\<close>)
+
+end
+
+lemma pos_mod_bound [simp]:
+ "k mod l < l" if "l > 0" for k l :: int
+proof -
+ obtain m and s where "k = sgn s * int m"
+ by (blast intro: int_sgnE elim: that)
+ moreover from that obtain n where "l = sgn 1 * int n"
+ by (cases l) auto
+ ultimately show ?thesis
+ using that by (simp only: modulo_int_unfold)
+ (simp add: mod_greater_zero_iff_not_dvd)
+qed
+
+lemma pos_mod_sign [simp]:
+ "0 \<le> k mod l" if "l > 0" for k l :: int
+proof -
+ obtain m and s where "k = sgn s * int m"
+ by (blast intro: int_sgnE elim: that)
+ moreover from that obtain n where "l = sgn 1 * int n"
+ by (cases l) auto
+ ultimately show ?thesis
+ using that by (simp only: modulo_int_unfold) simp
+qed
+
+
+subsection \<open>Code generation\<close>
+
+code_identifier
+ code_module Euclidean_Division \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
+
+end
--- a/src/HOL/Factorial.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Factorial.thy Mon Oct 09 22:08:05 2017 +0200
@@ -130,11 +130,11 @@
lemma fact_ge_self: "fact n \<ge> n"
by (cases "n = 0") (simp_all add: dvd_imp_le dvd_fact)
-lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd (fact m :: 'a::{semiring_div,linordered_semidom})"
+lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd (fact m :: 'a::linordered_semidom)"
by (induct m) (auto simp: le_Suc_eq)
-lemma fact_mod: "m \<le> n \<Longrightarrow> fact n mod (fact m :: 'a::{semiring_div,linordered_semidom}) = 0"
- by (auto simp add: fact_dvd)
+lemma fact_mod: "m \<le> n \<Longrightarrow> fact n mod (fact m :: 'a::{semidom_modulo, linordered_semidom}) = 0"
+ by (simp add: mod_eq_0_iff_dvd fact_dvd)
lemma fact_div_fact:
assumes "m \<ge> n"
--- a/src/HOL/GCD.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/GCD.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1510,21 +1510,6 @@
lemma Gcd_2 [simp]: "Gcd {a, b} = gcd a b"
by simp
-
-definition pairwise_coprime
- where "pairwise_coprime A = (\<forall>x y. x \<in> A \<and> y \<in> A \<and> x \<noteq> y \<longrightarrow> coprime x y)"
-
-lemma pairwise_coprimeI [intro?]:
- "(\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<noteq> y \<Longrightarrow> coprime x y) \<Longrightarrow> pairwise_coprime A"
- by (simp add: pairwise_coprime_def)
-
-lemma pairwise_coprimeD:
- "pairwise_coprime A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<noteq> y \<Longrightarrow> coprime x y"
- by (simp add: pairwise_coprime_def)
-
-lemma pairwise_coprime_subset: "pairwise_coprime A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> pairwise_coprime B"
- by (force simp: pairwise_coprime_def)
-
end
@@ -1995,11 +1980,7 @@
lemma gcd_mult_distrib_int: "\<bar>k\<bar> * gcd m n = gcd (k * m) (k * n)"
for k m n :: int
- apply (subst (1 2) gcd_abs_int)
- apply (subst (1 2) abs_mult)
- apply (rule gcd_mult_distrib_nat [transferred])
- apply auto
- done
+ by (simp add: gcd_int_def abs_mult nat_mult_distrib gcd_mult_distrib_nat [symmetric])
lemma coprime_crossproduct_nat:
fixes a b c d :: nat
@@ -2036,11 +2017,11 @@
lemma gcd_non_0_int: "y > 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
for x y :: int
apply (frule_tac b = y and a = x in pos_mod_sign)
- apply (simp del: pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib)
+ apply (simp del: Euclidean_Division.pos_mod_sign add: gcd_int_def abs_if nat_mod_distrib)
apply (auto simp add: gcd_non_0_nat nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
apply (frule_tac a = x in pos_mod_bound)
apply (subst (1 2) gcd.commute)
- apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2_nat nat_le_eq_zle)
+ apply (simp del: Euclidean_Division.pos_mod_bound add: nat_diff_distrib gcd_diff2_nat nat_le_eq_zle)
done
lemma gcd_red_int: "gcd x y = gcd y (x mod y)"
@@ -2480,10 +2461,7 @@
lemma lcm_pos_int: "m \<noteq> 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> lcm m n > 0"
for m n :: int
- apply (subst lcm_abs_int)
- apply (rule lcm_pos_nat [transferred])
- apply auto
- done
+ by (simp add: lcm_int_def lcm_pos_nat)
lemma dvd_pos_nat: "n > 0 \<Longrightarrow> m dvd n \<Longrightarrow> m > 0" (* FIXME move *)
for m n :: nat
--- a/src/HOL/Groups_Big.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Groups_Big.thy Mon Oct 09 22:08:05 2017 +0200
@@ -425,14 +425,14 @@
by (simp add: H)
qed
-lemma commute: "F (\<lambda>i. F (g i) B) A = F (\<lambda>j. F (\<lambda>i. g i j) A) B"
+lemma swap: "F (\<lambda>i. F (g i) B) A = F (\<lambda>j. F (\<lambda>i. g i j) A) B"
unfolding cartesian_product
by (rule reindex_bij_witness [where i = "\<lambda>(i, j). (j, i)" and j = "\<lambda>(i, j). (j, i)"]) auto
-lemma commute_restrict:
+lemma swap_restrict:
"finite A \<Longrightarrow> finite B \<Longrightarrow>
F (\<lambda>x. F (g x) {y. y \<in> B \<and> R x y}) A = F (\<lambda>y. F (\<lambda>x. g x y) {x. x \<in> A \<and> R x y}) B"
- by (simp add: inter_filter) (rule commute)
+ by (simp add: inter_filter) (rule swap)
lemma Plus:
fixes A :: "'b set" and B :: "'c set"
@@ -540,7 +540,7 @@
then have "sum g S = sum (\<lambda>x. sum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
by simp
also have "\<dots> = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
- by (rule sum.commute_restrict [OF fin finite_imageI [OF fin]])
+ by (rule sum.swap_restrict [OF fin finite_imageI [OF fin]])
finally show ?thesis .
qed
@@ -843,7 +843,7 @@
lemma sum_product:
fixes f :: "'a \<Rightarrow> 'b::semiring_0"
shows "sum f A * sum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
- by (simp add: sum_distrib_left sum_distrib_right) (rule sum.commute)
+ by (simp add: sum_distrib_left sum_distrib_right) (rule sum.swap)
lemma sum_mult_sum_if_inj:
fixes f :: "'a \<Rightarrow> 'b::semiring_0"
@@ -1021,7 +1021,7 @@
have "?l = sum (\<lambda>i. sum (\<lambda>x.1) {j\<in>t. R i j}) s"
by auto
also have "\<dots> = ?r"
- unfolding sum.commute_restrict [OF assms(1-2)]
+ unfolding sum.swap_restrict [OF assms(1-2)]
using assms(3) by auto
finally show ?thesis .
qed
--- a/src/HOL/Inequalities.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Inequalities.thy Mon Oct 09 22:08:05 2017 +0200
@@ -59,7 +59,7 @@
let ?S = "(\<Sum>j=0..<n. (\<Sum>k=0..<n. (a j - a k) * (b j - b k)))"
have "2 * (of_nat n * (\<Sum>j=0..<n. (a j * b j)) - (\<Sum>j=0..<n. b j) * (\<Sum>k=0..<n. a k)) = ?S"
by (simp only: one_add_one[symmetric] algebra_simps)
- (simp add: algebra_simps sum_subtractf sum.distrib sum.commute[of "\<lambda>i j. a i * b j"] sum_distrib_left)
+ (simp add: algebra_simps sum_subtractf sum.distrib sum.swap[of "\<lambda>i j. a i * b j"] sum_distrib_left)
also
{ fix i j::nat assume "i<n" "j<n"
hence "a i - a j \<le> 0 \<and> b i - b j \<ge> 0 \<or> a i - a j \<ge> 0 \<and> b i - b j \<le> 0"
--- a/src/HOL/Int.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Int.thy Mon Oct 09 22:08:05 2017 +0200
@@ -255,6 +255,10 @@
lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n"
by (induct n) simp_all
+lemma of_int_of_bool [simp]:
+ "of_int (of_bool P) = of_bool P"
+ by auto
+
end
context ring_char_0
@@ -548,6 +552,10 @@
lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
+lemma nat_of_bool [simp]:
+ "nat (of_bool P) = of_bool P"
+ by auto
+
text \<open>For termination proofs:\<close>
lemma measure_function_int[measure_function]: "is_measure (nat \<circ> abs)" ..
@@ -668,6 +676,31 @@
\<comment> \<open>Unfold all \<open>let\<close>s involving constants\<close>
by (fact Let_neg_numeral) \<comment> \<open>FIXME drop\<close>
+lemma sgn_mult_dvd_iff [simp]:
+ "sgn r * l dvd k \<longleftrightarrow> l dvd k \<and> (r = 0 \<longrightarrow> k = 0)" for k l r :: int
+ by (cases r rule: int_cases3) auto
+
+lemma mult_sgn_dvd_iff [simp]:
+ "l * sgn r dvd k \<longleftrightarrow> l dvd k \<and> (r = 0 \<longrightarrow> k = 0)" for k l r :: int
+ using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps)
+
+lemma dvd_sgn_mult_iff [simp]:
+ "l dvd sgn r * k \<longleftrightarrow> l dvd k \<or> r = 0" for k l r :: int
+ by (cases r rule: int_cases3) simp_all
+
+lemma dvd_mult_sgn_iff [simp]:
+ "l dvd k * sgn r \<longleftrightarrow> l dvd k \<or> r = 0" for k l r :: int
+ using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps)
+
+lemma int_sgnE:
+ fixes k :: int
+ obtains n and l where "k = sgn l * int n"
+proof -
+ have "k = sgn k * int (nat \<bar>k\<bar>)"
+ by (simp add: sgn_mult_abs)
+ then show ?thesis ..
+qed
+
text \<open>Unfold \<open>min\<close> and \<open>max\<close> on numerals.\<close>
lemmas max_number_of [simp] =
@@ -1629,19 +1662,10 @@
shows "a dvd (x + t) \<longleftrightarrow> a dvd ((x + c * d) + t)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof -
- from assms obtain k where "d = a * k" by (rule dvdE)
- show ?thesis
- proof
- assume ?lhs
- then obtain l where "x + t = a * l" by (rule dvdE)
- then have "x = a * l - t" by simp
- with \<open>d = a * k\<close> show ?rhs by simp
- next
- assume ?rhs
- then obtain l where "x + c * d + t = a * l" by (rule dvdE)
- then have "x = a * l - c * d - t" by simp
- with \<open>d = a * k\<close> show ?lhs by simp
- qed
+ from assms have "a dvd (x + t) \<longleftrightarrow> a dvd ((x + t) + c * d)"
+ by (simp add: dvd_add_left_iff)
+ then show ?thesis
+ by (simp add: ac_simps)
qed
--- a/src/HOL/Library/Code_Target_Nat.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Library/Code_Target_Nat.thy Mon Oct 09 22:08:05 2017 +0200
@@ -95,7 +95,7 @@
lemma [code]:
"Divides.divmod_nat m n = (m div n, m mod n)"
- by (fact divmod_nat_div_mod)
+ by (fact divmod_nat_def)
lemma [code]:
"divmod m n = map_prod nat_of_integer nat_of_integer (divmod m n)"
@@ -130,7 +130,7 @@
proof -
from div_mult_mod_eq have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp
show ?thesis
- by (simp add: Let_def divmod_nat_div_mod of_nat_add [symmetric])
+ by (simp add: Let_def divmod_nat_def of_nat_add [symmetric])
(simp add: * mult.commute of_nat_mult add.commute)
qed
--- a/src/HOL/Library/Function_Growth.thy Mon Oct 09 22:03:05 2017 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,454 +0,0 @@
-
-(* Author: Florian Haftmann, TU Muenchen *)
-
-section \<open>Comparing growth of functions on natural numbers by a preorder relation\<close>
-
-theory Function_Growth
-imports Main Preorder Discrete
-begin
-
-(* FIXME move *)
-
-context linorder
-begin
-
-lemma mono_invE:
- fixes f :: "'a \<Rightarrow> 'b::order"
- assumes "mono f"
- assumes "f x < f y"
- obtains "x < y"
-proof
- show "x < y"
- proof (rule ccontr)
- assume "\<not> x < y"
- then have "y \<le> x" by simp
- with \<open>mono f\<close> obtain "f y \<le> f x" by (rule monoE)
- with \<open>f x < f y\<close> show False by simp
- qed
-qed
-
-end
-
-lemma (in semidom_divide) power_diff:
- fixes a :: 'a
- assumes "a \<noteq> 0"
- assumes "m \<ge> n"
- shows "a ^ (m - n) = (a ^ m) div (a ^ n)"
-proof -
- define q where "q = m - n"
- with assms have "m = q + n" by (simp add: q_def)
- with q_def show ?thesis using \<open>a \<noteq> 0\<close> by (simp add: power_add)
-qed
-
-
-subsection \<open>Motivation\<close>
-
-text \<open>
- When comparing growth of functions in computer science, it is common to adhere
- on Landau Symbols (``O-Notation''). However these come at the cost of notational
- oddities, particularly writing \<open>f = O(g)\<close> for \<open>f \<in> O(g)\<close> etc.
-
- Here we suggest a different way, following Hardy (G.~H.~Hardy and J.~E.~Littlewood,
- Some problems of Diophantine approximation, Acta Mathematica 37 (1914), p.~225).
- We establish a quasi order relation \<open>\<lesssim>\<close> on functions such that
- \<open>f \<lesssim> g \<longleftrightarrow> f \<in> O(g)\<close>. From a didactic point of view, this does not only
- avoid the notational oddities mentioned above but also emphasizes the key insight
- of a growth hierarchy of functions:
- \<open>(\<lambda>n. 0) \<lesssim> (\<lambda>n. k) \<lesssim> Discrete.log \<lesssim> Discrete.sqrt \<lesssim> id \<lesssim> \<dots>\<close>.
-\<close>
-
-subsection \<open>Model\<close>
-
-text \<open>
- Our growth functions are of type \<open>\<nat> \<Rightarrow> \<nat>\<close>. This is different
- to the usual conventions for Landau symbols for which \<open>\<real> \<Rightarrow> \<real>\<close>
- would be appropriate, but we argue that \<open>\<real> \<Rightarrow> \<real>\<close> is more
- appropriate for analysis, whereas our setting is discrete.
-
- Note that we also restrict the additional coefficients to \<open>\<nat>\<close>, something
- we discuss at the particular definitions.
-\<close>
-
-subsection \<open>The \<open>\<lesssim>\<close> relation\<close>
-
-definition less_eq_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<lesssim>" 50)
-where
- "f \<lesssim> g \<longleftrightarrow> (\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m)"
-
-text \<open>
- This yields \<open>f \<lesssim> g \<longleftrightarrow> f \<in> O(g)\<close>. Note that \<open>c\<close> is restricted to
- \<open>\<nat>\<close>. This does not pose any problems since if \<open>f \<in> O(g)\<close> holds for
- a \<open>c \<in> \<real>\<close>, it also holds for \<open>\<lceil>c\<rceil> \<in> \<nat>\<close> by transitivity.
-\<close>
-
-lemma less_eq_funI [intro?]:
- assumes "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m"
- shows "f \<lesssim> g"
- unfolding less_eq_fun_def by (rule assms)
-
-lemma not_less_eq_funI:
- assumes "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * g m < f m"
- shows "\<not> f \<lesssim> g"
- using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
-
-lemma less_eq_funE [elim?]:
- assumes "f \<lesssim> g"
- obtains n c where "c > 0" and "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
- using assms unfolding less_eq_fun_def by blast
-
-lemma not_less_eq_funE:
- assumes "\<not> f \<lesssim> g" and "c > 0"
- obtains m where "m > n" and "c * g m < f m"
- using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
-
-
-subsection \<open>The \<open>\<approx>\<close> relation, the equivalence relation induced by \<open>\<lesssim>\<close>\<close>
-
-definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<cong>" 50)
-where
- "f \<cong> g \<longleftrightarrow>
- (\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m)"
-
-text \<open>
- This yields \<open>f \<cong> g \<longleftrightarrow> f \<in> \<Theta>(g)\<close>. Concerning \<open>c\<^sub>1\<close> and \<open>c\<^sub>2\<close>
- restricted to @{typ nat}, see note above on \<open>(\<lesssim>)\<close>.
-\<close>
-
-lemma equiv_funI:
- assumes "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
- shows "f \<cong> g"
- unfolding equiv_fun_def by (rule assms)
-
-lemma not_equiv_funI:
- assumes "\<And>c\<^sub>1 c\<^sub>2 n. c\<^sub>1 > 0 \<Longrightarrow> c\<^sub>2 > 0 \<Longrightarrow>
- \<exists>m>n. c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
- shows "\<not> f \<cong> g"
- using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
-
-lemma equiv_funE:
- assumes "f \<cong> g"
- obtains n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
- and "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
- using assms unfolding equiv_fun_def by blast
-
-lemma not_equiv_funE:
- fixes n c\<^sub>1 c\<^sub>2
- assumes "\<not> f \<cong> g" and "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
- obtains m where "m > n"
- and "c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
- using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
-
-
-subsection \<open>The \<open>\<prec>\<close> relation, the strict part of \<open>\<lesssim>\<close>\<close>
-
-definition less_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<prec>" 50)
-where
- "f \<prec> g \<longleftrightarrow> f \<lesssim> g \<and> \<not> g \<lesssim> f"
-
-lemma less_funI:
- assumes "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m"
- and "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * f m < g m"
- shows "f \<prec> g"
- using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
-
-lemma not_less_funI:
- assumes "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * g m < f m"
- and "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m"
- shows "\<not> f \<prec> g"
- using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
-
-lemma less_funE [elim?]:
- assumes "f \<prec> g"
- obtains n c where "c > 0" and "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
- and "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * f m < g m"
-proof -
- from assms have "f \<lesssim> g" and "\<not> g \<lesssim> f" by (simp_all add: less_fun_def)
- from \<open>f \<lesssim> g\<close> obtain n c where *:"c > 0" "m > n \<Longrightarrow> f m \<le> c * g m" for m
- by (rule less_eq_funE) blast
- { fix c n :: nat
- assume "c > 0"
- with \<open>\<not> g \<lesssim> f\<close> obtain m where "m > n" "c * f m < g m"
- by (rule not_less_eq_funE) blast
- then have **: "\<exists>m>n. c * f m < g m" by blast
- } note ** = this
- from * ** show thesis by (rule that)
-qed
-
-lemma not_less_funE:
- assumes "\<not> f \<prec> g" and "c > 0"
- obtains m where "m > n" and "c * g m < f m"
- | d q where "\<And>m. d > 0 \<Longrightarrow> m > q \<Longrightarrow> g q \<le> d * f q"
- using assms unfolding less_fun_def linorder_not_less [symmetric] by blast
-
-text \<open>
- I did not find a proof for \<open>f \<prec> g \<longleftrightarrow> f \<in> o(g)\<close>. Maybe this only
- holds if \<open>f\<close> and/or \<open>g\<close> are of a certain class of functions.
- However \<open>f \<in> o(g) \<longrightarrow> f \<prec> g\<close> is provable, and this yields a
- handy introduction rule.
-
- Note that D. Knuth ignores \<open>o\<close> altogether. So what \dots
-
- Something still has to be said about the coefficient \<open>c\<close> in
- the definition of \<open>(\<prec>)\<close>. In the typical definition of \<open>o\<close>,
- it occurs on the \emph{right} hand side of the \<open>(>)\<close>. The reason
- is that the situation is dual to the definition of \<open>O\<close>: the definition
- works since \<open>c\<close> may become arbitrary small. Since this is not possible
- within @{term \<nat>}, we push the coefficient to the left hand side instead such
- that it may become arbitrary big instead.
-\<close>
-
-lemma less_fun_strongI:
- assumes "\<And>c. c > 0 \<Longrightarrow> \<exists>n. \<forall>m>n. c * f m < g m"
- shows "f \<prec> g"
-proof (rule less_funI)
- have "1 > (0::nat)" by simp
- with assms [OF this] obtain n where *: "m > n \<Longrightarrow> 1 * f m < g m" for m
- by blast
- have "\<forall>m>n. f m \<le> 1 * g m"
- proof (rule allI, rule impI)
- fix m
- assume "m > n"
- with * have "1 * f m < g m" by simp
- then show "f m \<le> 1 * g m" by simp
- qed
- with \<open>1 > 0\<close> show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
- fix c n :: nat
- assume "c > 0"
- with assms obtain q where "m > q \<Longrightarrow> c * f m < g m" for m by blast
- then have "c * f (Suc (q + n)) < g (Suc (q + n))" by simp
- moreover have "Suc (q + n) > n" by simp
- ultimately show "\<exists>m>n. c * f m < g m" by blast
-qed
-
-
-subsection \<open>\<open>\<lesssim>\<close> is a preorder\<close>
-
-text \<open>This yields all lemmas relating \<open>\<lesssim>\<close>, \<open>\<prec>\<close> and \<open>\<cong>\<close>.\<close>
-
-interpretation fun_order: preorder_equiv less_eq_fun less_fun
- rewrites "fun_order.equiv = equiv_fun"
-proof -
- interpret preorder: preorder_equiv less_eq_fun less_fun
- proof
- fix f g h
- show "f \<lesssim> f"
- proof
- have "\<exists>n. \<forall>m>n. f m \<le> 1 * f m" by auto
- then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * f m" by blast
- qed
- show "f \<prec> g \<longleftrightarrow> f \<lesssim> g \<and> \<not> g \<lesssim> f"
- by (fact less_fun_def)
- assume "f \<lesssim> g" and "g \<lesssim> h"
- show "f \<lesssim> h"
- proof
- from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1
- where "c\<^sub>1 > 0" and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m"
- by rule blast
- from \<open>g \<lesssim> h\<close> obtain n\<^sub>2 c\<^sub>2
- where "c\<^sub>2 > 0" and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * h m"
- by rule blast
- have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m"
- proof (rule allI, rule impI)
- fix m
- assume Q: "m > max n\<^sub>1 n\<^sub>2"
- from P\<^sub>1 Q have *: "f m \<le> c\<^sub>1 * g m" by simp
- from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * h m" by simp
- with \<open>c\<^sub>1 > 0\<close> have "c\<^sub>1 * g m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by simp
- with * show "f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by (rule order_trans)
- qed
- then have "\<exists>n. \<forall>m>n. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by rule
- moreover from \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "c\<^sub>1 * c\<^sub>2 > 0" by simp
- ultimately show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * h m" by blast
- qed
- qed
- from preorder.preorder_equiv_axioms show "class.preorder_equiv less_eq_fun less_fun" .
- show "preorder_equiv.equiv less_eq_fun = equiv_fun"
- proof (rule ext, rule ext, unfold preorder.equiv_def)
- fix f g
- show "f \<lesssim> g \<and> g \<lesssim> f \<longleftrightarrow> f \<cong> g"
- proof
- assume "f \<cong> g"
- then obtain n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
- and *: "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
- by (rule equiv_funE) blast
- have "\<forall>m>n. f m \<le> c\<^sub>1 * g m"
- proof (rule allI, rule impI)
- fix m
- assume "m > n"
- with * show "f m \<le> c\<^sub>1 * g m" by simp
- qed
- with \<open>c\<^sub>1 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
- then have "f \<lesssim> g" ..
- have "\<forall>m>n. g m \<le> c\<^sub>2 * f m"
- proof (rule allI, rule impI)
- fix m
- assume "m > n"
- with * show "g m \<le> c\<^sub>2 * f m" by simp
- qed
- with \<open>c\<^sub>2 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
- then have "g \<lesssim> f" ..
- from \<open>f \<lesssim> g\<close> and \<open>g \<lesssim> f\<close> show "f \<lesssim> g \<and> g \<lesssim> f" ..
- next
- assume "f \<lesssim> g \<and> g \<lesssim> f"
- then have "f \<lesssim> g" and "g \<lesssim> f" by auto
- from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1 where "c\<^sub>1 > 0"
- and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m" by rule blast
- from \<open>g \<lesssim> f\<close> obtain n\<^sub>2 c\<^sub>2 where "c\<^sub>2 > 0"
- and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * f m" by rule blast
- have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
- proof (rule allI, rule impI)
- fix m
- assume Q: "m > max n\<^sub>1 n\<^sub>2"
- from P\<^sub>1 Q have "f m \<le> c\<^sub>1 * g m" by simp
- moreover from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * f m" by simp
- ultimately show "f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" ..
- qed
- with \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n.
- \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" by blast
- then show "f \<cong> g" by (rule equiv_funI)
- qed
- qed
-qed
-
-declare fun_order.antisym [intro?]
-
-
-subsection \<open>Simple examples\<close>
-
-text \<open>
- Most of these are left as constructive exercises for the reader. Note that additional
- preconditions to the functions may be necessary. The list here is by no means to be
- intended as complete construction set for typical functions, here surely something
- has to be added yet.
-\<close>
-
-text \<open>@{prop "(\<lambda>n. f n + k) \<cong> f"}\<close>
-
-lemma equiv_fun_mono_const:
- assumes "mono f" and "\<exists>n. f n > 0"
- shows "(\<lambda>n. f n + k) \<cong> f"
-proof (cases "k = 0")
- case True then show ?thesis by simp
-next
- case False
- show ?thesis
- proof
- show "(\<lambda>n. f n + k) \<lesssim> f"
- proof
- from \<open>\<exists>n. f n > 0\<close> obtain n where "f n > 0" ..
- have "\<forall>m>n. f m + k \<le> Suc k * f m"
- proof (rule allI, rule impI)
- fix m
- assume "n < m"
- with \<open>mono f\<close> have "f n \<le> f m"
- using less_imp_le_nat monoE by blast
- with \<open>0 < f n\<close> have "0 < f m" by auto
- then obtain l where "f m = Suc l" by (cases "f m") simp_all
- then show "f m + k \<le> Suc k * f m" by simp
- qed
- then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m + k \<le> c * f m" by blast
- qed
- show "f \<lesssim> (\<lambda>n. f n + k)"
- proof
- have "f m \<le> 1 * (f m + k)" for m by simp
- then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * (f m + k)" by blast
- qed
- qed
-qed
-
-lemma
- assumes "strict_mono f"
- shows "(\<lambda>n. f n + k) \<cong> f"
-proof (rule equiv_fun_mono_const)
- from assms show "mono f" by (rule strict_mono_mono)
- show "\<exists>n. 0 < f n"
- proof (rule ccontr)
- assume "\<not> (\<exists>n. 0 < f n)"
- then have "\<And>n. f n = 0" by simp
- then have "f 0 = f 1" by simp
- moreover from \<open>strict_mono f\<close> have "f 0 < f 1"
- by (simp add: strict_mono_def)
- ultimately show False by simp
- qed
-qed
-
-lemma
- "(\<lambda>n. Suc k * f n) \<cong> f"
-proof
- show "(\<lambda>n. Suc k * f n) \<lesssim> f"
- proof
- have "Suc k * f m \<le> Suc k * f m" for m by simp
- then show "\<exists>c>0. \<exists>n. \<forall>m>n. Suc k * f m \<le> c * f m" by blast
- qed
- show "f \<lesssim> (\<lambda>n. Suc k * f n)"
- proof
- have "f m \<le> 1 * (Suc k * f m)" for m by simp
- then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * (Suc k * f m)" by blast
- qed
-qed
-
-lemma
- "f \<lesssim> (\<lambda>n. f n + g n)"
- by rule auto
-
-lemma
- "(\<lambda>_. 0) \<prec> (\<lambda>n. Suc k)"
- by (rule less_fun_strongI) auto
-
-lemma
- "(\<lambda>_. k) \<prec> Discrete.log"
-proof (rule less_fun_strongI)
- fix c :: nat
- have "\<forall>m>2 ^ (Suc (c * k)). c * k < Discrete.log m"
- proof (rule allI, rule impI)
- fix m :: nat
- assume "2 ^ Suc (c * k) < m"
- then have "2 ^ Suc (c * k) \<le> m" by simp
- with log_mono have "Discrete.log (2 ^ (Suc (c * k))) \<le> Discrete.log m"
- by (blast dest: monoD)
- moreover have "c * k < Discrete.log (2 ^ (Suc (c * k)))" by simp
- ultimately show "c * k < Discrete.log m" by auto
- qed
- then show "\<exists>n. \<forall>m>n. c * k < Discrete.log m" ..
-qed
-
-(*lemma
- "Discrete.log \<prec> Discrete.sqrt"
-proof (rule less_fun_strongI)*)
-text \<open>@{prop "Discrete.log \<prec> Discrete.sqrt"}\<close>
-
-lemma
- "Discrete.sqrt \<prec> id"
-proof (rule less_fun_strongI)
- fix c :: nat
- assume "0 < c"
- have "\<forall>m>(Suc c)\<^sup>2. c * Discrete.sqrt m < id m"
- proof (rule allI, rule impI)
- fix m
- assume "(Suc c)\<^sup>2 < m"
- then have "(Suc c)\<^sup>2 \<le> m" by simp
- with mono_sqrt have "Discrete.sqrt ((Suc c)\<^sup>2) \<le> Discrete.sqrt m" by (rule monoE)
- then have "Suc c \<le> Discrete.sqrt m" by simp
- then have "c < Discrete.sqrt m" by simp
- moreover from \<open>(Suc c)\<^sup>2 < m\<close> have "Discrete.sqrt m > 0" by simp
- ultimately have "c * Discrete.sqrt m < Discrete.sqrt m * Discrete.sqrt m" by simp
- also have "\<dots> \<le> m" by (simp add: power2_eq_square [symmetric])
- finally show "c * Discrete.sqrt m < id m" by simp
- qed
- then show "\<exists>n. \<forall>m>n. c * Discrete.sqrt m < id m" ..
-qed
-
-lemma
- "id \<prec> (\<lambda>n. n\<^sup>2)"
- by (rule less_fun_strongI) (auto simp add: power2_eq_square)
-
-lemma
- "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. n ^ Suc k)"
- by (rule less_fun_strongI) auto
-
-(*lemma
- "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"
-proof (rule less_fun_strongI)*)
-text \<open>@{prop "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"}\<close>
-
-end
--- a/src/HOL/Library/Groups_Big_Fun.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy Mon Oct 09 22:08:05 2017 +0200
@@ -106,7 +106,7 @@
by (simp add: expand_superset [of "{a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}"] F.distrib)
qed
-lemma commute:
+lemma swap:
assumes "finite C"
assumes subset: "{a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C")
shows "G (\<lambda>a. G (g a)) = G (\<lambda>b. G (\<lambda>a. g a b))"
@@ -122,7 +122,7 @@
"{a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {a. \<exists>b. g a b \<noteq> \<^bold>1}"
"{a. F.F (\<lambda>aa. g aa a) {a. \<exists>b. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {b. \<exists>a. g a b \<noteq> \<^bold>1}"
by (auto elim: F.not_neutral_contains_not_neutral)
- from F.commute have
+ from F.swap have
"F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1} =
F.F (\<lambda>b. F.F (\<lambda>a. g a b) {a. \<exists>b. g a b \<noteq> \<^bold>1}) {b. \<exists>a. g a b \<noteq> \<^bold>1}" .
with subsets fins have "G (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) =
--- a/src/HOL/Library/Library.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Library/Library.thy Mon Oct 09 22:08:05 2017 +0200
@@ -17,6 +17,7 @@
Countable_Set_Type
Debug
Diagonal_Subsequence
+ Discrete
Disjoint_Sets
Dlist
Extended
@@ -28,7 +29,6 @@
FSet
FuncSet
Function_Division
- Function_Growth
Fun_Lexorder
Going_To_Filter
Groups_Big_Fun
--- a/src/HOL/Library/RBT_Impl.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Library/RBT_Impl.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1169,7 +1169,7 @@
apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
apfst (Branch B t1 k v) (rbtreeify_f n' kvs'))"
-by (subst rbtreeify_f.simps) (simp only: Let_def divmod_nat_div_mod prod.case)
+by (subst rbtreeify_f.simps) (simp only: Let_def divmod_nat_def prod.case)
lemma rbtreeify_g_code [code]:
"rbtreeify_g n kvs =
@@ -1180,7 +1180,7 @@
apfst (Branch B t1 k v) (rbtreeify_g n' kvs')
else case rbtreeify_f n' kvs of (t1, (k, v) # kvs') \<Rightarrow>
apfst (Branch B t1 k v) (rbtreeify_g n' kvs'))"
-by(subst rbtreeify_g.simps)(simp only: Let_def divmod_nat_div_mod prod.case)
+by(subst rbtreeify_g.simps)(simp only: Let_def divmod_nat_def prod.case)
lemma Suc_double_half: "Suc (2 * n) div 2 = n"
by simp
--- a/src/HOL/Limits.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Limits.thy Mon Oct 09 22:08:05 2017 +0200
@@ -820,6 +820,11 @@
lemmas tendsto_mult_right_zero =
bounded_bilinear.tendsto_right_zero [OF bounded_bilinear_mult]
+lemma tendsto_divide_zero:
+ fixes c :: "'a::real_normed_field"
+ shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x / c) \<longlongrightarrow> 0) F"
+ by (cases "c=0") (simp_all add: divide_inverse tendsto_mult_left_zero)
+
lemma tendsto_power [tendsto_intros]: "(f \<longlongrightarrow> a) F \<Longrightarrow> ((\<lambda>x. f x ^ n) \<longlongrightarrow> a ^ n) F"
for f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
by (induct n) (simp_all add: tendsto_mult)
--- a/src/HOL/Main.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Main.thy Mon Oct 09 22:08:05 2017 +0200
@@ -6,7 +6,16 @@
\<close>
theory Main
-imports Predicate_Compile Quickcheck_Narrowing Extraction Nunchaku BNF_Greatest_Fixpoint Filter Conditionally_Complete_Lattices Binomial GCD
+ imports
+ Predicate_Compile
+ Quickcheck_Narrowing
+ Extraction
+ Nunchaku
+ BNF_Greatest_Fixpoint Filter
+ Conditionally_Complete_Lattices
+ Binomial
+ GCD
+ Nat_Transfer
begin
text \<open>
--- a/src/HOL/Nat.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Nat.thy Mon Oct 09 22:08:05 2017 +0200
@@ -10,10 +10,6 @@
imports Inductive Typedef Fun Rings
begin
-named_theorems arith "arith facts -- only ground formulas"
-ML_file "Tools/arith_data.ML"
-
-
subsection \<open>Type \<open>ind\<close>\<close>
typedecl ind
@@ -1680,6 +1676,10 @@
by (simp add: add.commute)
qed
+lemma of_nat_of_bool [simp]:
+ "of_nat (of_bool P) = of_bool P"
+ by auto
+
end
declare of_nat_code [code]
@@ -1768,6 +1768,10 @@
lemma abs_of_nat [simp]: "\<bar>of_nat n\<bar> = of_nat n"
unfolding abs_if by auto
+lemma sgn_of_nat [simp]:
+ "sgn (of_nat n) = of_bool (n > 0)"
+ by simp
+
end
lemma of_nat_id [simp]: "of_nat n = n"
--- a/src/HOL/Nat_Transfer.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Nat_Transfer.thy Mon Oct 09 22:08:05 2017 +0200
@@ -3,7 +3,7 @@
section \<open>Generic transfer machinery; specific transfer from nats to ints and back.\<close>
theory Nat_Transfer
-imports Int
+imports Int Divides
begin
subsection \<open>Generic transfer machinery\<close>
@@ -21,7 +21,8 @@
text \<open>set up transfer direction\<close>
-lemma transfer_morphism_nat_int: "transfer_morphism nat (op <= (0::int))" ..
+lemma transfer_morphism_nat_int [no_atp]:
+ "transfer_morphism nat (op <= (0::int))" ..
declare transfer_morphism_nat_int [transfer add
mode: manual
@@ -31,7 +32,7 @@
text \<open>basic functions and relations\<close>
-lemma transfer_nat_int_numerals [transfer key: transfer_morphism_nat_int]:
+lemma transfer_nat_int_numerals [no_atp, transfer key: transfer_morphism_nat_int]:
"(0::nat) = nat 0"
"(1::nat) = nat 1"
"(2::nat) = nat 2"
@@ -46,15 +47,17 @@
lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y"
by (simp add: tsub_def)
-lemma transfer_nat_int_functions [transfer key: transfer_morphism_nat_int]:
+lemma transfer_nat_int_functions [no_atp, transfer key: transfer_morphism_nat_int]:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)"
"(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
by (auto simp add: eq_nat_nat_iff nat_mult_distrib
- nat_power_eq tsub_def)
+ nat_power_eq tsub_def nat_div_distrib nat_mod_distrib)
-lemma transfer_nat_int_function_closures [transfer key: transfer_morphism_nat_int]:
+lemma transfer_nat_int_function_closures [no_atp, transfer key: transfer_morphism_nat_int]:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0"
@@ -64,9 +67,16 @@
"(2::int) >= 0"
"(3::int) >= 0"
"int z >= 0"
- by (auto simp add: zero_le_mult_iff tsub_def)
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
+ "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
+ apply (auto simp add: zero_le_mult_iff tsub_def pos_imp_zdiv_nonneg_iff)
+ apply (cases "y = 0")
+ apply (auto simp add: pos_imp_zdiv_nonneg_iff)
+ apply (cases "y = 0")
+ apply auto
+ done
-lemma transfer_nat_int_relations [transfer key: transfer_morphism_nat_int]:
+lemma transfer_nat_int_relations [no_atp, transfer key: transfer_morphism_nat_int]:
"x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
(nat (x::int) = nat y) = (x = y)"
"x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
@@ -94,7 +104,7 @@
then show "\<exists>x. P x" by auto
qed
-lemma transfer_nat_int_quantifiers [transfer key: transfer_morphism_nat_int]:
+lemma transfer_nat_int_quantifiers [no_atp, transfer key: transfer_morphism_nat_int]:
"(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))"
"(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))"
by (rule all_nat, rule ex_nat)
@@ -126,7 +136,7 @@
where
"nat_set S = (ALL x:S. x >= 0)"
-lemma transfer_nat_int_set_functions:
+lemma transfer_nat_int_set_functions [no_atp]:
"card A = card (int ` A)"
"{} = nat ` ({}::int set)"
"A Un B = nat ` (int ` A Un int ` B)"
@@ -144,7 +154,7 @@
apply auto
done
-lemma transfer_nat_int_set_function_closures:
+lemma transfer_nat_int_set_function_closures [no_atp]:
"nat_set {}"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
@@ -154,7 +164,7 @@
unfolding nat_set_def apply auto
done
-lemma transfer_nat_int_set_relations:
+lemma transfer_nat_int_set_relations [no_atp]:
"(finite A) = (finite (int ` A))"
"(x : A) = (int x : int ` A)"
"(A = B) = (int ` A = int ` B)"
@@ -169,11 +179,11 @@
apply (drule_tac x = "int x" in spec, auto)
done
-lemma transfer_nat_int_set_return_embed: "nat_set A \<Longrightarrow>
+lemma transfer_nat_int_set_return_embed [no_atp]: "nat_set A \<Longrightarrow>
(int ` nat ` A = A)"
by (auto simp add: nat_set_def image_def)
-lemma transfer_nat_int_set_cong: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow>
+lemma transfer_nat_int_set_cong [no_atp]: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow>
{(x::int). x >= 0 & P x} = {x. x >= 0 & P' x}"
by auto
@@ -189,7 +199,7 @@
text \<open>sum and prod\<close>
(* this handles the case where the *domain* of f is nat *)
-lemma transfer_nat_int_sum_prod:
+lemma transfer_nat_int_sum_prod [no_atp]:
"sum f A = sum (%x. f (nat x)) (int ` A)"
"prod f A = prod (%x. f (nat x)) (int ` A)"
apply (subst sum.reindex)
@@ -199,14 +209,14 @@
done
(* this handles the case where the *range* of f is nat *)
-lemma transfer_nat_int_sum_prod2:
+lemma transfer_nat_int_sum_prod2 [no_atp]:
"sum f A = nat(sum (%x. int (f x)) A)"
"prod f A = nat(prod (%x. int (f x)) A)"
apply (simp only: int_sum [symmetric] nat_int)
apply (simp only: int_prod [symmetric] nat_int)
done
-lemma transfer_nat_int_sum_prod_closure:
+lemma transfer_nat_int_sum_prod_closure [no_atp]:
"nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> sum f A >= 0"
"nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> prod f A >= 0"
unfolding nat_set_def
@@ -236,7 +246,7 @@
Also, why aren't sum.cong and prod.cong enough,
with the previously mentioned rule turned on? *)
-lemma transfer_nat_int_sum_prod_cong:
+lemma transfer_nat_int_sum_prod_cong [no_atp]:
"A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
sum f A = sum g B"
"A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
@@ -257,7 +267,7 @@
text \<open>set up transfer direction\<close>
-lemma transfer_morphism_int_nat: "transfer_morphism int (\<lambda>n. True)" ..
+lemma transfer_morphism_int_nat [no_atp]: "transfer_morphism int (\<lambda>n. True)" ..
declare transfer_morphism_int_nat [transfer add
mode: manual
@@ -273,21 +283,23 @@
where
"is_nat x = (x >= 0)"
-lemma transfer_int_nat_numerals:
+lemma transfer_int_nat_numerals [no_atp]:
"0 = int 0"
"1 = int 1"
"2 = int 2"
"3 = int 3"
by auto
-lemma transfer_int_nat_functions:
+lemma transfer_int_nat_functions [no_atp]:
"(int x) + (int y) = int (x + y)"
"(int x) * (int y) = int (x * y)"
"tsub (int x) (int y) = int (x - y)"
"(int x)^n = int (x^n)"
- by (auto simp add: of_nat_mult tsub_def of_nat_power)
+ "(int x) div (int y) = int (x div y)"
+ "(int x) mod (int y) = int (x mod y)"
+ by (auto simp add: zdiv_int zmod_int tsub_def)
-lemma transfer_int_nat_function_closures:
+lemma transfer_int_nat_function_closures [no_atp]:
"is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)"
"is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)"
"is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)"
@@ -297,9 +309,11 @@
"is_nat 2"
"is_nat 3"
"is_nat (int z)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
+ "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
by (simp_all only: is_nat_def transfer_nat_int_function_closures)
-lemma transfer_int_nat_relations:
+lemma transfer_int_nat_relations [no_atp]:
"(int x = int y) = (x = y)"
"(int x < int y) = (x < y)"
"(int x <= int y) = (x <= y)"
@@ -316,7 +330,7 @@
text \<open>first-order quantifiers\<close>
-lemma transfer_int_nat_quantifiers:
+lemma transfer_int_nat_quantifiers [no_atp]:
"(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))"
"(EX (x::int) >= 0. P x) = (EX (x::nat). P (int x))"
apply (subst all_nat)
@@ -341,7 +355,7 @@
text \<open>operations with sets\<close>
-lemma transfer_int_nat_set_functions:
+lemma transfer_int_nat_set_functions [no_atp]:
"nat_set A \<Longrightarrow> card A = card (nat ` A)"
"{} = int ` ({}::nat set)"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)"
@@ -353,7 +367,7 @@
transfer_nat_int_set_return_embed nat_0_le
cong: transfer_nat_int_set_cong)
-lemma transfer_int_nat_set_function_closures:
+lemma transfer_int_nat_set_function_closures [no_atp]:
"nat_set {}"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
@@ -362,7 +376,7 @@
"nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x"
by (simp_all only: transfer_nat_int_set_function_closures is_nat_def)
-lemma transfer_int_nat_set_relations:
+lemma transfer_int_nat_set_relations [no_atp]:
"nat_set A \<Longrightarrow> finite A = finite (nat ` A)"
"is_nat x \<Longrightarrow> nat_set A \<Longrightarrow> (x : A) = (nat x : nat ` A)"
"nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A = B) = (nat ` A = nat ` B)"
@@ -371,12 +385,12 @@
by (simp_all only: is_nat_def transfer_nat_int_set_relations
transfer_nat_int_set_return_embed nat_0_le)
-lemma transfer_int_nat_set_return_embed: "nat ` int ` A = A"
+lemma transfer_int_nat_set_return_embed [no_atp]: "nat ` int ` A = A"
by (simp only: transfer_nat_int_set_relations
transfer_nat_int_set_function_closures
transfer_nat_int_set_return_embed nat_0_le)
-lemma transfer_int_nat_set_cong: "(!!x. P x = P' x) \<Longrightarrow>
+lemma transfer_int_nat_set_cong [no_atp]: "(!!x. P x = P' x) \<Longrightarrow>
{(x::nat). P x} = {x. P' x}"
by auto
@@ -392,7 +406,7 @@
text \<open>sum and prod\<close>
(* this handles the case where the *domain* of f is int *)
-lemma transfer_int_nat_sum_prod:
+lemma transfer_int_nat_sum_prod [no_atp]:
"nat_set A \<Longrightarrow> sum f A = sum (%x. f (int x)) (nat ` A)"
"nat_set A \<Longrightarrow> prod f A = prod (%x. f (int x)) (nat ` A)"
apply (subst sum.reindex)
@@ -403,7 +417,7 @@
done
(* this handles the case where the *range* of f is int *)
-lemma transfer_int_nat_sum_prod2:
+lemma transfer_int_nat_sum_prod2 [no_atp]:
"(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> sum f A = int(sum (%x. nat (f x)) A)"
"(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow>
prod f A = int(prod (%x. nat (f x)) A)"
@@ -414,4 +428,6 @@
return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2
cong: sum.cong prod.cong]
+declare transfer_morphism_int_nat [transfer add return: even_int_iff]
+
end
--- a/src/HOL/Nonstandard_Analysis/StarDef.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Nonstandard_Analysis/StarDef.thy Mon Oct 09 22:08:05 2017 +0200
@@ -772,9 +772,6 @@
instance star :: (semidom_divide) semidom_divide
by (intro_classes; transfer) simp_all
-instance star :: (semiring_div) semiring_div
- by (intro_classes; transfer) (simp_all add: div_mult_mod_eq)
-
instance star :: (ring_no_zero_divisors) ring_no_zero_divisors ..
instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
instance star :: (idom) idom ..
@@ -821,6 +818,26 @@
instance star :: (linordered_field) linordered_field ..
+instance star :: (algebraic_semidom) algebraic_semidom ..
+
+instantiation star :: (normalization_semidom) normalization_semidom
+begin
+
+definition unit_factor_star :: "'a star \<Rightarrow> 'a star"
+ where [transfer_unfold]: "unit_factor_star = *f* unit_factor"
+
+definition normalize_star :: "'a star \<Rightarrow> 'a star"
+ where [transfer_unfold]: "normalize_star = *f* normalize"
+
+instance
+ by standard (transfer; simp add: is_unit_unit_factor unit_factor_mult)+
+
+end
+
+instance star :: (semidom_modulo) semidom_modulo
+ by standard (transfer; simp)
+
+
subsection \<open>Power\<close>
@@ -904,57 +921,6 @@
instance star :: (ring_char_0) ring_char_0 ..
-instance star :: (semiring_parity) semiring_parity
- apply intro_classes
- apply (transfer, rule odd_one)
- apply (transfer, erule (1) odd_even_add)
- apply (transfer, erule even_multD)
- apply (transfer, erule odd_ex_decrement)
- done
-
-instance star :: (semiring_div_parity) semiring_div_parity
- apply intro_classes
- apply (transfer, rule parity)
- apply (transfer, rule one_mod_two_eq_one)
- apply (transfer, rule zero_not_eq_two)
- done
-
-instantiation star :: (semiring_numeral_div) semiring_numeral_div
-begin
-
-definition divmod_star :: "num \<Rightarrow> num \<Rightarrow> 'a star \<times> 'a star"
- where divmod_star_def: "divmod_star m n = (numeral m div numeral n, numeral m mod numeral n)"
-
-definition divmod_step_star :: "num \<Rightarrow> 'a star \<times> 'a star \<Rightarrow> 'a star \<times> 'a star"
- where "divmod_step_star l qr =
- (let (q, r) = qr
- in if r \<ge> numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))"
-
-instance
-proof
- show "divmod m n = (numeral m div numeral n :: 'a star, numeral m mod numeral n)" for m n
- by (fact divmod_star_def)
- show "divmod_step l qr =
- (let (q, r) = qr
- in if r \<ge> numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))"
- for l and qr :: "'a star \<times> 'a star"
- by (fact divmod_step_star_def)
-qed (transfer,
- fact
- semiring_numeral_div_class.div_less
- semiring_numeral_div_class.mod_less
- semiring_numeral_div_class.div_positive
- semiring_numeral_div_class.mod_less_eq_dividend
- semiring_numeral_div_class.pos_mod_bound
- semiring_numeral_div_class.pos_mod_sign
- semiring_numeral_div_class.mod_mult2_eq
- semiring_numeral_div_class.div_mult2_eq
- semiring_numeral_div_class.discrete)+
-
-end
-
-declare divmod_algorithm_code [where ?'a = "'a::semiring_numeral_div star", code]
-
subsection \<open>Finite class\<close>
--- a/src/HOL/NthRoot.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/NthRoot.thy Mon Oct 09 22:08:05 2017 +0200
@@ -218,7 +218,7 @@
lemma real_root_inverse: "root n (inverse x) = inverse (root n x)"
by (auto split: split_root elim!: sgn_power_injE
- simp: inverse_sgn power_inverse)
+ simp: power_inverse)
lemma real_root_divide: "root n (x / y) = root n x / root n y"
by (simp add: divide_inverse real_root_mult real_root_inverse)
@@ -715,7 +715,7 @@
have "x n \<le> sqrt (2 / real n)" if "2 < n" for n :: nat
proof -
have "1 + (real (n - 1) * n) / 2 * (x n)\<^sup>2 = 1 + of_nat (n choose 2) * (x n)\<^sup>2"
- by (auto simp add: choose_two of_nat_div mod_eq_0_iff_dvd)
+ by (auto simp add: choose_two field_char_0_class.of_nat_div mod_eq_0_iff_dvd)
also have "\<dots> \<le> (\<Sum>k\<in>{0, 2}. of_nat (n choose k) * x n^k)"
by (simp add: x_def)
also have "\<dots> \<le> (\<Sum>k=0..n. of_nat (n choose k) * x n^k)"
--- a/src/HOL/Number_Theory/Cong.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Number_Theory/Cong.thy Mon Oct 09 22:08:05 2017 +0200
@@ -86,8 +86,7 @@
"x \<ge> 0 \<Longrightarrow> y \<ge> 0 \<Longrightarrow> m \<ge> 0 \<Longrightarrow> [nat x = nat y] (mod (nat m)) \<longleftrightarrow> [x = y] (mod m)"
for x y m :: int
unfolding cong_int_def cong_nat_def
- by (metis Divides.transfer_int_nat_functions(2) nat_0_le nat_mod_distrib)
-
+ by (metis int_nat_eq nat_mod_distrib zmod_int)
declare transfer_morphism_nat_int [transfer add return: transfer_nat_int_cong]
--- a/src/HOL/Number_Theory/Gauss.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Number_Theory/Gauss.thy Mon Oct 09 22:08:05 2017 +0200
@@ -196,7 +196,8 @@
moreover have "int y = 0 \<or> 0 < int y" for y
by linarith
ultimately show "0 < x mod int p"
- by (metis (no_types) B_greater_zero Divides.transfer_int_nat_functions(2) zero_less_imp_eq_int)
+ using B_greater_zero [of x]
+ by (auto simp add: mod_int_pos_iff intro: neq_le_trans)
qed
lemma F_subset: "F \<subseteq> {x. 0 < x \<and> x \<le> ((int p - 1) div 2)}"
--- a/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Mon Oct 09 22:08:05 2017 +0200
@@ -381,7 +381,7 @@
from that have *: "(fst x + (snd x - 1) * int p) mod int p = fst x"
by force
from that have "(fst x + (snd x - 1) * int p) div int p + 1 = snd x"
- by (auto simp: semiring_numeral_div_class.div_less)
+ by (auto simp: div_pos_pos_trivial)
with * show "f_3 (g_3 x) = x"
by (simp add: f_3_def g_3_def)
qed
--- a/src/HOL/Number_Theory/Residues.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Number_Theory/Residues.thy Mon Oct 09 22:08:05 2017 +0200
@@ -365,8 +365,7 @@
done
finally have "fact (p - 1) mod p = \<ominus> \<one>" .
then show ?thesis
- by (metis of_nat_fact Divides.transfer_int_nat_functions(2)
- cong_int_def res_neg_eq res_one_eq)
+ by (simp add: cong_int_def res_neg_eq res_one_eq zmod_int)
qed
lemma wilson_theorem:
--- a/src/HOL/Number_Theory/Totient.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Number_Theory/Totient.thy Mon Oct 09 22:08:05 2017 +0200
@@ -335,19 +335,30 @@
qed
lemma totient_prod_coprime:
- assumes "pairwise_coprime (f ` A)" "inj_on f A"
- shows "totient (prod f A) = prod (\<lambda>x. totient (f x)) A"
+ assumes "pairwise coprime (f ` A)" "inj_on f A"
+ shows "totient (prod f A) = (\<Prod>a\<in>A. totient (f a))"
using assms
proof (induction A rule: infinite_finite_induct)
case (insert x A)
- from insert.prems and insert.hyps have *: "coprime (prod f A) (f x)"
- by (intro prod_coprime[OF pairwise_coprimeD[OF insert.prems(1)]]) (auto simp: inj_on_def)
+ have *: "coprime (prod f A) (f x)"
+ proof (rule prod_coprime)
+ fix y
+ assume "y \<in> A"
+ with \<open>x \<notin> A\<close> have "y \<noteq> x"
+ by auto
+ with \<open>x \<notin> A\<close> \<open>y \<in> A\<close> \<open>inj_on f (insert x A)\<close> have "f y \<noteq> f x"
+ using inj_onD [of f "insert x A" y x]
+ by auto
+ with \<open>y \<in> A\<close> show "coprime (f y) (f x)"
+ using pairwiseD [OF \<open>pairwise coprime (f ` insert x A)\<close>]
+ by auto
+ qed
from insert.hyps have "prod f (insert x A) = prod f A * f x" by simp
also have "totient \<dots> = totient (prod f A) * totient (f x)"
using insert.hyps insert.prems by (intro totient_mult_coprime *)
- also have "totient (prod f A) = (\<Prod>x\<in>A. totient (f x))"
- using insert.prems by (intro insert.IH) (auto dest: pairwise_coprime_subset)
- also from insert.hyps have "\<dots> * totient (f x) = (\<Prod>x\<in>insert x A. totient (f x))" by simp
+ also have "totient (prod f A) = (\<Prod>a\<in>A. totient (f a))"
+ using insert.prems by (intro insert.IH) (auto dest: pairwise_subset)
+ also from insert.hyps have "\<dots> * totient (f x) = (\<Prod>a\<in>insert x A. totient (f a))" by simp
finally show ?case .
qed simp_all
@@ -375,8 +386,8 @@
by (rule prime_factorization_nat)
also have "totient \<dots> = (\<Prod>x\<in>prime_factors n. totient (x ^ multiplicity x n))"
proof (rule totient_prod_coprime)
- show "pairwise_coprime ((\<lambda>p. p ^ multiplicity p n) ` prime_factors n)"
- proof (standard, clarify, goal_cases)
+ show "pairwise coprime ((\<lambda>p. p ^ multiplicity p n) ` prime_factors n)"
+ proof (rule pairwiseI, clarify)
fix p q assume "p \<in># prime_factorization n" "q \<in># prime_factorization n"
"p ^ multiplicity p n \<noteq> q ^ multiplicity q n"
thus "coprime (p ^ multiplicity p n) (q ^ multiplicity q n)"
--- a/src/HOL/Numeral_Simprocs.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Numeral_Simprocs.thy Mon Oct 09 22:08:05 2017 +0200
@@ -169,8 +169,8 @@
(* TODO: remove comm_ring_1 constraint if possible *)
simproc_setup int_div_cancel_numeral_factors
- ("((l::'a::{semiring_div,comm_ring_1,ring_char_0}) * m) div n"
- |"(l::'a::{semiring_div,comm_ring_1,ring_char_0}) div (m * n)") =
+ ("((l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) * m) div n"
+ |"(l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) div (m * n)") =
\<open>fn phi => Numeral_Simprocs.div_cancel_numeral_factor\<close>
simproc_setup divide_cancel_numeral_factor
@@ -194,13 +194,13 @@
\<open>fn phi => Numeral_Simprocs.less_cancel_factor\<close>
simproc_setup int_div_cancel_factor
- ("((l::'a::semiring_div) * m) div n"
- |"(l::'a::semiring_div) div (m * n)") =
+ ("((l::'a::euclidean_semiring_cancel) * m) div n"
+ |"(l::'a::euclidean_semiring_cancel) div (m * n)") =
\<open>fn phi => Numeral_Simprocs.div_cancel_factor\<close>
simproc_setup int_mod_cancel_factor
- ("((l::'a::semiring_div) * m) mod n"
- |"(l::'a::semiring_div) mod (m * n)") =
+ ("((l::'a::euclidean_semiring_cancel) * m) mod n"
+ |"(l::'a::euclidean_semiring_cancel) mod (m * n)") =
\<open>fn phi => Numeral_Simprocs.mod_cancel_factor\<close>
simproc_setup dvd_cancel_factor
--- a/src/HOL/Parity.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Parity.thy Mon Oct 09 22:08:05 2017 +0200
@@ -6,19 +6,73 @@
section \<open>Parity in rings and semirings\<close>
theory Parity
- imports Nat_Transfer Euclidean_Division
+ imports Euclidean_Division
begin
subsection \<open>Ring structures with parity and \<open>even\<close>/\<open>odd\<close> predicates\<close>
-class semiring_parity = comm_semiring_1_cancel + numeral +
- assumes odd_one [simp]: "\<not> 2 dvd 1"
- assumes odd_even_add: "\<not> 2 dvd a \<Longrightarrow> \<not> 2 dvd b \<Longrightarrow> 2 dvd a + b"
- assumes even_multD: "2 dvd a * b \<Longrightarrow> 2 dvd a \<or> 2 dvd b"
- assumes odd_ex_decrement: "\<not> 2 dvd a \<Longrightarrow> \<exists>b. a = b + 1"
+class semiring_parity = linordered_semidom + unique_euclidean_semiring +
+ assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n"
+ and odd_imp_mod_2_eq_1: "\<not> 2 dvd a \<Longrightarrow> a mod 2 = 1"
+
+context semiring_parity
begin
-subclass semiring_numeral ..
+lemma of_nat_dvd_iff:
+ "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q")
+proof (cases "m = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ show ?thesis
+ proof
+ assume ?Q
+ then show ?P
+ by (auto elim: dvd_class.dvdE)
+ next
+ assume ?P
+ with False have "of_nat n = of_nat n div of_nat m * of_nat m"
+ by simp
+ then have "of_nat n = of_nat (n div m * m)"
+ by (simp add: of_nat_div)
+ then have "n = n div m * m"
+ by (simp only: of_nat_eq_iff)
+ then have "n = m * (n div m)"
+ by (simp add: ac_simps)
+ then show ?Q ..
+ qed
+qed
+
+lemma of_nat_mod:
+ "of_nat (m mod n) = of_nat m mod of_nat n"
+proof -
+ have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m"
+ by (simp add: div_mult_mod_eq)
+ also have "of_nat m = of_nat (m div n * n + m mod n)"
+ by simp
+ finally show ?thesis
+ by (simp only: of_nat_div of_nat_mult of_nat_add) simp
+qed
+
+lemma one_div_two_eq_zero [simp]:
+ "1 div 2 = 0"
+proof -
+ from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0"
+ by (simp only:) simp
+ then show ?thesis
+ by simp
+qed
+
+lemma one_mod_two_eq_one [simp]:
+ "1 mod 2 = 1"
+proof -
+ from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1"
+ by (simp only:) simp
+ then show ?thesis
+ by simp
+qed
abbreviation even :: "'a \<Rightarrow> bool"
where "even a \<equiv> 2 dvd a"
@@ -26,11 +80,27 @@
abbreviation odd :: "'a \<Rightarrow> bool"
where "odd a \<equiv> \<not> 2 dvd a"
-lemma even_zero [simp]: "even 0"
- by (fact dvd_0_right)
+lemma even_iff_mod_2_eq_zero:
+ "even a \<longleftrightarrow> a mod 2 = 0"
+ by (fact dvd_eq_mod_eq_0)
+
+lemma odd_iff_mod_2_eq_one:
+ "odd a \<longleftrightarrow> a mod 2 = 1"
+ by (auto dest: odd_imp_mod_2_eq_1)
-lemma even_plus_one_iff [simp]: "even (a + 1) \<longleftrightarrow> odd a"
- by (auto simp add: dvd_add_right_iff intro: odd_even_add)
+lemma parity_cases [case_names even odd]:
+ assumes "even a \<Longrightarrow> a mod 2 = 0 \<Longrightarrow> P"
+ assumes "odd a \<Longrightarrow> a mod 2 = 1 \<Longrightarrow> P"
+ shows P
+ using assms by (cases "even a") (simp_all add: odd_iff_mod_2_eq_one)
+
+lemma not_mod_2_eq_1_eq_0 [simp]:
+ "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0"
+ by (cases a rule: parity_cases) simp_all
+
+lemma not_mod_2_eq_0_eq_1 [simp]:
+ "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1"
+ by (cases a rule: parity_cases) simp_all
lemma evenE [elim?]:
assumes "even a"
@@ -41,22 +111,107 @@
assumes "odd a"
obtains b where "a = 2 * b + 1"
proof -
- from assms obtain b where *: "a = b + 1"
- by (blast dest: odd_ex_decrement)
- with assms have "even (b + 2)" by simp
- then have "even b" by simp
- then obtain c where "b = 2 * c" ..
- with * have "a = 2 * c + 1" by simp
- with that show thesis .
+ have "a = 2 * (a div 2) + a mod 2"
+ by (simp add: mult_div_mod_eq)
+ with assms have "a = 2 * (a div 2) + 1"
+ by (simp add: odd_iff_mod_2_eq_one)
+ then show ?thesis ..
+qed
+
+lemma mod_2_eq_odd:
+ "a mod 2 = of_bool (odd a)"
+ by (auto elim: oddE)
+
+lemma one_mod_2_pow_eq [simp]:
+ "1 mod (2 ^ n) = of_bool (n > 0)"
+proof -
+ have "1 mod (2 ^ n) = (of_bool (n > 0) :: nat)"
+ by (induct n) (simp_all add: mod_mult2_eq)
+ then have "of_nat (1 mod (2 ^ n)) = of_bool (n > 0)"
+ by simp
+ then show ?thesis
+ by (simp add: of_nat_mod)
+qed
+
+lemma even_of_nat [simp]:
+ "even (of_nat a) \<longleftrightarrow> even a"
+proof -
+ have "even (of_nat a) \<longleftrightarrow> of_nat 2 dvd of_nat a"
+ by simp
+ also have "\<dots> \<longleftrightarrow> even a"
+ by (simp only: of_nat_dvd_iff)
+ finally show ?thesis .
+qed
+
+lemma even_zero [simp]:
+ "even 0"
+ by (fact dvd_0_right)
+
+lemma odd_one [simp]:
+ "odd 1"
+proof -
+ have "\<not> (2 :: nat) dvd 1"
+ by simp
+ then have "\<not> of_nat 2 dvd of_nat 1"
+ unfolding of_nat_dvd_iff by simp
+ then show ?thesis
+ by simp
qed
-lemma even_times_iff [simp]: "even (a * b) \<longleftrightarrow> even a \<or> even b"
- by (auto dest: even_multD)
+lemma odd_even_add:
+ "even (a + b)" if "odd a" and "odd b"
+proof -
+ from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1"
+ by (blast elim: oddE)
+ then have "a + b = 2 * c + 2 * d + (1 + 1)"
+ by (simp only: ac_simps)
+ also have "\<dots> = 2 * (c + d + 1)"
+ by (simp add: algebra_simps)
+ finally show ?thesis ..
+qed
+
+lemma even_add [simp]:
+ "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)"
+ by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add)
+
+lemma odd_add [simp]:
+ "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)"
+ by simp
+
+lemma even_plus_one_iff [simp]:
+ "even (a + 1) \<longleftrightarrow> odd a"
+ by (auto simp add: dvd_add_right_iff intro: odd_even_add)
+
+lemma even_mult_iff [simp]:
+ "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q")
+proof
+ assume ?Q
+ then show ?P
+ by auto
+next
+ assume ?P
+ show ?Q
+ proof (rule ccontr)
+ assume "\<not> (even a \<or> even b)"
+ then have "odd a" and "odd b"
+ by auto
+ then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1"
+ by (blast elim: oddE)
+ then have "a * b = (2 * r + 1) * (2 * s + 1)"
+ by simp
+ also have "\<dots> = 2 * (2 * r * s + r + s) + 1"
+ by (simp add: algebra_simps)
+ finally have "odd (a * b)"
+ by simp
+ with \<open>?P\<close> show False
+ by auto
+ qed
+qed
lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))"
proof -
have "even (2 * numeral n)"
- unfolding even_times_iff by simp
+ unfolding even_mult_iff by simp
then have "even (numeral n + numeral n)"
unfolding mult_2 .
then show ?thesis
@@ -77,15 +232,26 @@
then show False by simp
qed
-lemma even_add [simp]: "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)"
- by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add)
-
-lemma odd_add [simp]: "odd (a + b) \<longleftrightarrow> (\<not> (odd a \<longleftrightarrow> odd b))"
- by simp
-
lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0"
by (induct n) auto
+lemma even_succ_div_two [simp]:
+ "even a \<Longrightarrow> (a + 1) div 2 = a div 2"
+ by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero)
+
+lemma odd_succ_div_two [simp]:
+ "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1"
+ by (auto elim!: oddE simp add: add.assoc)
+
+lemma even_two_times_div_two:
+ "even a \<Longrightarrow> 2 * (a div 2) = a"
+ by (fact dvd_mult_div_cancel)
+
+lemma odd_two_times_div_two_succ [simp]:
+ "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a"
+ using mult_div_mod_eq [of 2 a]
+ by (simp add: even_iff_mod_2_eq_zero)
+
end
class ring_parity = ring + semiring_parity
@@ -93,80 +259,48 @@
subclass comm_ring_1 ..
-lemma even_minus [simp]: "even (- a) \<longleftrightarrow> even a"
+lemma even_minus [simp]:
+ "even (- a) \<longleftrightarrow> even a"
by (fact dvd_minus_iff)
-lemma even_diff [simp]: "even (a - b) \<longleftrightarrow> even (a + b)"
+lemma even_diff [simp]:
+ "even (a - b) \<longleftrightarrow> even (a + b)"
using even_add [of a "- b"] by simp
end
-subsection \<open>Instances for @{typ nat} and @{typ int}\<close>
+subsection \<open>Instance for @{typ nat}\<close>
-lemma even_Suc_Suc_iff [simp]: "2 dvd Suc (Suc n) \<longleftrightarrow> 2 dvd n"
+instance nat :: semiring_parity
+ by standard (simp_all add: dvd_eq_mod_eq_0)
+
+lemma even_Suc_Suc_iff [simp]:
+ "even (Suc (Suc n)) \<longleftrightarrow> even n"
using dvd_add_triv_right_iff [of 2 n] by simp
-lemma even_Suc [simp]: "2 dvd Suc n \<longleftrightarrow> \<not> 2 dvd n"
- by (induct n) auto
+lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n"
+ using even_plus_one_iff [of n] by simp
-lemma even_diff_nat [simp]: "2 dvd (m - n) \<longleftrightarrow> m < n \<or> 2 dvd (m + n)"
- for m n :: nat
+lemma even_diff_nat [simp]:
+ "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat
proof (cases "n \<le> m")
case True
then have "m - n + n * 2 = m + n" by (simp add: mult_2_right)
- moreover have "2 dvd (m - n) \<longleftrightarrow> 2 dvd (m - n + n * 2)" by simp
- ultimately have "2 dvd (m - n) \<longleftrightarrow> 2 dvd (m + n)" by (simp only:)
+ moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp
+ ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:)
then show ?thesis by auto
next
case False
then show ?thesis by simp
qed
-instance nat :: semiring_parity
-proof
- show "\<not> 2 dvd (1 :: nat)"
- by (rule notI, erule dvdE) simp
-next
- fix m n :: nat
- assume "\<not> 2 dvd m"
- moreover assume "\<not> 2 dvd n"
- ultimately have *: "2 dvd Suc m \<and> 2 dvd Suc n"
- by simp
- then have "2 dvd (Suc m + Suc n)"
- by (blast intro: dvd_add)
- also have "Suc m + Suc n = m + n + 2"
- by simp
- finally show "2 dvd (m + n)"
- using dvd_add_triv_right_iff [of 2 "m + n"] by simp
-next
- fix m n :: nat
- assume *: "2 dvd (m * n)"
- show "2 dvd m \<or> 2 dvd n"
- proof (rule disjCI)
- assume "\<not> 2 dvd n"
- then have "2 dvd (Suc n)" by simp
- then obtain r where "Suc n = 2 * r" ..
- moreover from * obtain s where "m * n = 2 * s" ..
- then have "2 * s + m = m * Suc n" by simp
- ultimately have " 2 * s + m = 2 * (m * r)"
- by (simp add: algebra_simps)
- then have "m = 2 * (m * r - s)" by simp
- then show "2 dvd m" ..
- qed
-next
- fix n :: nat
- assume "\<not> 2 dvd n"
- then show "\<exists>m. n = m + 1"
- by (cases n) simp_all
-qed
-
-lemma odd_pos: "odd n \<Longrightarrow> 0 < n"
- for n :: nat
+lemma odd_pos:
+ "odd n \<Longrightarrow> 0 < n" for n :: nat
by (auto elim: oddE)
-lemma Suc_double_not_eq_double: "Suc (2 * m) \<noteq> 2 * n"
- for m n :: nat
+lemma Suc_double_not_eq_double:
+ "Suc (2 * m) \<noteq> 2 * n"
proof
assume "Suc (2 * m) = 2 * n"
moreover have "odd (Suc (2 * m))" and "even (2 * n)"
@@ -174,55 +308,58 @@
ultimately show False by simp
qed
-lemma double_not_eq_Suc_double: "2 * m \<noteq> Suc (2 * n)"
- for m n :: nat
+lemma double_not_eq_Suc_double:
+ "2 * m \<noteq> Suc (2 * n)"
using Suc_double_not_eq_double [of n m] by simp
-lemma even_diff_iff [simp]: "2 dvd (k - l) \<longleftrightarrow> 2 dvd (k + l)"
- for k l :: int
- using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: mult_2_right)
-
-lemma even_abs_add_iff [simp]: "2 dvd (\<bar>k\<bar> + l) \<longleftrightarrow> 2 dvd (k + l)"
- for k l :: int
- by (cases "k \<ge> 0") (simp_all add: ac_simps)
-
-lemma even_add_abs_iff [simp]: "2 dvd (k + \<bar>l\<bar>) \<longleftrightarrow> 2 dvd (k + l)"
- for k l :: int
- using even_abs_add_iff [of l k] by (simp add: ac_simps)
-
lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n"
by (auto elim: oddE)
-instance int :: ring_parity
-proof
- show "\<not> 2 dvd (1 :: int)"
- by (simp add: dvd_int_unfold_dvd_nat)
-next
- fix k l :: int
- assume "\<not> 2 dvd k"
- moreover assume "\<not> 2 dvd l"
- ultimately have "2 dvd (nat \<bar>k\<bar> + nat \<bar>l\<bar>)"
- by (auto simp add: dvd_int_unfold_dvd_nat intro: odd_even_add)
- then have "2 dvd (\<bar>k\<bar> + \<bar>l\<bar>)"
- by (simp add: dvd_int_unfold_dvd_nat nat_add_distrib)
- then show "2 dvd (k + l)"
+lemma even_Suc_div_two [simp]:
+ "even n \<Longrightarrow> Suc n div 2 = n div 2"
+ using even_succ_div_two [of n] by simp
+
+lemma odd_Suc_div_two [simp]:
+ "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)"
+ using odd_succ_div_two [of n] by simp
+
+lemma odd_two_times_div_two_nat [simp]:
+ assumes "odd n"
+ shows "2 * (n div 2) = n - (1 :: nat)"
+proof -
+ from assms have "2 * (n div 2) + 1 = n"
+ by (rule odd_two_times_div_two_succ)
+ then have "Suc (2 * (n div 2)) - 1 = n - 1"
by simp
-next
- fix k l :: int
- assume "2 dvd (k * l)"
- then show "2 dvd k \<or> 2 dvd l"
- by (simp add: dvd_int_unfold_dvd_nat even_multD nat_abs_mult_distrib)
-next
- fix k :: int
- have "k = (k - 1) + 1" by simp
- then show "\<exists>l. k = l + 1" ..
+ then show ?thesis
+ by simp
qed
-lemma even_int_iff [simp]: "even (int n) \<longleftrightarrow> even n"
- by (simp add: dvd_int_iff)
-
-lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k"
- by (simp add: even_int_iff [symmetric])
+lemma parity_induct [case_names zero even odd]:
+ assumes zero: "P 0"
+ assumes even: "\<And>n. P n \<Longrightarrow> P (2 * n)"
+ assumes odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))"
+ shows "P n"
+proof (induct n rule: less_induct)
+ case (less n)
+ show "P n"
+ proof (cases "n = 0")
+ case True with zero show ?thesis by simp
+ next
+ case False
+ with less have hyp: "P (n div 2)" by simp
+ show ?thesis
+ proof (cases "even n")
+ case True
+ with hyp even [of "n div 2"] show ?thesis
+ by simp
+ next
+ case False
+ with hyp odd [of "n div 2"] show ?thesis
+ by simp
+ qed
+ qed
+qed
subsection \<open>Parity and powers\<close>
@@ -236,6 +373,10 @@
lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)"
by (auto elim: oddE)
+lemma uminus_power_if:
+ "(- a) ^ n = (if even n then a ^ n else - (a ^ n))"
+ by auto
+
lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
by simp
@@ -313,9 +454,6 @@
qed
qed
-lemma (in comm_ring_1) uminus_power_if: "(- x) ^ n = (if even n then x^n else - (x ^ n))"
- by auto
-
text \<open>Simplify, when the exponent is a numeral\<close>
lemma zero_le_power_eq_numeral [simp]:
@@ -346,8 +484,39 @@
end
-subsubsection \<open>Tool setup\<close>
+subsection \<open>Instance for @{typ int}\<close>
-declare transfer_morphism_int_nat [transfer add return: even_int_iff]
+instance int :: ring_parity
+proof
+ fix k l :: int
+ show "k mod 2 = 1" if "\<not> 2 dvd k"
+ proof (rule order_antisym)
+ have "0 \<le> k mod 2" and "k mod 2 < 2"
+ by auto
+ moreover have "k mod 2 \<noteq> 0"
+ using that by (simp add: dvd_eq_mod_eq_0)
+ ultimately have "0 < k mod 2"
+ by (simp only: less_le) simp
+ then show "1 \<le> k mod 2"
+ by simp
+ from \<open>k mod 2 < 2\<close> show "k mod 2 \<le> 1"
+ by (simp only: less_le) simp
+ qed
+qed (simp_all add: dvd_eq_mod_eq_0 divide_int_def)
+
+lemma even_diff_iff [simp]:
+ "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int
+ using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: mult_2_right)
+
+lemma even_abs_add_iff [simp]:
+ "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int
+ by (cases "k \<ge> 0") (simp_all add: ac_simps)
+
+lemma even_add_abs_iff [simp]:
+ "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int
+ using even_abs_add_iff [of l k] by (simp add: ac_simps)
+
+lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k"
+ by (simp add: even_of_nat [of "nat k", where ?'a = int, symmetric])
end
--- a/src/HOL/Probability/Probability_Mass_Function.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Oct 09 22:08:05 2017 +0200
@@ -797,7 +797,7 @@
(insert assms, auto simp: scaleR_sum_left)
qed
also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R (\<Sum>i\<in>(\<Union>x\<in>A. set_pmf (f x)). pmf (f j) i *\<^sub>R h i))"
- by (subst sum.commute) (simp add: scaleR_sum_right)
+ by (subst sum.swap) (simp add: scaleR_sum_right)
also have "\<dots> = (\<Sum>j\<in>A. pmf p j *\<^sub>R measure_pmf.expectation (f j) h)"
proof (intro sum.cong refl, goal_cases)
case (1 x)
@@ -2141,7 +2141,7 @@
by (intro sum.cong refl, subst sum_list_map_filter', subst sum_list_sum_nth) simp
also have "\<dots> = (\<Sum>xa = 0..<length xs. (\<Sum>x\<in>set (map fst xs).
if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
- by (rule sum.commute)
+ by (rule sum.swap)
also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then
(\<Sum>x\<in>set (map fst xs). if x = fst (xs ! xa) then snd (xs ! xa) else 0) else 0)"
by (auto intro!: sum.cong sum.neutral simp del: sum.delta)
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Oct 09 22:08:05 2017 +0200
@@ -321,7 +321,7 @@
apply metis
done
also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
- by (subst sum.commute) rule
+ by (subst sum.swap) rule
also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
by (auto simp add: sum_distrib_left vimage_def intro!: sum.cong prob_conj_imp1)
finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
--- a/src/HOL/ROOT Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/ROOT Mon Oct 09 22:08:05 2017 +0200
@@ -71,7 +71,6 @@
Computational_Algebra
(*conflicting type class instantiations and dependent applications*)
Field_as_Ring
- Polynomial_Factorial
session "HOL-Hahn_Banach" in Hahn_Banach = "HOL-Library" +
description {*
@@ -548,6 +547,7 @@
Executable_Relation
Execute_Choice
Functions
+ Function_Growth
Gauge_Integration
Groebner_Examples
Guess
--- a/src/HOL/Rat.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Rat.thy Mon Oct 09 22:08:05 2017 +0200
@@ -199,12 +199,12 @@
(* We cannot state these two rules earlier because of pending sort hypotheses *)
lemma div_add_self1_no_field [simp]:
- assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: semiring_div) \<noteq> 0"
+ assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: euclidean_semiring_cancel) \<noteq> 0"
shows "(b + a) div b = a div b + 1"
using assms(2) by (fact div_add_self1)
lemma div_add_self2_no_field [simp]:
- assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: semiring_div) \<noteq> 0"
+ assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: euclidean_semiring_cancel) \<noteq> 0"
shows "(a + b) div b = a div b + 1"
using assms(2) by (fact div_add_self2)
--- a/src/HOL/Real.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Real.thy Mon Oct 09 22:08:05 2017 +0200
@@ -23,6 +23,11 @@
subsection \<open>Preliminary lemmas\<close>
+text{*Useful in convergence arguments*}
+lemma inverse_of_nat_le:
+ fixes n::nat shows "\<lbrakk>n \<le> m; n\<noteq>0\<rbrakk> \<Longrightarrow> 1 / of_nat m \<le> (1::'a::linordered_field) / of_nat n"
+ by (simp add: frac_le)
+
lemma inj_add_left [simp]: "inj (op + x)"
for x :: "'a::cancel_semigroup_add"
by (meson add_left_imp_eq injI)
--- a/src/HOL/Real_Vector_Spaces.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1469,17 +1469,14 @@
begin
lemma pos_bounded: "\<exists>K>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
- apply (insert bounded)
- apply (erule exE)
- apply (rule_tac x="max 1 K" in exI)
- apply safe
- apply (rule order_less_le_trans [OF zero_less_one max.cobounded1])
- apply (drule spec)
- apply (drule spec)
- apply (erule order_trans)
- apply (rule mult_left_mono [OF max.cobounded2])
- apply (intro mult_nonneg_nonneg norm_ge_zero)
- done
+proof -
+ obtain K where "\<And>a b. norm (a ** b) \<le> norm a * norm b * K"
+ using bounded by blast
+ then have "norm (a ** b) \<le> norm a * norm b * (max 1 K)" for a b
+ by (rule order.trans) (simp add: mult_left_mono)
+ then show ?thesis
+ by force
+qed
lemma nonneg_bounded: "\<exists>K\<ge>0. \<forall>a b. norm (a ** b) \<le> norm a * norm b * K"
using pos_bounded by (auto intro: order_less_imp_le)
--- a/src/HOL/Rings.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Rings.thy Mon Oct 09 22:08:05 2017 +0200
@@ -118,6 +118,13 @@
end
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
+begin
+
+lemma (in semiring_1) of_bool_conj:
+ "of_bool (P \<and> Q) = of_bool P * of_bool Q"
+ by auto
+
+end
text \<open>Abstract divisibility\<close>
@@ -1558,6 +1565,117 @@
end
+text \<open>Quotient and remainder in integral domains\<close>
+
+class semidom_modulo = algebraic_semidom + semiring_modulo
+begin
+
+lemma mod_0 [simp]: "0 mod a = 0"
+ using div_mult_mod_eq [of 0 a] by simp
+
+lemma mod_by_0 [simp]: "a mod 0 = a"
+ using div_mult_mod_eq [of a 0] by simp
+
+lemma mod_by_1 [simp]:
+ "a mod 1 = 0"
+proof -
+ from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp
+ then have "a + a mod 1 = a + 0" by simp
+ then show ?thesis by (rule add_left_imp_eq)
+qed
+
+lemma mod_self [simp]:
+ "a mod a = 0"
+ using div_mult_mod_eq [of a a] by simp
+
+lemma dvd_imp_mod_0 [simp]:
+ assumes "a dvd b"
+ shows "b mod a = 0"
+ using assms minus_div_mult_eq_mod [of b a] by simp
+
+lemma mod_0_imp_dvd:
+ assumes "a mod b = 0"
+ shows "b dvd a"
+proof -
+ have "b dvd ((a div b) * b)" by simp
+ also have "(a div b) * b = a"
+ using div_mult_mod_eq [of a b] by (simp add: assms)
+ finally show ?thesis .
+qed
+
+lemma mod_eq_0_iff_dvd:
+ "a mod b = 0 \<longleftrightarrow> b dvd a"
+ by (auto intro: mod_0_imp_dvd)
+
+lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]:
+ "a dvd b \<longleftrightarrow> b mod a = 0"
+ by (simp add: mod_eq_0_iff_dvd)
+
+lemma dvd_mod_iff:
+ assumes "c dvd b"
+ shows "c dvd a mod b \<longleftrightarrow> c dvd a"
+proof -
+ from assms have "(c dvd a mod b) \<longleftrightarrow> (c dvd ((a div b) * b + a mod b))"
+ by (simp add: dvd_add_right_iff)
+ also have "(a div b) * b + a mod b = a"
+ using div_mult_mod_eq [of a b] by simp
+ finally show ?thesis .
+qed
+
+lemma dvd_mod_imp_dvd:
+ assumes "c dvd a mod b" and "c dvd b"
+ shows "c dvd a"
+ using assms dvd_mod_iff [of c b a] by simp
+
+lemma dvd_minus_mod [simp]:
+ "b dvd a - a mod b"
+ by (simp add: minus_mod_eq_div_mult)
+
+lemma cancel_div_mod_rules:
+ "((a div b) * b + a mod b) + c = a + c"
+ "(b * (a div b) + a mod b) + c = a + c"
+ by (simp_all add: div_mult_mod_eq mult_div_mod_eq)
+
+end
+
+text \<open>Interlude: basic tool support for algebraic and arithmetic calculations\<close>
+
+named_theorems arith "arith facts -- only ground formulas"
+ML_file "Tools/arith_data.ML"
+
+ML_file "~~/src/Provers/Arith/cancel_div_mod.ML"
+
+ML \<open>
+structure Cancel_Div_Mod_Ring = Cancel_Div_Mod
+(
+ val div_name = @{const_name divide};
+ val mod_name = @{const_name modulo};
+ val mk_binop = HOLogic.mk_binop;
+ val mk_sum = Arith_Data.mk_sum;
+ val dest_sum = Arith_Data.dest_sum;
+
+ val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules};
+
+ val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
+ @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps})
+)
+\<close>
+
+simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") =
+ \<open>K Cancel_Div_Mod_Ring.proc\<close>
+
+class idom_modulo = idom + semidom_modulo
+begin
+
+subclass idom_divide ..
+
+lemma div_diff [simp]:
+ "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> (a - b) div c = a div c - b div c"
+ using div_add [of _ _ "- b"] by (simp add: dvd_neg_div)
+
+end
+
+
class ordered_semiring = semiring + ordered_comm_monoid_add +
assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
@@ -2208,6 +2326,10 @@
by (auto simp add: abs_mult)
qed
+lemma sgn_not_eq_imp:
+ "sgn a = - sgn b" if "sgn b \<noteq> sgn a" and "sgn a \<noteq> 0" and "sgn b \<noteq> 0"
+ using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg)
+
lemma abs_dvd_iff [simp]: "\<bar>m\<bar> dvd k \<longleftrightarrow> m dvd k"
by (simp add: abs_if)
@@ -2241,6 +2363,10 @@
lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \<longleftrightarrow> - 1 < a"
by (fact minus_less_iff)
+lemma add_less_zeroD:
+ shows "x+y < 0 \<Longrightarrow> x<0 \<or> y<0"
+ by (auto simp: not_less intro: le_less_trans [of _ "x+y"])
+
end
text \<open>Simprules for comparisons where common factors can be cancelled.\<close>
--- a/src/HOL/SMT.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/SMT.thy Mon Oct 09 22:08:05 2017 +0200
@@ -124,8 +124,14 @@
lemma int_nat_nneg: "\<forall>i. i \<ge> 0 \<longrightarrow> int (nat i) = i" by simp
lemma int_nat_neg: "\<forall>i. i < 0 \<longrightarrow> int (nat i) = 0" by simp
-lemmas nat_zero_as_int = transfer_nat_int_numerals(1)
-lemmas nat_one_as_int = transfer_nat_int_numerals(2)
+lemma nat_zero_as_int:
+ "0 = nat 0"
+ by simp
+
+lemma nat_one_as_int:
+ "1 = nat 1"
+ by simp
+
lemma nat_numeral_as_int: "numeral = (\<lambda>i. nat (numeral i))" by simp
lemma nat_less_as_int: "op < = (\<lambda>a b. int a < int b)" by simp
lemma nat_leq_as_int: "op \<le> = (\<lambda>a b. int a \<le> int b)" by simp
--- a/src/HOL/Set.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Set.thy Mon Oct 09 22:08:05 2017 +0200
@@ -1849,12 +1849,34 @@
definition pairwise :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
where "pairwise R S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> R x y)"
+lemma pairwiseI:
+ "pairwise R S" if "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> R x y"
+ using that by (simp add: pairwise_def)
+
+lemma pairwiseD:
+ "R x y" and "R y x"
+ if "pairwise R S" "x \<in> S" and "y \<in> S" and "x \<noteq> y"
+ using that by (simp_all add: pairwise_def)
+
+lemma pairwise_empty [simp]: "pairwise P {}"
+ by (simp add: pairwise_def)
+
+lemma pairwise_singleton [simp]: "pairwise P {A}"
+ by (simp add: pairwise_def)
+
+lemma pairwise_insert:
+ "pairwise r (insert x s) \<longleftrightarrow> (\<forall>y. y \<in> s \<and> y \<noteq> x \<longrightarrow> r x y \<and> r y x) \<and> pairwise r s"
+ by (force simp: pairwise_def)
+
lemma pairwise_subset: "pairwise P S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> pairwise P T"
by (force simp: pairwise_def)
lemma pairwise_mono: "\<lbrakk>pairwise P A; \<And>x y. P x y \<Longrightarrow> Q x y\<rbrakk> \<Longrightarrow> pairwise Q A"
by (auto simp: pairwise_def)
+lemma pairwise_image: "pairwise r (f ` s) \<longleftrightarrow> pairwise (\<lambda>x y. (f x \<noteq> f y) \<longrightarrow> r (f x) (f y)) s"
+ by (force simp: pairwise_def)
+
definition disjnt :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
where "disjnt A B \<longleftrightarrow> A \<inter> B = {}"
@@ -1882,19 +1904,6 @@
lemma disjnt_subset2 : "\<lbrakk>disjnt X Y; Z \<subseteq> Y\<rbrakk> \<Longrightarrow> disjnt X Z"
by (auto simp: disjnt_def)
-lemma pairwise_empty [simp]: "pairwise P {}"
- by (simp add: pairwise_def)
-
-lemma pairwise_singleton [simp]: "pairwise P {A}"
- by (simp add: pairwise_def)
-
-lemma pairwise_insert:
- "pairwise r (insert x s) \<longleftrightarrow> (\<forall>y. y \<in> s \<and> y \<noteq> x \<longrightarrow> r x y \<and> r y x) \<and> pairwise r s"
- by (force simp: pairwise_def)
-
-lemma pairwise_image: "pairwise r (f ` s) \<longleftrightarrow> pairwise (\<lambda>x y. (f x \<noteq> f y) \<longrightarrow> r (f x) (f y)) s"
- by (force simp: pairwise_def)
-
lemma disjoint_image_subset: "\<lbrakk>pairwise disjnt \<A>; \<And>X. X \<in> \<A> \<Longrightarrow> f X \<subseteq> X\<rbrakk> \<Longrightarrow> pairwise disjnt (f `\<A>)"
unfolding disjnt_def pairwise_def by fast
--- a/src/HOL/Tools/arith_data.ML Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Tools/arith_data.ML Mon Oct 09 22:08:05 2017 +0200
@@ -71,7 +71,7 @@
(*this version ALWAYS includes a trailing zero*)
fun long_mk_sum T [] = mk_number T 0
- | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts);
+ | long_mk_sum T (t :: ts) = mk_plus (t, long_mk_sum T ts);
(*decompose additions AND subtractions as a sum*)
fun dest_summing (pos, Const (@{const_name Groups.plus}, _) $ t $ u, ts) =
--- a/src/HOL/Tools/legacy_transfer.ML Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Tools/legacy_transfer.ML Mon Oct 09 22:08:05 2017 +0200
@@ -154,7 +154,6 @@
{ inj = inj1, embed = embed1, return = return1, cong = cong1, hints = hints1 }
{ inj = inj2, embed = embed2, return = return2, cong = cong2, hints = hints2 } =
let
- fun add_del eq del add = union eq add #> subtract eq del;
val guessed = if guess
then map (fn thm => transfer_thm
((a, D), (if null inj1 then inj2 else inj1, [], [], cong1)) [] ctxt thm RS sym) embed1
--- a/src/HOL/Tools/nat_numeral_simprocs.ML Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML Mon Oct 09 22:08:05 2017 +0200
@@ -41,17 +41,9 @@
handle TERM _ => find_first_numeral (t::past) terms)
| find_first_numeral past [] = raise TERM("find_first_numeral", []);
-val zero = mk_number 0;
-val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
+val mk_sum = Arith_Data.mk_sum HOLogic.natT;
-(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
-fun mk_sum [] = zero
- | mk_sum [t,u] = mk_plus (t, u)
- | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
-
-(*this version ALWAYS includes a trailing zero*)
-fun long_mk_sum [] = HOLogic.zero
- | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+val long_mk_sum = Arith_Data.long_mk_sum HOLogic.natT;
val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
--- a/src/HOL/Word/Word.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Word/Word.thy Mon Oct 09 22:08:05 2017 +0200
@@ -2024,7 +2024,7 @@
lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse]
-lemmas thd = refl [THEN [2] split_div_lemma [THEN iffD2], THEN conjunct1]
+lemmas thd = times_div_less_eq_dividend
lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend dtle
@@ -3724,7 +3724,7 @@
lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size]
(* alternative proof of word_rcat_rsplit *)
-lemmas tdle = iffD2 [OF split_div_lemma refl, THEN conjunct1]
+lemmas tdle = times_div_less_eq_dividend
lemmas dtle = xtr4 [OF tdle mult.commute]
lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w"
@@ -3734,7 +3734,7 @@
apply (simp_all add: word_size
refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]])
apply safe
- apply (erule xtr7, rule len_gt_0 [THEN dtle])+
+ apply (erule xtr7, rule dtle)+
done
lemma size_word_rsplit_rcat_size:
--- a/src/HOL/Word/Word_Miscellaneous.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/Word/Word_Miscellaneous.thy Mon Oct 09 22:08:05 2017 +0200
@@ -248,9 +248,7 @@
lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *)
-lemma int_mod_le: "0 \<le> a \<Longrightarrow> a mod n \<le> a"
- for a :: int
- by (fact Divides.semiring_numeral_div_class.mod_less_eq_dividend) (* FIXME: delete *)
+lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *)
lemma mod_add_if_z:
"x < z \<Longrightarrow> y < z \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> z \<Longrightarrow>
@@ -306,8 +304,8 @@
lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm]
lemmas dme = div_mult_mod_eq
-lemmas dtle = xtr3 [OF dme [symmetric] le_add1]
-lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle]
+lemmas dtle = div_times_less_eq_dividend
+lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend]
lemma td_gal: "0 < c \<Longrightarrow> a \<ge> b * c \<longleftrightarrow> a div c \<ge> b"
for a b c :: nat
@@ -318,15 +316,13 @@
lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified]
-lemma div_mult_le: "a div b * b \<le> a"
- for a b :: nat
- by (fact dtle)
+lemmas div_mult_le = div_times_less_eq_dividend
-lemmas sdl = split_div_lemma [THEN iffD1, symmetric]
+lemmas sdl = div_nat_eqI
lemma given_quot: "f > 0 \<Longrightarrow> (f * l + (f - 1)) div f = l"
for f l :: nat
- by (rule sdl, assumption) (simp (no_asm))
+ by (rule div_nat_eqI) (simp_all)
lemma given_quot_alt: "f > 0 \<Longrightarrow> (l * f + f - Suc 0) div f = l"
for f l :: nat
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Function_Growth.thy Mon Oct 09 22:08:05 2017 +0200
@@ -0,0 +1,457 @@
+
+(* Author: Florian Haftmann, TU Muenchen *)
+
+section \<open>Comparing growth of functions on natural numbers by a preorder relation\<close>
+
+theory Function_Growth
+imports
+ Main
+ "HOL-Library.Preorder"
+ "HOL-Library.Discrete"
+begin
+
+(* FIXME move *)
+
+context linorder
+begin
+
+lemma mono_invE:
+ fixes f :: "'a \<Rightarrow> 'b::order"
+ assumes "mono f"
+ assumes "f x < f y"
+ obtains "x < y"
+proof
+ show "x < y"
+ proof (rule ccontr)
+ assume "\<not> x < y"
+ then have "y \<le> x" by simp
+ with \<open>mono f\<close> obtain "f y \<le> f x" by (rule monoE)
+ with \<open>f x < f y\<close> show False by simp
+ qed
+qed
+
+end
+
+lemma (in semidom_divide) power_diff:
+ fixes a :: 'a
+ assumes "a \<noteq> 0"
+ assumes "m \<ge> n"
+ shows "a ^ (m - n) = (a ^ m) div (a ^ n)"
+proof -
+ define q where "q = m - n"
+ with assms have "m = q + n" by (simp add: q_def)
+ with q_def show ?thesis using \<open>a \<noteq> 0\<close> by (simp add: power_add)
+qed
+
+
+subsection \<open>Motivation\<close>
+
+text \<open>
+ When comparing growth of functions in computer science, it is common to adhere
+ on Landau Symbols (``O-Notation''). However these come at the cost of notational
+ oddities, particularly writing \<open>f = O(g)\<close> for \<open>f \<in> O(g)\<close> etc.
+
+ Here we suggest a different way, following Hardy (G.~H.~Hardy and J.~E.~Littlewood,
+ Some problems of Diophantine approximation, Acta Mathematica 37 (1914), p.~225).
+ We establish a quasi order relation \<open>\<lesssim>\<close> on functions such that
+ \<open>f \<lesssim> g \<longleftrightarrow> f \<in> O(g)\<close>. From a didactic point of view, this does not only
+ avoid the notational oddities mentioned above but also emphasizes the key insight
+ of a growth hierarchy of functions:
+ \<open>(\<lambda>n. 0) \<lesssim> (\<lambda>n. k) \<lesssim> Discrete.log \<lesssim> Discrete.sqrt \<lesssim> id \<lesssim> \<dots>\<close>.
+\<close>
+
+subsection \<open>Model\<close>
+
+text \<open>
+ Our growth functions are of type \<open>\<nat> \<Rightarrow> \<nat>\<close>. This is different
+ to the usual conventions for Landau symbols for which \<open>\<real> \<Rightarrow> \<real>\<close>
+ would be appropriate, but we argue that \<open>\<real> \<Rightarrow> \<real>\<close> is more
+ appropriate for analysis, whereas our setting is discrete.
+
+ Note that we also restrict the additional coefficients to \<open>\<nat>\<close>, something
+ we discuss at the particular definitions.
+\<close>
+
+subsection \<open>The \<open>\<lesssim>\<close> relation\<close>
+
+definition less_eq_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<lesssim>" 50)
+where
+ "f \<lesssim> g \<longleftrightarrow> (\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m)"
+
+text \<open>
+ This yields \<open>f \<lesssim> g \<longleftrightarrow> f \<in> O(g)\<close>. Note that \<open>c\<close> is restricted to
+ \<open>\<nat>\<close>. This does not pose any problems since if \<open>f \<in> O(g)\<close> holds for
+ a \<open>c \<in> \<real>\<close>, it also holds for \<open>\<lceil>c\<rceil> \<in> \<nat>\<close> by transitivity.
+\<close>
+
+lemma less_eq_funI [intro?]:
+ assumes "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m"
+ shows "f \<lesssim> g"
+ unfolding less_eq_fun_def by (rule assms)
+
+lemma not_less_eq_funI:
+ assumes "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * g m < f m"
+ shows "\<not> f \<lesssim> g"
+ using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
+
+lemma less_eq_funE [elim?]:
+ assumes "f \<lesssim> g"
+ obtains n c where "c > 0" and "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
+ using assms unfolding less_eq_fun_def by blast
+
+lemma not_less_eq_funE:
+ assumes "\<not> f \<lesssim> g" and "c > 0"
+ obtains m where "m > n" and "c * g m < f m"
+ using assms unfolding less_eq_fun_def linorder_not_le [symmetric] by blast
+
+
+subsection \<open>The \<open>\<approx>\<close> relation, the equivalence relation induced by \<open>\<lesssim>\<close>\<close>
+
+definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<cong>" 50)
+where
+ "f \<cong> g \<longleftrightarrow>
+ (\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m)"
+
+text \<open>
+ This yields \<open>f \<cong> g \<longleftrightarrow> f \<in> \<Theta>(g)\<close>. Concerning \<open>c\<^sub>1\<close> and \<open>c\<^sub>2\<close>
+ restricted to @{typ nat}, see note above on \<open>(\<lesssim>)\<close>.
+\<close>
+
+lemma equiv_funI:
+ assumes "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
+ shows "f \<cong> g"
+ unfolding equiv_fun_def by (rule assms)
+
+lemma not_equiv_funI:
+ assumes "\<And>c\<^sub>1 c\<^sub>2 n. c\<^sub>1 > 0 \<Longrightarrow> c\<^sub>2 > 0 \<Longrightarrow>
+ \<exists>m>n. c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
+ shows "\<not> f \<cong> g"
+ using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
+
+lemma equiv_funE:
+ assumes "f \<cong> g"
+ obtains n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
+ and "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
+ using assms unfolding equiv_fun_def by blast
+
+lemma not_equiv_funE:
+ fixes n c\<^sub>1 c\<^sub>2
+ assumes "\<not> f \<cong> g" and "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
+ obtains m where "m > n"
+ and "c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
+ using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
+
+
+subsection \<open>The \<open>\<prec>\<close> relation, the strict part of \<open>\<lesssim>\<close>\<close>
+
+definition less_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<prec>" 50)
+where
+ "f \<prec> g \<longleftrightarrow> f \<lesssim> g \<and> \<not> g \<lesssim> f"
+
+lemma less_funI:
+ assumes "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m"
+ and "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * f m < g m"
+ shows "f \<prec> g"
+ using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
+
+lemma not_less_funI:
+ assumes "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * g m < f m"
+ and "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m"
+ shows "\<not> f \<prec> g"
+ using assms unfolding less_fun_def less_eq_fun_def linorder_not_less [symmetric] by blast
+
+lemma less_funE [elim?]:
+ assumes "f \<prec> g"
+ obtains n c where "c > 0" and "\<And>m. m > n \<Longrightarrow> f m \<le> c * g m"
+ and "\<And>c n. c > 0 \<Longrightarrow> \<exists>m>n. c * f m < g m"
+proof -
+ from assms have "f \<lesssim> g" and "\<not> g \<lesssim> f" by (simp_all add: less_fun_def)
+ from \<open>f \<lesssim> g\<close> obtain n c where *:"c > 0" "m > n \<Longrightarrow> f m \<le> c * g m" for m
+ by (rule less_eq_funE) blast
+ { fix c n :: nat
+ assume "c > 0"
+ with \<open>\<not> g \<lesssim> f\<close> obtain m where "m > n" "c * f m < g m"
+ by (rule not_less_eq_funE) blast
+ then have **: "\<exists>m>n. c * f m < g m" by blast
+ } note ** = this
+ from * ** show thesis by (rule that)
+qed
+
+lemma not_less_funE:
+ assumes "\<not> f \<prec> g" and "c > 0"
+ obtains m where "m > n" and "c * g m < f m"
+ | d q where "\<And>m. d > 0 \<Longrightarrow> m > q \<Longrightarrow> g q \<le> d * f q"
+ using assms unfolding less_fun_def linorder_not_less [symmetric] by blast
+
+text \<open>
+ I did not find a proof for \<open>f \<prec> g \<longleftrightarrow> f \<in> o(g)\<close>. Maybe this only
+ holds if \<open>f\<close> and/or \<open>g\<close> are of a certain class of functions.
+ However \<open>f \<in> o(g) \<longrightarrow> f \<prec> g\<close> is provable, and this yields a
+ handy introduction rule.
+
+ Note that D. Knuth ignores \<open>o\<close> altogether. So what \dots
+
+ Something still has to be said about the coefficient \<open>c\<close> in
+ the definition of \<open>(\<prec>)\<close>. In the typical definition of \<open>o\<close>,
+ it occurs on the \emph{right} hand side of the \<open>(>)\<close>. The reason
+ is that the situation is dual to the definition of \<open>O\<close>: the definition
+ works since \<open>c\<close> may become arbitrary small. Since this is not possible
+ within @{term \<nat>}, we push the coefficient to the left hand side instead such
+ that it may become arbitrary big instead.
+\<close>
+
+lemma less_fun_strongI:
+ assumes "\<And>c. c > 0 \<Longrightarrow> \<exists>n. \<forall>m>n. c * f m < g m"
+ shows "f \<prec> g"
+proof (rule less_funI)
+ have "1 > (0::nat)" by simp
+ with assms [OF this] obtain n where *: "m > n \<Longrightarrow> 1 * f m < g m" for m
+ by blast
+ have "\<forall>m>n. f m \<le> 1 * g m"
+ proof (rule allI, rule impI)
+ fix m
+ assume "m > n"
+ with * have "1 * f m < g m" by simp
+ then show "f m \<le> 1 * g m" by simp
+ qed
+ with \<open>1 > 0\<close> show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
+ fix c n :: nat
+ assume "c > 0"
+ with assms obtain q where "m > q \<Longrightarrow> c * f m < g m" for m by blast
+ then have "c * f (Suc (q + n)) < g (Suc (q + n))" by simp
+ moreover have "Suc (q + n) > n" by simp
+ ultimately show "\<exists>m>n. c * f m < g m" by blast
+qed
+
+
+subsection \<open>\<open>\<lesssim>\<close> is a preorder\<close>
+
+text \<open>This yields all lemmas relating \<open>\<lesssim>\<close>, \<open>\<prec>\<close> and \<open>\<cong>\<close>.\<close>
+
+interpretation fun_order: preorder_equiv less_eq_fun less_fun
+ rewrites "fun_order.equiv = equiv_fun"
+proof -
+ interpret preorder: preorder_equiv less_eq_fun less_fun
+ proof
+ fix f g h
+ show "f \<lesssim> f"
+ proof
+ have "\<exists>n. \<forall>m>n. f m \<le> 1 * f m" by auto
+ then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * f m" by blast
+ qed
+ show "f \<prec> g \<longleftrightarrow> f \<lesssim> g \<and> \<not> g \<lesssim> f"
+ by (fact less_fun_def)
+ assume "f \<lesssim> g" and "g \<lesssim> h"
+ show "f \<lesssim> h"
+ proof
+ from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1
+ where "c\<^sub>1 > 0" and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m"
+ by rule blast
+ from \<open>g \<lesssim> h\<close> obtain n\<^sub>2 c\<^sub>2
+ where "c\<^sub>2 > 0" and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * h m"
+ by rule blast
+ have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m"
+ proof (rule allI, rule impI)
+ fix m
+ assume Q: "m > max n\<^sub>1 n\<^sub>2"
+ from P\<^sub>1 Q have *: "f m \<le> c\<^sub>1 * g m" by simp
+ from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * h m" by simp
+ with \<open>c\<^sub>1 > 0\<close> have "c\<^sub>1 * g m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by simp
+ with * show "f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by (rule order_trans)
+ qed
+ then have "\<exists>n. \<forall>m>n. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by rule
+ moreover from \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "c\<^sub>1 * c\<^sub>2 > 0" by simp
+ ultimately show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * h m" by blast
+ qed
+ qed
+ from preorder.preorder_equiv_axioms show "class.preorder_equiv less_eq_fun less_fun" .
+ show "preorder_equiv.equiv less_eq_fun = equiv_fun"
+ proof (rule ext, rule ext, unfold preorder.equiv_def)
+ fix f g
+ show "f \<lesssim> g \<and> g \<lesssim> f \<longleftrightarrow> f \<cong> g"
+ proof
+ assume "f \<cong> g"
+ then obtain n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
+ and *: "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
+ by (rule equiv_funE) blast
+ have "\<forall>m>n. f m \<le> c\<^sub>1 * g m"
+ proof (rule allI, rule impI)
+ fix m
+ assume "m > n"
+ with * show "f m \<le> c\<^sub>1 * g m" by simp
+ qed
+ with \<open>c\<^sub>1 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
+ then have "f \<lesssim> g" ..
+ have "\<forall>m>n. g m \<le> c\<^sub>2 * f m"
+ proof (rule allI, rule impI)
+ fix m
+ assume "m > n"
+ with * show "g m \<le> c\<^sub>2 * f m" by simp
+ qed
+ with \<open>c\<^sub>2 > 0\<close> have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
+ then have "g \<lesssim> f" ..
+ from \<open>f \<lesssim> g\<close> and \<open>g \<lesssim> f\<close> show "f \<lesssim> g \<and> g \<lesssim> f" ..
+ next
+ assume "f \<lesssim> g \<and> g \<lesssim> f"
+ then have "f \<lesssim> g" and "g \<lesssim> f" by auto
+ from \<open>f \<lesssim> g\<close> obtain n\<^sub>1 c\<^sub>1 where "c\<^sub>1 > 0"
+ and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m" by rule blast
+ from \<open>g \<lesssim> f\<close> obtain n\<^sub>2 c\<^sub>2 where "c\<^sub>2 > 0"
+ and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * f m" by rule blast
+ have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
+ proof (rule allI, rule impI)
+ fix m
+ assume Q: "m > max n\<^sub>1 n\<^sub>2"
+ from P\<^sub>1 Q have "f m \<le> c\<^sub>1 * g m" by simp
+ moreover from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * f m" by simp
+ ultimately show "f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" ..
+ qed
+ with \<open>c\<^sub>1 > 0\<close> \<open>c\<^sub>2 > 0\<close> have "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n.
+ \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" by blast
+ then show "f \<cong> g" by (rule equiv_funI)
+ qed
+ qed
+qed
+
+declare fun_order.antisym [intro?]
+
+
+subsection \<open>Simple examples\<close>
+
+text \<open>
+ Most of these are left as constructive exercises for the reader. Note that additional
+ preconditions to the functions may be necessary. The list here is by no means to be
+ intended as complete construction set for typical functions, here surely something
+ has to be added yet.
+\<close>
+
+text \<open>@{prop "(\<lambda>n. f n + k) \<cong> f"}\<close>
+
+lemma equiv_fun_mono_const:
+ assumes "mono f" and "\<exists>n. f n > 0"
+ shows "(\<lambda>n. f n + k) \<cong> f"
+proof (cases "k = 0")
+ case True then show ?thesis by simp
+next
+ case False
+ show ?thesis
+ proof
+ show "(\<lambda>n. f n + k) \<lesssim> f"
+ proof
+ from \<open>\<exists>n. f n > 0\<close> obtain n where "f n > 0" ..
+ have "\<forall>m>n. f m + k \<le> Suc k * f m"
+ proof (rule allI, rule impI)
+ fix m
+ assume "n < m"
+ with \<open>mono f\<close> have "f n \<le> f m"
+ using less_imp_le_nat monoE by blast
+ with \<open>0 < f n\<close> have "0 < f m" by auto
+ then obtain l where "f m = Suc l" by (cases "f m") simp_all
+ then show "f m + k \<le> Suc k * f m" by simp
+ qed
+ then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m + k \<le> c * f m" by blast
+ qed
+ show "f \<lesssim> (\<lambda>n. f n + k)"
+ proof
+ have "f m \<le> 1 * (f m + k)" for m by simp
+ then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * (f m + k)" by blast
+ qed
+ qed
+qed
+
+lemma
+ assumes "strict_mono f"
+ shows "(\<lambda>n. f n + k) \<cong> f"
+proof (rule equiv_fun_mono_const)
+ from assms show "mono f" by (rule strict_mono_mono)
+ show "\<exists>n. 0 < f n"
+ proof (rule ccontr)
+ assume "\<not> (\<exists>n. 0 < f n)"
+ then have "\<And>n. f n = 0" by simp
+ then have "f 0 = f 1" by simp
+ moreover from \<open>strict_mono f\<close> have "f 0 < f 1"
+ by (simp add: strict_mono_def)
+ ultimately show False by simp
+ qed
+qed
+
+lemma
+ "(\<lambda>n. Suc k * f n) \<cong> f"
+proof
+ show "(\<lambda>n. Suc k * f n) \<lesssim> f"
+ proof
+ have "Suc k * f m \<le> Suc k * f m" for m by simp
+ then show "\<exists>c>0. \<exists>n. \<forall>m>n. Suc k * f m \<le> c * f m" by blast
+ qed
+ show "f \<lesssim> (\<lambda>n. Suc k * f n)"
+ proof
+ have "f m \<le> 1 * (Suc k * f m)" for m by simp
+ then show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * (Suc k * f m)" by blast
+ qed
+qed
+
+lemma
+ "f \<lesssim> (\<lambda>n. f n + g n)"
+ by rule auto
+
+lemma
+ "(\<lambda>_. 0) \<prec> (\<lambda>n. Suc k)"
+ by (rule less_fun_strongI) auto
+
+lemma
+ "(\<lambda>_. k) \<prec> Discrete.log"
+proof (rule less_fun_strongI)
+ fix c :: nat
+ have "\<forall>m>2 ^ (Suc (c * k)). c * k < Discrete.log m"
+ proof (rule allI, rule impI)
+ fix m :: nat
+ assume "2 ^ Suc (c * k) < m"
+ then have "2 ^ Suc (c * k) \<le> m" by simp
+ with log_mono have "Discrete.log (2 ^ (Suc (c * k))) \<le> Discrete.log m"
+ by (blast dest: monoD)
+ moreover have "c * k < Discrete.log (2 ^ (Suc (c * k)))" by simp
+ ultimately show "c * k < Discrete.log m" by auto
+ qed
+ then show "\<exists>n. \<forall>m>n. c * k < Discrete.log m" ..
+qed
+
+(*lemma
+ "Discrete.log \<prec> Discrete.sqrt"
+proof (rule less_fun_strongI)*)
+text \<open>@{prop "Discrete.log \<prec> Discrete.sqrt"}\<close>
+
+lemma
+ "Discrete.sqrt \<prec> id"
+proof (rule less_fun_strongI)
+ fix c :: nat
+ assume "0 < c"
+ have "\<forall>m>(Suc c)\<^sup>2. c * Discrete.sqrt m < id m"
+ proof (rule allI, rule impI)
+ fix m
+ assume "(Suc c)\<^sup>2 < m"
+ then have "(Suc c)\<^sup>2 \<le> m" by simp
+ with mono_sqrt have "Discrete.sqrt ((Suc c)\<^sup>2) \<le> Discrete.sqrt m" by (rule monoE)
+ then have "Suc c \<le> Discrete.sqrt m" by simp
+ then have "c < Discrete.sqrt m" by simp
+ moreover from \<open>(Suc c)\<^sup>2 < m\<close> have "Discrete.sqrt m > 0" by simp
+ ultimately have "c * Discrete.sqrt m < Discrete.sqrt m * Discrete.sqrt m" by simp
+ also have "\<dots> \<le> m" by (simp add: power2_eq_square [symmetric])
+ finally show "c * Discrete.sqrt m < id m" by simp
+ qed
+ then show "\<exists>n. \<forall>m>n. c * Discrete.sqrt m < id m" ..
+qed
+
+lemma
+ "id \<prec> (\<lambda>n. n\<^sup>2)"
+ by (rule less_fun_strongI) (auto simp add: power2_eq_square)
+
+lemma
+ "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. n ^ Suc k)"
+ by (rule less_fun_strongI) auto
+
+(*lemma
+ "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"
+proof (rule less_fun_strongI)*)
+text \<open>@{prop "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. 2 ^ n)"}\<close>
+
+end
--- a/src/HOL/ex/Simproc_Tests.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/ex/Simproc_Tests.thy Mon Oct 09 22:08:05 2017 +0200
@@ -199,7 +199,7 @@
subsection \<open>\<open>int_div_cancel_numeral_factors\<close>\<close>
notepad begin
- fix x y z :: "'a::{semiring_div,comm_ring_1,ring_char_0}"
+ fix x y z :: "'a::{unique_euclidean_semiring,comm_ring_1,ring_char_0}"
{
assume "(3*x) div (4*y) = z" have "(9*x) div (12*y) = z"
by (tactic \<open>test @{context} [@{simproc int_div_cancel_numeral_factors}]\<close>) fact
@@ -325,7 +325,7 @@
subsection \<open>\<open>int_div_cancel_factor\<close>\<close>
notepad begin
- fix a b c d k uu x y :: "'a::semiring_div"
+ fix a b c d k uu x y :: "'a::unique_euclidean_semiring"
{
assume "(if k = 0 then 0 else x div y) = uu"
have "(x*k) div (k*y) = uu"
--- a/src/HOL/ex/Word_Type.thy Mon Oct 09 22:03:05 2017 +0200
+++ b/src/HOL/ex/Word_Type.thy Mon Oct 09 22:08:05 2017 +0200
@@ -11,7 +11,7 @@
subsection \<open>Truncating bit representations of numeric types\<close>
-class semiring_bits = semiring_div_parity +
+class semiring_bits = semiring_parity +
assumes semiring_bits: "(1 + 2 * a) mod of_nat (2 * n) = 1 + 2 * (a mod of_nat n)"
begin
@@ -27,28 +27,16 @@
by (simp add: bitrunc_eq_mod)
lemma bitrunc_Suc [simp]:
- "bitrunc (Suc n) a = bitrunc n (a div 2) * 2 + a mod 2"
+ "bitrunc (Suc n) a = bitrunc n (a div 2) * 2 + of_bool (odd a)"
proof -
- define b and c
- where "b = a div 2" and "c = a mod 2"
- then have a: "a = b * 2 + c"
- and "c = 0 \<or> c = 1"
- by (simp_all add: div_mult_mod_eq parity)
- from \<open>c = 0 \<or> c = 1\<close>
- have "bitrunc (Suc n) (b * 2 + c) = bitrunc n b * 2 + c"
- proof
- assume "c = 0"
- moreover have "(2 * b) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n)"
- by (simp add: mod_mult_mult1)
- ultimately show ?thesis
- by (simp add: bitrunc_eq_mod ac_simps)
- next
- assume "c = 1"
- with semiring_bits [of b "2 ^ n"] show ?thesis
- by (simp add: bitrunc_eq_mod ac_simps)
- qed
- with a show ?thesis
- by (simp add: b_def c_def)
+ have "1 + 2 * (a div 2) mod (2 * 2 ^ n) = (a div 2 * 2 + a mod 2) mod (2 * 2 ^ n)"
+ if "odd a"
+ using that semiring_bits [of "a div 2" "2 ^ n"]
+ by (simp add: algebra_simps odd_iff_mod_2_eq_one mult_mod_right)
+ also have "\<dots> = a mod (2 * 2 ^ n)"
+ by (simp only: div_mult_mod_eq)
+ finally show ?thesis
+ by (simp add: bitrunc_eq_mod algebra_simps mult_mod_right)
qed
lemma bitrunc_of_0 [simp]:
@@ -57,11 +45,11 @@
lemma bitrunc_plus:
"bitrunc n (bitrunc n a + bitrunc n b) = bitrunc n (a + b)"
- by (simp add: bitrunc_eq_mod mod_add_eq)
+ by (simp add: bitrunc_eq_mod mod_simps)
lemma bitrunc_of_1_eq_0_iff [simp]:
"bitrunc n 1 = 0 \<longleftrightarrow> n = 0"
- by (induct n) simp_all
+ by (simp add: bitrunc_eq_mod)
end
@@ -113,7 +101,7 @@
lemma signed_bitrunc_Suc [simp]:
"signed_bitrunc (Suc n) k = signed_bitrunc n (k div 2) * 2 + k mod 2"
- using zero_not_eq_two by (simp add: signed_bitrunc_eq_bitrunc algebra_simps)
+ by (simp add: odd_iff_mod_2_eq_one signed_bitrunc_eq_bitrunc algebra_simps)
lemma signed_bitrunc_of_0 [simp]:
"signed_bitrunc n 0 = 0"
--- a/src/Provers/Arith/cancel_div_mod.ML Mon Oct 09 22:03:05 2017 +0200
+++ b/src/Provers/Arith/cancel_div_mod.ML Mon Oct 09 22:08:05 2017 +0200
@@ -15,7 +15,7 @@
val div_name: string
val mod_name: string
val mk_binop: string -> term * term -> term
- val mk_sum: term list -> term
+ val mk_sum: typ -> term list -> term
val dest_sum: term -> term list
(*logic*)
val div_mod_eqs: thm list
@@ -34,16 +34,16 @@
functor Cancel_Div_Mod(Data: CANCEL_DIV_MOD_DATA): CANCEL_DIV_MOD =
struct
-fun coll_div_mod (Const(@{const_name Groups.plus},_) $ s $ t) dms =
+fun coll_div_mod (Const (@{const_name Groups.plus}, _) $ s $ t) dms =
coll_div_mod t (coll_div_mod s dms)
- | coll_div_mod (Const(@{const_name Groups.times},_) $ m $ (Const(d,_) $ s $ n))
- (dms as (divs,mods)) =
- if d = Data.div_name andalso m=n then ((s,n)::divs,mods) else dms
- | coll_div_mod (Const(@{const_name Groups.times},_) $ (Const(d,_) $ s $ n) $ m)
- (dms as (divs,mods)) =
- if d = Data.div_name andalso m=n then ((s,n)::divs,mods) else dms
- | coll_div_mod (Const(m,_) $ s $ n) (dms as (divs,mods)) =
- if m = Data.mod_name then (divs,(s,n)::mods) else dms
+ | coll_div_mod (Const (@{const_name Groups.times}, _) $ m $ (Const (d, _) $ s $ n))
+ (dms as (divs, mods)) =
+ if d = Data.div_name andalso m = n then ((s, n) :: divs, mods) else dms
+ | coll_div_mod (Const (@{const_name Groups.times}, _) $ (Const (d, _) $ s $ n) $ m)
+ (dms as (divs, mods)) =
+ if d = Data.div_name andalso m = n then ((s, n) :: divs, mods) else dms
+ | coll_div_mod (Const (m, _) $ s $ n) (dms as (divs, mods)) =
+ if m = Data.mod_name then (divs, (s, n) :: mods) else dms
| coll_div_mod _ dms = dms;
@@ -58,16 +58,18 @@
val mk_plus = Data.mk_binop @{const_name Groups.plus};
val mk_times = Data.mk_binop @{const_name Groups.times};
-fun rearrange t pq =
- let val ts = Data.dest_sum t;
- val dpq = Data.mk_binop Data.div_name pq
- val d1 = mk_times (snd pq,dpq) and d2 = mk_times (dpq,snd pq)
- val d = if member (op =) ts d1 then d1 else d2
- val m = Data.mk_binop Data.mod_name pq
- in mk_plus(mk_plus(d,m),Data.mk_sum(ts |> remove (op =) d |> remove (op =) m)) end
+fun rearrange T t pq =
+ let
+ val ts = Data.dest_sum t;
+ val dpq = Data.mk_binop Data.div_name pq;
+ val d1 = mk_times (snd pq, dpq) and d2 = mk_times (dpq, snd pq);
+ val d = if member (op =) ts d1 then d1 else d2;
+ val m = Data.mk_binop Data.mod_name pq;
+ in mk_plus (mk_plus (d, m), Data.mk_sum T (ts |> remove (op =) d |> remove (op =) m)) end
fun cancel ctxt t pq =
- let val teqt' = Data.prove_eq_sums ctxt (t, rearrange t pq)
+ let
+ val teqt' = Data.prove_eq_sums ctxt (t, rearrange (fastype_of t) t pq);
in hd (Data.div_mod_eqs RL [teqt' RS transitive_thm]) end;
fun proc ctxt ct =