diff --git a/src/HOL/Combinatorics/Orbits.thy b/src/HOL/Combinatorics/Orbits.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Combinatorics/Orbits.thy @@ -0,0 +1,594 @@ +(* Author: Lars Noschinski +*) + +section \Permutation orbits\ + +theory Orbits +imports + "HOL-Library.FuncSet" + "HOL-Combinatorics.Permutations" +begin + +subsection \Orbits and cyclic permutations\ + +inductive_set orbit :: "('a \ 'a) \ 'a \ 'a set" for f x where + base: "f x \ orbit f x" | + step: "y \ orbit f x \ f y \ orbit f x" + +definition cyclic_on :: "('a \ 'a) \ 'a set \ bool" where + "cyclic_on f S \ (\s\S. S = orbit f s)" + +lemma orbit_altdef: "orbit f x = {(f ^^ n) x | n. 0 < n}" (is "?L = ?R") +proof (intro set_eqI iffI) + fix y assume "y \ ?L" then show "y \ ?R" + by (induct rule: orbit.induct) (auto simp: exI[where x=1] exI[where x="Suc n" for n]) +next + fix y assume "y \ ?R" + then obtain n where "y = (f ^^ n) x" "0 < n" by blast + then show "y \ ?L" + proof (induction n arbitrary: y) + case (Suc n) then show ?case by (cases "n = 0") (auto intro: orbit.intros) + qed simp +qed + +lemma orbit_trans: + assumes "s \ orbit f t" "t \ orbit f u" shows "s \ orbit f u" + using assms by induct (auto intro: orbit.intros) + +lemma orbit_subset: + assumes "s \ orbit f (f t)" shows "s \ orbit f t" + using assms by (induct) (auto intro: orbit.intros) + +lemma orbit_sim_step: + assumes "s \ orbit f t" shows "f s \ orbit f (f t)" + using assms by induct (auto intro: orbit.intros) + +lemma orbit_step: + assumes "y \ orbit f x" "f x \ y" shows "y \ orbit f (f x)" + using assms +proof induction + case (step y) then show ?case by (cases "x = y") (auto intro: orbit.intros) +qed simp + +lemma self_in_orbit_trans: + assumes "s \ orbit f s" "t \ orbit f s" shows "t \ orbit f t" + using assms(2,1) by induct (auto intro: orbit_sim_step) + +lemma orbit_swap: + assumes "s \ orbit f s" "t \ orbit f s" shows "s \ orbit f t" + using assms(2,1) +proof induction + case base then show ?case by (cases "f s = s") (auto intro: orbit_step) +next + case (step x) then show ?case by (cases "f x = s") (auto intro: orbit_step) +qed + +lemma permutation_self_in_orbit: + assumes "permutation f" shows "s \ orbit f s" + unfolding orbit_altdef using permutation_self[OF assms, of s] by simp metis + +lemma orbit_altdef_self_in: + assumes "s \ orbit f s" shows "orbit f s = {(f ^^ n) s | n. True}" +proof (intro set_eqI iffI) + fix x assume "x \ {(f ^^ n) s | n. True}" + then obtain n where "x = (f ^^ n) s" by auto + then show "x \ orbit f s" using assms by (cases "n = 0") (auto simp: orbit_altdef) +qed (auto simp: orbit_altdef) + +lemma orbit_altdef_permutation: + assumes "permutation f" shows "orbit f s = {(f ^^ n) s | n. True}" + using assms by (intro orbit_altdef_self_in permutation_self_in_orbit) + +lemma orbit_altdef_bounded: + assumes "(f ^^ n) s = s" "0 < n" shows "orbit f s = {(f ^^ m) s| m. m < n}" +proof - + from assms have "s \ orbit f s" + by (auto simp add: orbit_altdef) metis + then have "orbit f s = {(f ^^ m) s|m. True}" by (rule orbit_altdef_self_in) + also have "\ = {(f ^^ m) s| m. m < n}" + using assms + by (auto simp: funpow_mod_eq intro: exI[where x="m mod n" for m]) + finally show ?thesis . +qed + +lemma funpow_in_orbit: + assumes "s \ orbit f t" shows "(f ^^ n) s \ orbit f t" + using assms by (induct n) (auto intro: orbit.intros) + +lemma finite_orbit: + assumes "s \ orbit f s" shows "finite (orbit f s)" +proof - + from assms obtain n where n: "0 < n" "(f ^^ n) s = s" + by (auto simp: orbit_altdef) + then show ?thesis by (auto simp: orbit_altdef_bounded) +qed + +lemma self_in_orbit_step: + assumes "s \ orbit f s" shows "orbit f (f s) = orbit f s" +proof (intro set_eqI iffI) + fix t assume "t \ orbit f s" then show "t \ orbit f (f s)" + using assms by (auto intro: orbit_step orbit_sim_step) +qed (auto intro: orbit_subset) + +lemma permutation_orbit_step: + assumes "permutation f" shows "orbit f (f s) = orbit f s" + using assms by (intro self_in_orbit_step permutation_self_in_orbit) + +lemma orbit_nonempty: + "orbit f s \ {}" + using orbit.base by fastforce + +lemma orbit_inv_eq: + assumes "permutation f" + shows "orbit (inv f) x = orbit f x" (is "?L = ?R") +proof - + { fix g y assume A: "permutation g" "y \ orbit (inv g) x" + have "y \ orbit g x" + proof - + have inv_g: "\y. x = g y \ inv g x = y" "\y. inv g (g y) = y" + by (metis A(1) bij_inv_eq_iff permutation_bijective)+ + + { fix y assume "y \ orbit g x" + then have "inv g y \ orbit g x" + by (cases) (simp_all add: inv_g A(1) permutation_self_in_orbit) + } note inv_g_in_orb = this + + from A(2) show ?thesis + by induct (simp_all add: inv_g_in_orb A permutation_self_in_orbit) + qed + } note orb_inv_ss = this + + have "inv (inv f) = f" + by (simp add: assms inv_inv_eq permutation_bijective) + then show ?thesis + using orb_inv_ss[OF assms] orb_inv_ss[OF permutation_inverse[OF assms]] by auto +qed + +lemma cyclic_on_alldef: + "cyclic_on f S \ S \ {} \ (\s\S. S = orbit f s)" + unfolding cyclic_on_def by (auto intro: orbit.step orbit_swap orbit_trans) + +lemma cyclic_on_funpow_in: + assumes "cyclic_on f S" "s \ S" shows "(f^^n) s \ S" + using assms unfolding cyclic_on_def by (auto intro: funpow_in_orbit) + +lemma finite_cyclic_on: + assumes "cyclic_on f S" shows "finite S" + using assms by (auto simp: cyclic_on_def finite_orbit) + +lemma cyclic_on_singleI: + assumes "s \ S" "S = orbit f s" shows "cyclic_on f S" + using assms unfolding cyclic_on_def by blast + +lemma cyclic_on_inI: + assumes "cyclic_on f S" "s \ S" shows "f s \ S" + using assms by (auto simp: cyclic_on_def intro: orbit.intros) + +lemma orbit_inverse: + assumes self: "a \ orbit g a" + and eq: "\x. x \ orbit g a \ g' (f x) = f (g x)" + shows "f ` orbit g a = orbit g' (f a)" (is "?L = ?R") +proof (intro set_eqI iffI) + fix x assume "x \ ?L" + then obtain x0 where "x0 \ orbit g a" "x = f x0" by auto + then show "x \ ?R" + proof (induct arbitrary: x) + case base then show ?case by (auto simp: self orbit.base eq[symmetric]) + next + case step then show ?case by cases (auto simp: eq[symmetric] orbit.intros) + qed +next + fix x assume "x \ ?R" + then show "x \ ?L" + proof (induct arbitrary: ) + case base then show ?case by (auto simp: self orbit.base eq) + next + case step then show ?case by cases (auto simp: eq orbit.intros) + qed +qed + +lemma cyclic_on_image: + assumes "cyclic_on f S" + assumes "\x. x \ S \ g (h x) = h (f x)" + shows "cyclic_on g (h ` S)" + using assms by (auto simp: cyclic_on_def) (meson orbit_inverse) + +lemma cyclic_on_f_in: + assumes "f permutes S" "cyclic_on f A" "f x \ A" + shows "x \ A" +proof - + from assms have fx_in_orb: "f x \ orbit f (f x)" by (auto simp: cyclic_on_alldef) + from assms have "A = orbit f (f x)" by (auto simp: cyclic_on_alldef) + moreover + then have "\ = orbit f x" using \f x \ A\ by (auto intro: orbit_step orbit_subset) + ultimately + show ?thesis by (metis (no_types) orbit.simps permutes_inverses(2)[OF assms(1)]) +qed + +lemma orbit_cong0: + assumes "x \ A" "f \ A \ A" "\y. y \ A \ f y = g y" shows "orbit f x = orbit g x" +proof - + { fix n have "(f ^^ n) x = (g ^^ n) x \ (f ^^ n) x \ A" + by (induct n rule: nat.induct) (insert assms, auto) + } then show ?thesis by (auto simp: orbit_altdef) +qed + +lemma orbit_cong: + assumes self_in: "t \ orbit f t" and eq: "\s. s \ orbit f t \ g s = f s" + shows "orbit g t = orbit f t" + using assms(1) _ assms(2) by (rule orbit_cong0) (auto simp: orbit.step eq) + +lemma cyclic_cong: + assumes "\s. s \ S \ f s = g s" shows "cyclic_on f S = cyclic_on g S" +proof - + have "(\s\S. orbit f s = orbit g s) \ cyclic_on f S = cyclic_on g S" + by (metis cyclic_on_alldef cyclic_on_def) + then show ?thesis by (metis assms orbit_cong cyclic_on_def) +qed + +lemma permutes_comp_preserves_cyclic1: + assumes "g permutes B" "cyclic_on f C" + assumes "A \ B = {}" "C \ A" + shows "cyclic_on (f o g) C" +proof - + have *: "\c. c \ C \ f (g c) = f c" + using assms by (subst permutes_not_in [of g]) auto + with assms(2) show ?thesis by (simp cong: cyclic_cong) +qed + +lemma permutes_comp_preserves_cyclic2: + assumes "f permutes A" "cyclic_on g C" + assumes "A \ B = {}" "C \ B" + shows "cyclic_on (f o g) C" +proof - + obtain c where c: "c \ C" "C = orbit g c" "c \ orbit g c" + using \cyclic_on g C\ by (auto simp: cyclic_on_def) + then have "\c. c \ C \ f (g c) = g c" + using assms c by (subst permutes_not_in [of f]) (auto intro: orbit.intros) + with assms(2) show ?thesis by (simp cong: cyclic_cong) +qed + +lemma permutes_orbit_subset: + assumes "f permutes S" "x \ S" shows "orbit f x \ S" +proof + fix y assume "y \ orbit f x" + then show "y \ S" by induct (auto simp: permutes_in_image assms) +qed + +lemma cyclic_on_orbit': + assumes "permutation f" shows "cyclic_on f (orbit f x)" + unfolding cyclic_on_alldef using orbit_nonempty[of f x] + by (auto intro: assms orbit_swap orbit_trans permutation_self_in_orbit) + +lemma cyclic_on_orbit: + assumes "f permutes S" "finite S" shows "cyclic_on f (orbit f x)" + using assms by (intro cyclic_on_orbit') (auto simp: permutation_permutes) + +lemma orbit_cyclic_eq3: + assumes "cyclic_on f S" "y \ S" shows "orbit f y = S" + using assms unfolding cyclic_on_alldef by simp + +lemma orbit_eq_singleton_iff: "orbit f x = {x} \ f x = x" (is "?L \ ?R") +proof + assume A: ?R + { fix y assume "y \ orbit f x" then have "y = x" + by induct (auto simp: A) + } then show ?L by (metis orbit_nonempty singletonI subsetI subset_singletonD) +next + assume A: ?L + then have "\y. y \ orbit f x \ f x = y" + by - (erule orbit.cases, simp_all) + then show ?R using A by blast +qed + +lemma eq_on_cyclic_on_iff1: + assumes "cyclic_on f S" "x \ S" + obtains "f x \ S" "f x = x \ card S = 1" +proof + from assms show "f x \ S" by (auto simp: cyclic_on_def intro: orbit.intros) + from assms have "S = orbit f x" by (auto simp: cyclic_on_alldef) + then have "f x = x \ S = {x}" by (metis orbit_eq_singleton_iff) + then show "f x = x \ card S = 1" using \x \ S\ by (auto simp: card_Suc_eq) +qed + +lemma orbit_eqI: + "y = f x \ y \ orbit f x" + "z = f y \y \ orbit f x \z \ orbit f x" + by (metis orbit.base) (metis orbit.step) + + +subsection \Decomposition of arbitrary permutations\ + +definition perm_restrict :: "('a \ 'a) \ 'a set \ ('a \ 'a)" where + "perm_restrict f S x \ if x \ S then f x else x" + +lemma perm_restrict_comp: + assumes "A \ B = {}" "cyclic_on f B" + shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" +proof - + have "\x. x \ B \ f x \ B" using \cyclic_on f B\ by (rule cyclic_on_inI) + with assms show ?thesis by (auto simp: perm_restrict_def fun_eq_iff) +qed + +lemma perm_restrict_simps: + "x \ S \ perm_restrict f S x = f x" + "x \ S \ perm_restrict f S x = x" + by (auto simp: perm_restrict_def) + +lemma perm_restrict_perm_restrict: + "perm_restrict (perm_restrict f A) B = perm_restrict f (A \ B)" + by (auto simp: perm_restrict_def) + +lemma perm_restrict_union: + assumes "perm_restrict f A permutes A" "perm_restrict f B permutes B" "A \ B = {}" + shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \ B)" + using assms by (auto simp: fun_eq_iff perm_restrict_def permutes_def) (metis Diff_iff Diff_triv) + +lemma perm_restrict_id[simp]: + assumes "f permutes S" shows "perm_restrict f S = f" + using assms by (auto simp: permutes_def perm_restrict_def) + +lemma cyclic_on_perm_restrict: + "cyclic_on (perm_restrict f S) S \ cyclic_on f S" + by (simp add: perm_restrict_def cong: cyclic_cong) + +lemma perm_restrict_diff_cyclic: + assumes "f permutes S" "cyclic_on f A" + shows "perm_restrict f (S - A) permutes (S - A)" +proof - + { fix y + have "\x. perm_restrict f (S - A) x = y" + proof cases + assume A: "y \ S - A" + with \f permutes S\ obtain x where "f x = y" "x \ S" + unfolding permutes_def by auto metis + moreover + with A have "x \ A" by (metis Diff_iff assms(2) cyclic_on_inI) + ultimately + have "perm_restrict f (S - A) x = y" by (simp add: perm_restrict_simps) + then show ?thesis .. + next + assume "y \ S - A" + then have "perm_restrict f (S - A) y = y" by (simp add: perm_restrict_simps) + then show ?thesis .. + qed + } note X = this + + { fix x y assume "perm_restrict f (S - A) x = perm_restrict f (S - A) y" + with assms have "x = y" + by (auto simp: perm_restrict_def permutes_def split: if_splits intro: cyclic_on_f_in) + } note Y = this + + show ?thesis by (auto simp: permutes_def perm_restrict_simps X intro: Y) +qed + +lemma permutes_decompose: + assumes "f permutes S" "finite S" + shows "\C. (\c \ C. cyclic_on f c) \ \C = S \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {})" + using assms(2,1) +proof (induction arbitrary: f rule: finite_psubset_induct) + case (psubset S) + + show ?case + proof (cases "S = {}") + case True then show ?thesis by (intro exI[where x="{}"]) auto + next + case False + then obtain s where "s \ S" by auto + with \f permutes S\ have "orbit f s \ S" + by (rule permutes_orbit_subset) + have cyclic_orbit: "cyclic_on f (orbit f s)" + using \f permutes S\ \finite S\ by (rule cyclic_on_orbit) + + let ?f' = "perm_restrict f (S - orbit f s)" + + have "f s \ S" using \f permutes S\ \s \ S\ by (auto simp: permutes_in_image) + then have "S - orbit f s \ S" using orbit.base[of f s] \s \ S\ by blast + moreover + have "?f' permutes (S - orbit f s)" + using \f permutes S\ cyclic_orbit by (rule perm_restrict_diff_cyclic) + ultimately + obtain C where C: "\c. c \ C \ cyclic_on ?f' c" "\C = S - orbit f s" + "\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}" + using psubset.IH by metis + + { fix c assume "c \ C" + then have *: "\x. x \ c \ perm_restrict f (S - orbit f s) x = f x" + using C(2) \f permutes S\ by (auto simp add: perm_restrict_def) + then have "cyclic_on f c" using C(1)[OF \c \ C\] by (simp cong: cyclic_cong add: *) + } note in_C_cyclic = this + + have Un_ins: "\(insert (orbit f s) C) = S" + using \\C = _\ \orbit f s \ S\ by blast + + have Disj_ins: "(\c1 \ insert (orbit f s) C. \c2 \ insert (orbit f s) C. c1 \ c2 \ c1 \ c2 = {})" + using C by auto + + show ?thesis + by (intro conjI Un_ins Disj_ins exI[where x="insert (orbit f s) C"]) + (auto simp: cyclic_orbit in_C_cyclic) + qed +qed + + +subsection \Function-power distance between values\ + +definition funpow_dist :: "('a \ 'a) \ 'a \ 'a \ nat" where + "funpow_dist f x y \ LEAST n. (f ^^ n) x = y" + +abbreviation funpow_dist1 :: "('a \ 'a) \ 'a \ 'a \ nat" where + "funpow_dist1 f x y \ Suc (funpow_dist f (f x) y)" + +lemma funpow_dist_0: + assumes "x = y" shows "funpow_dist f x y = 0" + using assms unfolding funpow_dist_def by (intro Least_eq_0) simp + +lemma funpow_dist_least: + assumes "n < funpow_dist f x y" shows "(f ^^ n) x \ y" +proof (rule notI) + assume "(f ^^ n) x = y" + then have "funpow_dist f x y \ n" unfolding funpow_dist_def by (rule Least_le) + with assms show False by linarith +qed + +lemma funpow_dist1_least: + assumes "0 < n" "n < funpow_dist1 f x y" shows "(f ^^ n) x \ y" +proof (rule notI) + assume "(f ^^ n) x = y" + then have "(f ^^ (n - 1)) (f x) = y" + using \0 < n\ by (cases n) (simp_all add: funpow_swap1) + then have "funpow_dist f (f x) y \ n - 1" unfolding funpow_dist_def by (rule Least_le) + with assms show False by simp +qed + +lemma funpow_dist_prop: + "y \ orbit f x \ (f ^^ funpow_dist f x y) x = y" + unfolding funpow_dist_def by (rule LeastI_ex) (auto simp: orbit_altdef) + +lemma funpow_dist_0_eq: + assumes "y \ orbit f x" shows "funpow_dist f x y = 0 \ x = y" + using assms by (auto simp: funpow_dist_0 dest: funpow_dist_prop) + +lemma funpow_dist_step: + assumes "x \ y" "y \ orbit f x" shows "funpow_dist f x y = Suc (funpow_dist f (f x) y)" +proof - + from \y \ _\ obtain n where "(f ^^ n) x = y" by (auto simp: orbit_altdef) + with \x \ y\ obtain n' where [simp]: "n = Suc n'" by (cases n) auto + + show ?thesis + unfolding funpow_dist_def + proof (rule Least_Suc2) + show "(f ^^ n) x = y" by fact + then show "(f ^^ n') (f x) = y" by (simp add: funpow_swap1) + show "(f ^^ 0) x \ y" using \x \ y\ by simp + show "\k. ((f ^^ Suc k) x = y) = ((f ^^ k) (f x) = y)" + by (simp add: funpow_swap1) + qed +qed + +lemma funpow_dist1_prop: + assumes "y \ orbit f x" shows "(f ^^ funpow_dist1 f x y) x = y" + by (metis assms funpow_dist_prop funpow_dist_step funpow_simps_right(2) o_apply self_in_orbit_step) + +(*XXX simplify? *) +lemma funpow_neq_less_funpow_dist: + assumes "y \ orbit f x" "m \ funpow_dist f x y" "n \ funpow_dist f x y" "m \ n" + shows "(f ^^ m) x \ (f ^^ n) x" +proof (rule notI) + assume A: "(f ^^ m) x = (f ^^ n) x" + + define m' n' where "m' = min m n" and "n' = max m n" + with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' \ funpow_dist f x y" + by (auto simp: min_def max_def) + + have "y = (f ^^ funpow_dist f x y) x" + using \y \ _\ by (simp only: funpow_dist_prop) + also have "\ = (f ^^ ((funpow_dist f x y - n') + n')) x" + using \n' \ _\ by simp + also have "\ = (f ^^ ((funpow_dist f x y - n') + m')) x" + by (simp add: funpow_add \(f ^^ m') x = _\) + also have "(f ^^ ((funpow_dist f x y - n') + m')) x \ y" + using A' by (intro funpow_dist_least) linarith + finally show "False" by simp +qed + +(* XXX reduce to funpow_neq_less_funpow_dist? *) +lemma funpow_neq_less_funpow_dist1: + assumes "y \ orbit f x" "m < funpow_dist1 f x y" "n < funpow_dist1 f x y" "m \ n" + shows "(f ^^ m) x \ (f ^^ n) x" +proof (rule notI) + assume A: "(f ^^ m) x = (f ^^ n) x" + + define m' n' where "m' = min m n" and "n' = max m n" + with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' < funpow_dist1 f x y" + by (auto simp: min_def max_def) + + have "y = (f ^^ funpow_dist1 f x y) x" + using \y \ _\ by (simp only: funpow_dist1_prop) + also have "\ = (f ^^ ((funpow_dist1 f x y - n') + n')) x" + using \n' < _\ by simp + also have "\ = (f ^^ ((funpow_dist1 f x y - n') + m')) x" + by (simp add: funpow_add \(f ^^ m') x = _\) + also have "(f ^^ ((funpow_dist1 f x y - n') + m')) x \ y" + using A' by (intro funpow_dist1_least) linarith+ + finally show "False" by simp +qed + +lemma inj_on_funpow_dist: + assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0..funpow_dist f x y}" + using funpow_neq_less_funpow_dist[OF assms] by (intro inj_onI) auto + +lemma inj_on_funpow_dist1: + assumes "y \ orbit f x" shows "inj_on (\n. (f ^^ n) x) {0.. orbit f x" + shows "orbit f x = (\n. (f ^^ n) x) ` {0.. orbit f x" by (auto simp: orbit_altdef) + then show ?thesis by (rule funpow_dist1_prop) +qed + +lemma funpow_dist1_dist: + assumes "funpow_dist1 f x y < funpow_dist1 f x z" + assumes "{y,z} \ orbit f x" + shows "funpow_dist1 f x z = funpow_dist1 f x y + funpow_dist1 f y z" (is "?L = ?R") +proof - + define n where \n = funpow_dist1 f x z - funpow_dist1 f x y - 1\ + with assms have *: \funpow_dist1 f x z = Suc (funpow_dist1 f x y + n)\ + by simp + have x_z: "(f ^^ funpow_dist1 f x z) x = z" using assms by (blast intro: funpow_dist1_prop) + have x_y: "(f ^^ funpow_dist1 f x y) x = y" using assms by (blast intro: funpow_dist1_prop) + + have "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y + = (f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) ((f ^^ funpow_dist1 f x y) x)" + using x_y by simp + also have "\ = z" + using assms x_z by (simp add: * funpow_add ac_simps funpow_swap1) + finally have y_z_diff: "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y = z" . + then have "(f ^^ funpow_dist1 f y z) y = z" + using assms by (intro funpow_dist1_prop1) auto + then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" + using x_y by simp + then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" + by (simp add: * funpow_add funpow_swap1) + show ?thesis + proof (rule antisym) + from y_z_diff have "(f ^^ funpow_dist1 f y z) y = z" + using assms by (intro funpow_dist1_prop1) auto + then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z" + using x_y by simp + then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z" + by (simp add: * funpow_add funpow_swap1) + then have "funpow_dist1 f x z \ funpow_dist1 f y z + funpow_dist1 f x y" + using funpow_dist1_least not_less by fastforce + then show "?L \ ?R" by presburger + next + have "funpow_dist1 f y z \ funpow_dist1 f x z - funpow_dist1 f x y" + using y_z_diff assms(1) by (metis not_less zero_less_diff funpow_dist1_least) + then show "?R \ ?L" by linarith + qed +qed + +lemma funpow_dist1_le_self: + assumes "(f ^^ m) x = x" "0 < m" "y \ orbit f x" + shows "funpow_dist1 f x y \ m" +proof (cases "x = y") + case True with assms show ?thesis by (auto dest!: funpow_dist1_least) +next + case False + have "(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x" + using assms by (simp add: funpow_mod_eq) + with False \y \ orbit f x\ have "funpow_dist1 f x y \ funpow_dist1 f x y mod m" + by auto (metis \(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x\ funpow_dist1_prop funpow_dist_least funpow_dist_step leI) + with \m > 0\ show ?thesis + by (auto intro: order_trans) +qed + +end \ No newline at end of file diff --git a/src/HOL/Combinatorics/Permutations.thy b/src/HOL/Combinatorics/Permutations.thy --- a/src/HOL/Combinatorics/Permutations.thy +++ b/src/HOL/Combinatorics/Permutations.thy @@ -1,1628 +1,1670 @@ (* Author: Amine Chaieb, University of Cambridge *) section \Permutations, both general and specifically on finite sets.\ theory Permutations imports "HOL-Library.Multiset" "HOL-Library.Disjoint_Sets" begin subsection \Auxiliary\ abbreviation (input) fixpoints :: \('a \ 'a) \ 'a set\ where \fixpoints f \ {x. f x = x}\ lemma inj_on_fixpoints: \inj_on f (fixpoints f)\ by (rule inj_onI) simp lemma bij_betw_fixpoints: \bij_betw f (fixpoints f) (fixpoints f)\ using inj_on_fixpoints by (auto simp add: bij_betw_def) subsection \Basic definition and consequences\ definition permutes :: \('a \ 'a) \ 'a set \ bool\ (infixr \permutes\ 41) where \p permutes S \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)\ lemma bij_imp_permutes: \p permutes S\ if \bij_betw p S S\ and stable: \\x. x \ S \ p x = x\ proof - note \bij_betw p S S\ moreover have \bij_betw p (- S) (- S)\ by (auto simp add: stable intro!: bij_betw_imageI inj_onI) ultimately have \bij_betw p (S \ - S) (S \ - S)\ by (rule bij_betw_combine) simp then have \\!x. p x = y\ for y by (simp add: bij_iff) with stable show ?thesis by (simp add: permutes_def) qed context fixes p :: \'a \ 'a\ and S :: \'a set\ assumes perm: \p permutes S\ begin lemma permutes_inj: \inj p\ using perm by (auto simp: permutes_def inj_on_def) lemma permutes_image: \p ` S = S\ proof (rule set_eqI) fix x show \x \ p ` S \ x \ S\ proof assume \x \ p ` S\ then obtain y where \y \ S\ \p y = x\ by blast with perm show \x \ S\ by (cases \y = x\) (auto simp add: permutes_def) next assume \x \ S\ with perm obtain y where \y \ S\ \p y = x\ by (metis permutes_def) then show \x \ p ` S\ by blast qed qed lemma permutes_not_in: \x \ S \ p x = x\ using perm by (auto simp: permutes_def) lemma permutes_image_complement: \p ` (- S) = - S\ by (auto simp add: permutes_not_in) lemma permutes_in_image: \p x \ S \ x \ S\ using permutes_image permutes_inj by (auto dest: inj_image_mem_iff) lemma permutes_surj: \surj p\ proof - have \p ` (S \ - S) = p ` S \ p ` (- S)\ by (rule image_Un) then show ?thesis by (simp add: permutes_image permutes_image_complement) qed lemma permutes_inv_o: shows "p \ inv p = id" and "inv p \ p = id" using permutes_inj permutes_surj unfolding inj_iff [symmetric] surj_iff [symmetric] by auto lemma permutes_inverses: shows "p (inv p x) = x" and "inv p (p x) = x" using permutes_inv_o [unfolded fun_eq_iff o_def] by auto lemma permutes_inv_eq: \inv p y = x \ p x = y\ by (auto simp add: permutes_inverses) lemma permutes_inj_on: \inj_on p A\ by (rule inj_on_subset [of _ UNIV]) (auto intro: permutes_inj) lemma permutes_bij: \bij p\ unfolding bij_def by (metis permutes_inj permutes_surj) lemma permutes_imp_bij: \bij_betw p S S\ by (simp add: bij_betw_def permutes_image permutes_inj_on) lemma permutes_subset: \p permutes T\ if \S \ T\ proof (rule bij_imp_permutes) define R where \R = T - S\ with that have \T = R \ S\ \R \ S = {}\ by auto then have \p x = x\ if \x \ R\ for x using that by (auto intro: permutes_not_in) then have \p ` R = R\ by simp with \T = R \ S\ show \bij_betw p T T\ by (simp add: bij_betw_def permutes_inj_on image_Un permutes_image) fix x assume \x \ T\ with \T = R \ S\ show \p x = x\ by (simp add: permutes_not_in) qed lemma permutes_imp_permutes_insert: \p permutes insert x S\ by (rule permutes_subset) auto end lemma permutes_id [simp]: \id permutes S\ by (auto intro: bij_imp_permutes) lemma permutes_empty [simp]: \p permutes {} \ p = id\ proof assume \p permutes {}\ then show \p = id\ by (auto simp add: fun_eq_iff permutes_not_in) next assume \p = id\ then show \p permutes {}\ by simp qed lemma permutes_sing [simp]: \p permutes {a} \ p = id\ proof assume perm: \p permutes {a}\ show \p = id\ proof fix x from perm have \p ` {a} = {a}\ by (rule permutes_image) with perm show \p x = id x\ by (cases \x = a\) (auto simp add: permutes_not_in) qed next assume \p = id\ then show \p permutes {a}\ by simp qed lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" by (simp add: permutes_def) lemma permutes_swap_id: "a \ S \ b \ S \ Fun.swap a b id permutes S" by (rule bij_imp_permutes) (auto simp add: swap_id_eq) lemma permutes_superset: \p permutes T\ if \p permutes S\ \\x. x \ S - T \ p x = x\ proof - define R U where \R = T \ S\ and \U = S - T\ then have \T = R \ (T - S)\ \S = R \ U\ \R \ U = {}\ by auto from that \U = S - T\ have \p ` U = U\ by simp from \p permutes S\ have \bij_betw p (R \ U) (R \ U)\ by (simp add: permutes_imp_bij \S = R \ U\) moreover have \bij_betw p U U\ using that \U = S - T\ by (simp add: bij_betw_def permutes_inj_on) ultimately have \bij_betw p R R\ using \R \ U = {}\ \R \ U = {}\ by (rule bij_betw_partition) then have \p permutes R\ proof (rule bij_imp_permutes) fix x assume \x \ R\ with \R = T \ S\ \p permutes S\ show \p x = x\ by (cases \x \ S\) (auto simp add: permutes_not_in that(2)) qed then have \p permutes R \ (T - S)\ by (rule permutes_subset) simp with \T = R \ (T - S)\ show ?thesis by simp qed lemma permutes_bij_inv_into: \<^marker>\contributor \Lukas Bulwahn\\ fixes A :: "'a set" and B :: "'b set" assumes "p permutes A" and "bij_betw f A B" shows "(\x. if x \ B then f (p (inv_into A f x)) else x) permutes B" proof (rule bij_imp_permutes) from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A" by (auto simp add: permutes_imp_bij bij_betw_inv_into) then have "bij_betw (f \ p \ inv_into A f) B B" by (simp add: bij_betw_trans) then show "bij_betw (\x. if x \ B then f (p (inv_into A f x)) else x) B B" by (subst bij_betw_cong[where g="f \ p \ inv_into A f"]) auto next fix x assume "x \ B" then show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto qed lemma permutes_image_mset: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" shows "image_mset p (mset_set A) = mset_set A" using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image) lemma permutes_implies_image_mset_eq: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" "\x. x \ A \ f x = f' (p x)" shows "image_mset f' (mset_set A) = image_mset f (mset_set A)" proof - have "f x = f' (p x)" if "x \# mset_set A" for x using assms(2)[of x] that by (cases "finite A") auto with assms have "image_mset f (mset_set A) = image_mset (f' \ p) (mset_set A)" by (auto intro!: image_mset_cong) also have "\ = image_mset f' (image_mset p (mset_set A))" by (simp add: image_mset.compositionality) also have "\ = image_mset f' (mset_set A)" proof - from assms permutes_image_mset have "image_mset p (mset_set A) = mset_set A" by blast then show ?thesis by simp qed finally show ?thesis .. qed subsection \Group properties\ lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S" unfolding permutes_def o_def by metis lemma permutes_inv: assumes "p permutes S" shows "inv p permutes S" using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis lemma permutes_inv_inv: assumes "p permutes S" shows "inv (inv p) = p" unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]] by blast lemma permutes_invI: assumes perm: "p permutes S" and inv: "\x. x \ S \ p' (p x) = x" and outside: "\x. x \ S \ p' x = x" shows "inv p = p'" proof show "inv p x = p' x" for x proof (cases "x \ S") case True from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses) also from permutes_inv[OF perm] True have "\ = inv p x" by (subst inv) (simp_all add: permutes_in_image) finally show ?thesis .. next case False with permutes_inv[OF perm] show ?thesis by (simp_all add: outside permutes_not_in) qed qed lemma permutes_vimage: "f permutes A \ f -` A = A" by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv]) subsection \Mapping permutations with bijections\ lemma bij_betw_permutations: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A} {\. \ permutes B}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 show ?case using permutes_bij_inv_into[OF _ assms] by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) { fix \ assume "\ permutes B" from permutes_bij_inv_into[OF this bij_inv] and assms have "(\x. if x \ A then inv_into A f (\ (f x)) else x) permutes A" by (simp add: inv_into_inv_into_eq cong: if_cong) } from this show ?case by (auto simp: permutes_inv) next case 1 thus ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed lemma bij_betw_derangements: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A \ (\x\A. \ x \ x)} {\. \ permutes B \ (\x\B. \ x \ x)}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 have "?f \ x \ x" if "\ permutes A" "\x. x \ A \ \ x \ x" "x \ B" for \ x using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on inv_into_f_f inv_into_into permutes_imp_bij) with permutes_bij_inv_into[OF _ assms] show ?case by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) have "?g \ permutes A" if "\ permutes B" for \ using permutes_bij_inv_into[OF that bij_inv] and assms by (simp add: inv_into_inv_into_eq cong: if_cong) moreover have "?g \ x \ x" if "\ permutes B" "\x. x \ B \ \ x \ x" "x \ A" for \ x using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij) ultimately show ?case by auto next case 1 thus ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed subsection \The number of permutations on a finite set\ lemma permutes_insert_lemma: assumes "p permutes (insert a S)" shows "Fun.swap a (p a) id \ p permutes S" apply (rule permutes_superset[where S = "insert a S"]) apply (rule permutes_compose[OF assms]) apply (rule permutes_swap_id, simp) using permutes_in_image[OF assms, of a] apply simp apply (auto simp add: Ball_def Fun.swap_def) done lemma permutes_insert: "{p. p permutes (insert a S)} = (\(b, p). Fun.swap a b id \ p) ` {(b, p). b \ insert a S \ p \ {p. p permutes S}}" proof - have "p permutes insert a S \ (\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S)" for p proof - have "\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S" if p: "p permutes insert a S" proof - let ?b = "p a" let ?q = "Fun.swap a (p a) id \ p" have *: "p = Fun.swap a ?b id \ ?q" by (simp add: fun_eq_iff o_assoc) have **: "?b \ insert a S" unfolding permutes_in_image[OF p] by simp from permutes_insert_lemma[OF p] * ** show ?thesis by blast qed moreover have "p permutes insert a S" if bq: "p = Fun.swap a b id \ q" "b \ insert a S" "q permutes S" for b q proof - from permutes_subset[OF bq(3), of "insert a S"] have q: "q permutes insert a S" by auto have a: "a \ insert a S" by simp from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis by simp qed ultimately show ?thesis by blast qed then show ?thesis by auto qed lemma card_permutations: assumes "card S = n" and "finite S" shows "card {p. p permutes S} = fact n" using assms(2,1) proof (induct arbitrary: n) case empty then show ?case by simp next case (insert x F) { fix n assume card_insert: "card (insert x F) = n" let ?xF = "{p. p permutes insert x F}" let ?pF = "{p. p permutes F}" let ?pF' = "{(b, p). b \ insert x F \ p \ ?pF}" let ?g = "(\(b, p). Fun.swap x b id \ p)" have xfgpF': "?xF = ?g ` ?pF'" by (rule permutes_insert[of x F]) from \x \ F\ \finite F\ card_insert have Fs: "card F = n - 1" by auto from \finite F\ insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" by auto then have "finite ?pF" by (auto intro: card_ge_0_finite) with \finite F\ card.insert_remove have pF'f: "finite ?pF'" apply (simp only: Collect_case_prod Collect_mem_eq) apply (rule finite_cartesian_product) apply simp_all done have ginj: "inj_on ?g ?pF'" proof - { fix b p c q assume bp: "(b, p) \ ?pF'" assume cq: "(c, q) \ ?pF'" assume eq: "?g (b, p) = ?g (c, q)" from bp cq have pF: "p permutes F" and qF: "q permutes F" by auto from pF \x \ F\ eq have "b = ?g (b, p) x" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) also from qF \x \ F\ eq have "\ = ?g (c, q) x" by (auto simp: fun_upd_def fun_eq_iff) also from qF \x \ F\ have "\ = c" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) finally have "b = c" . then have "Fun.swap x b id = Fun.swap x c id" by simp with eq have "Fun.swap x b id \ p = Fun.swap x b id \ q" by simp then have "Fun.swap x b id \ (Fun.swap x b id \ p) = Fun.swap x b id \ (Fun.swap x b id \ q)" by simp then have "p = q" by (simp add: o_assoc) with \b = c\ have "(b, p) = (c, q)" by simp } then show ?thesis unfolding inj_on_def by blast qed from \x \ F\ \finite F\ card_insert have "n \ 0" by auto then have "\m. n = Suc m" by presburger then obtain m where n: "n = Suc m" by blast from pFs card_insert have *: "card ?xF = fact n" unfolding xfgpF' card_image[OF ginj] using \finite F\ \finite ?pF\ by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n) from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" by (simp add: xfgpF' n) from * have "card ?xF = fact n" unfolding xFf by blast } with insert show ?case by simp qed lemma finite_permutations: assumes "finite S" shows "finite {p. p permutes S}" using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite) subsection \Hence a sort of induction principle composing by swaps\ lemma permutes_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id: \P id\ and swap: \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (Fun.swap a b id \ p)\ using \finite S\ \p permutes S\ swap proof (induction S arbitrary: p) case empty with id show ?case by (simp only: permutes_empty) next case (insert x S p) define q where \q = Fun.swap x (p x) id \ p\ then have swap_q: \Fun.swap x (p x) id \ q = p\ by (simp add: o_assoc) from \p permutes insert x S\ have \q permutes S\ by (simp add: q_def permutes_insert_lemma) then have \q permutes insert x S\ by (simp add: permutes_imp_permutes_insert) from \q permutes S\ have \P q\ by (auto intro: insert.IH insert.prems(2) permutes_imp_permutes_insert) have \x \ insert x S\ by simp moreover from \p permutes insert x S\ have \p x \ insert x S\ using permutes_in_image [of p \insert x S\ x] by simp ultimately have \P (Fun.swap x (p x) id \ q)\ using \q permutes insert x S\ \P q\ by (rule insert.prems(2)) then show ?case by (simp add: swap_q) qed lemma permutes_rev_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id': \P id\ and swap': \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (Fun.swap a b p)\ using \p permutes S\ \finite S\ proof (induction rule: permutes_induct) case id from id' show ?case . next case (swap a b p) have \P (Fun.swap (inv p a) (inv p b) p)\ by (rule swap') (auto simp add: swap permutes_in_image permutes_inv) also have \Fun.swap (inv p a) (inv p b) p = Fun.swap a b id \ p\ by (rule bij_swap_comp [symmetric]) (rule permutes_bij, rule swap) finally show ?case . qed subsection \Permutations of index set for iterated operations\ lemma (in comm_monoid_set) permute: assumes "p permutes S" shows "F g S = F (g \ p) S" proof - from \p permutes S\ have "inj p" by (rule permutes_inj) then have "inj_on p S" by (auto intro: subset_inj_on) then have "F g (p ` S) = F (g \ p) S" by (rule reindex) moreover from \p permutes S\ have "p ` S = S" by (rule permutes_image) ultimately show ?thesis by simp qed subsection \Permutations as transposition sequences\ inductive swapidseq :: "nat \ ('a \ 'a) \ bool" where id[simp]: "swapidseq 0 id" | comp_Suc: "swapidseq n p \ a \ b \ swapidseq (Suc n) (Fun.swap a b id \ p)" declare id[unfolded id_def, simp] definition "permutation p \ (\n. swapidseq n p)" subsection \Some closure properties of the set of permutations, with lengths\ lemma permutation_id[simp]: "permutation id" unfolding permutation_def by (rule exI[where x=0]) simp declare permutation_id[unfolded id_def, simp] lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)" apply clarsimp using comp_Suc[of 0 id a b] apply simp done lemma permutation_swap_id: "permutation (Fun.swap a b id)" proof (cases "a = b") case True then show ?thesis by simp next case False then show ?thesis unfolding permutation_def using swapidseq_swap[of a b] by blast qed lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q \ swapidseq (n + m) (p \ q)" proof (induct n p arbitrary: m q rule: swapidseq.induct) case (id m q) then show ?case by simp next case (comp_Suc n p a b m q) have eq: "Suc n + m = Suc (n + m)" by arith show ?case apply (simp only: eq comp_assoc) apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) apply blast+ done qed lemma permutation_compose: "permutation p \ permutation q \ permutation (p \ q)" unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis lemma swapidseq_endswap: "swapidseq n p \ a \ b \ swapidseq (Suc n) (p \ Fun.swap a b id)" by (induct n p rule: swapidseq.induct) (use swapidseq_swap[of a b] in \auto simp add: comp_assoc intro: swapidseq.comp_Suc\) lemma swapidseq_inverse_exists: "swapidseq n p \ \q. swapidseq n q \ p \ q = id \ q \ p = id" proof (induct n p rule: swapidseq.induct) case id then show ?case by (rule exI[where x=id]) simp next case (comp_Suc n p a b) from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast let ?q = "q \ Fun.swap a b id" note H = comp_Suc.hyps from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (Fun.swap a b id)" by simp from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q" by simp have "Fun.swap a b id \ p \ ?q = Fun.swap a b id \ (p \ q) \ Fun.swap a b id" by (simp add: o_assoc) also have "\ = id" by (simp add: q(2)) finally have ***: "Fun.swap a b id \ p \ ?q = id" . have "?q \ (Fun.swap a b id \ p) = q \ (Fun.swap a b id \ Fun.swap a b id) \ p" by (simp only: o_assoc) then have "?q \ (Fun.swap a b id \ p) = id" by (simp add: q(3)) with ** *** show ?case by blast qed lemma swapidseq_inverse: assumes "swapidseq n p" shows "swapidseq n (inv p)" using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto lemma permutation_inverse: "permutation p \ permutation (inv p)" using permutation_def swapidseq_inverse by blast subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ lemma swap_id_common:" a \ c \ b \ c \ Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) lemma swap_id_common': "a \ b \ a \ c \ Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" by (simp add: fun_eq_iff Fun.swap_def) subsection \The identity map only has even transposition sequences\ lemma symmetry_lemma: assumes "\a b c d. P a b c d \ P a b d c" and "\a b c d. a \ b \ c \ d \ a = c \ b = d \ a = c \ b \ d \ a \ c \ b = d \ a \ c \ a \ d \ b \ c \ b \ d \ P a b c d" shows "\a b c d. a \ b \ c \ d \ P a b c d" using assms by metis lemma swap_general: "a \ b \ c \ d \ Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id)" proof - assume neq: "a \ b" "c \ d" have "a \ b \ c \ d \ (Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id))" apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d]) apply (simp_all only: swap_commute) apply (case_tac "a = c \ b = d") apply (clarsimp simp only: swap_commute swap_id_idempotent) apply (case_tac "a = c \ b \ d") apply (rule disjI2) apply (rule_tac x="b" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (case_tac "a \ c \ b = d") apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="c" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) done with neq show ?thesis by metis qed lemma swapidseq_id_iff[simp]: "swapidseq 0 p \ p = id" using swapidseq.cases[of 0 p "p = id"] by auto lemma swapidseq_cases: "swapidseq n p \ n = 0 \ p = id \ (\a b q m. n = Suc m \ p = Fun.swap a b id \ q \ swapidseq m q \ a \ b)" apply (rule iffI) apply (erule swapidseq.cases[of n p]) apply simp apply (rule disjI2) apply (rule_tac x= "a" in exI) apply (rule_tac x= "b" in exI) apply (rule_tac x= "pa" in exI) apply (rule_tac x= "na" in exI) apply simp apply auto apply (rule comp_Suc, simp_all) done lemma fixing_swapidseq_decrease: assumes "swapidseq n p" and "a \ b" and "(Fun.swap a b id \ p) a = a" shows "n \ 0 \ swapidseq (n - 1) (Fun.swap a b id \ p)" using assms proof (induct n arbitrary: p a b) case 0 then show ?case by (auto simp add: Fun.swap_def fun_upd_def) next case (Suc n p a b) from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id \ q" "swapidseq m q" "c \ d" "n = m" by auto consider "Fun.swap a b id \ Fun.swap c d id = id" | x y z where "x \ a" "y \ a" "z \ a" "x \ y" "Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id" using swap_general[OF Suc.prems(2) cdqm(4)] by metis then show ?case proof cases case 1 then show ?thesis by (simp only: cdqm o_assoc) (simp add: cdqm) next case prems: 2 then have az: "a \ z" by simp from prems have *: "(Fun.swap x y id \ h) a = a \ h a = a" for h by (simp add: Fun.swap_def) from cdqm(2) have "Fun.swap a b id \ p = Fun.swap a b id \ (Fun.swap c d id \ q)" by simp then have "Fun.swap a b id \ p = Fun.swap x y id \ (Fun.swap a z id \ q)" by (simp add: o_assoc prems) then have "(Fun.swap a b id \ p) a = (Fun.swap x y id \ (Fun.swap a z id \ q)) a" by simp then have "(Fun.swap x y id \ (Fun.swap a z id \ q)) a = a" unfolding Suc by metis then have "(Fun.swap a z id \ q) a = a" by (simp only: *) from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this] have **: "swapidseq (n - 1) (Fun.swap a z id \ q)" "n \ 0" by blast+ from \n \ 0\ have ***: "Suc n - 1 = Suc (n - 1)" by auto show ?thesis apply (simp only: cdqm(2) prems o_assoc ***) apply (simp only: Suc_not_Zero simp_thms comp_assoc) apply (rule comp_Suc) using ** prems apply blast+ done qed qed lemma swapidseq_identity_even: assumes "swapidseq n (id :: 'a \ 'a)" shows "even n" using \swapidseq n id\ proof (induct n rule: nat_less_induct) case H: (1 n) consider "n = 0" | a b :: 'a and q m where "n = Suc m" "id = Fun.swap a b id \ q" "swapidseq m q" "a \ b" using H(2)[unfolded swapidseq_cases[of n id]] by auto then show ?case proof cases case 1 then show ?thesis by presburger next case h: 2 from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]] have m: "m \ 0" "swapidseq (m - 1) (id :: 'a \ 'a)" by auto from h m have mn: "m - 1 < n" by arith from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis by presburger qed qed subsection \Therefore we have a welldefined notion of parity\ definition "evenperm p = even (SOME n. swapidseq n p)" lemma swapidseq_even_even: assumes m: "swapidseq m p" and n: "swapidseq n p" shows "even m \ even n" proof - from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis by arith qed lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b" shows "evenperm p = b" unfolding n[symmetric] evenperm_def apply (rule swapidseq_even_even[where p = p]) apply (rule someI[where x = n]) using p apply blast+ done subsection \And it has the expected composition properties\ lemma evenperm_id[simp]: "evenperm id = True" by (rule evenperm_unique[where n = 0]) simp_all lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)" by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap) lemma evenperm_comp: assumes "permutation p" "permutation q" shows "evenperm (p \ q) \ evenperm p = evenperm q" proof - from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q" unfolding permutation_def by blast have "even (n + m) \ (even n \ even m)" by arith from evenperm_unique[OF n refl] evenperm_unique[OF m refl] and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis by blast qed lemma evenperm_inv: assumes "permutation p" shows "evenperm (inv p) = evenperm p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast show ?thesis by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]) qed subsection \A more abstract characterization of permutations\ lemma permutation_bijective: assumes "permutation p" shows "bij p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast then show ?thesis unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done qed lemma permutation_finite_support: assumes "permutation p" shows "finite {x. p x \ x}" proof - from assms obtain n where "swapidseq n p" unfolding permutation_def by blast then show ?thesis proof (induct n p rule: swapidseq.induct) case id then show ?case by simp next case (comp_Suc n p a b) let ?S = "insert a (insert b {x. p x \ x})" from comp_Suc.hyps(2) have *: "finite ?S" by simp from \a \ b\ have **: "{x. (Fun.swap a b id \ p) x \ x} \ ?S" by (auto simp: Fun.swap_def) show ?case by (rule finite_subset[OF ** *]) qed qed lemma permutation_lemma: assumes "finite S" and "bij p" and "\x. x \ S \ p x = x" shows "permutation p" using assms proof (induct S arbitrary: p rule: finite_induct) case empty then show ?case by simp next case (insert a F p) let ?r = "Fun.swap a (p a) id \ p" let ?q = "Fun.swap a (p a) id \ ?r" have *: "?r a = a" by (simp add: Fun.swap_def) from insert * have **: "\x. x \ F \ ?r x = x" by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3)) have "bij ?r" by (rule bij_swap_compose_bij[OF insert(4)]) have "permutation ?r" by (rule insert(3)[OF \bij ?r\ **]) then have "permutation ?q" by (simp add: permutation_compose permutation_swap_id) then show ?case by (simp add: o_assoc) qed lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}" (is "?lhs \ ?b \ ?f") proof assume ?lhs with permutation_bijective permutation_finite_support show "?b \ ?f" by auto next assume "?b \ ?f" then have "?f" "?b" by blast+ from permutation_lemma[OF this] show ?lhs by blast qed lemma permutation_inverse_works: assumes "permutation p" shows "inv p \ p = id" and "p \ inv p = id" using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff) lemma permutation_inverse_compose: assumes p: "permutation p" and q: "permutation q" shows "inv (p \ q) = inv q \ inv p" proof - note ps = permutation_inverse_works[OF p] note qs = permutation_inverse_works[OF q] have "p \ q \ (inv q \ inv p) = p \ (q \ inv q) \ inv p" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have *: "p \ q \ (inv q \ inv p) = id" . have "inv q \ inv p \ (p \ q) = inv q \ (inv p \ p) \ q" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have **: "inv q \ inv p \ (p \ q) = id" . show ?thesis by (rule inv_unique_comp[OF * **]) qed subsection \Relation to \permutes\\ lemma permutes_imp_permutation: \permutation p\ if \finite S\ \p permutes S\ proof - from \p permutes S\ have \{x. p x \ x} \ S\ by (auto dest: permutes_not_in) then have \finite {x. p x \ x}\ using \finite S\ by (rule finite_subset) moreover from \p permutes S\ have \bij p\ by (auto dest: permutes_bij) ultimately show ?thesis by (simp add: permutation) qed lemma permutation_permutesE: assumes \permutation p\ obtains S where \finite S\ \p permutes S\ proof - from assms have fin: \finite {x. p x \ x}\ by (simp add: permutation) from assms have \bij p\ by (simp add: permutation) also have \UNIV = {x. p x \ x} \ {x. p x = x}\ by auto finally have \bij_betw p {x. p x \ x} {x. p x \ x}\ by (rule bij_betw_partition) (auto simp add: bij_betw_fixpoints) then have \p permutes {x. p x \ x}\ by (auto intro: bij_imp_permutes) with fin show thesis .. qed lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)" by (auto elim: permutation_permutesE intro: permutes_imp_permutation) subsection \Sign of a permutation as a real number\ definition sign :: \('a \ 'a) \ int\ \ \TODO: prefer less generic name\ where \sign p = (if evenperm p then (1::int) else -1)\ lemma sign_nz: "sign p \ 0" by (simp add: sign_def) lemma sign_id: "sign id = 1" by (simp add: sign_def) lemma sign_inverse: "permutation p \ sign (inv p) = sign p" by (simp add: sign_def evenperm_inv) lemma sign_compose: "permutation p \ permutation q \ sign (p \ q) = sign p * sign q" by (simp add: sign_def evenperm_comp) lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)" by (simp add: sign_def evenperm_swap) lemma sign_idempotent: "sign p * sign p = 1" by (simp add: sign_def) subsection \Permuting a list\ text \This function permutes a list by applying a permutation to the indices.\ definition permute_list :: "(nat \ nat) \ 'a list \ 'a list" where "permute_list f xs = map (\i. xs ! (f i)) [0.. g) xs = permute_list g (permute_list f xs)" using assms[THEN permutes_in_image] by (auto simp add: permute_list_def) lemma permute_list_ident [simp]: "permute_list (\x. x) xs = xs" by (simp add: permute_list_def map_nth) lemma permute_list_id [simp]: "permute_list id xs = xs" by (simp add: id_def) lemma mset_permute_list [simp]: fixes xs :: "'a list" assumes "f permutes {.. x < length xs" for x using permutes_in_image[OF assms] by auto have "count (mset (permute_list f xs)) y = card ((\i. xs ! f i) -` {y} \ {..i. xs ! f i) -` {y} \ {.. y = xs ! i}" by auto also from assms have "card \ = card {i. i < length xs \ y = xs ! i}" by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj) also have "\ = count (mset xs) y" by (simp add: count_mset length_filter_conv_card) finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp qed lemma set_permute_list [simp]: assumes "f permutes {.. i < length ys" for i by simp have "permute_list f (zip xs ys) = map (\i. zip xs ys ! f i) [0.. = map (\(x, y). (xs ! f x, ys ! f y)) (zip [0.. = zip (permute_list f xs) (permute_list f ys)" by (simp_all add: permute_list_def zip_map_map) finally show ?thesis . qed lemma map_of_permute: assumes "\ permutes fst ` set xs" shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)" (is "_ = map_of (map ?f _)") proof from assms have "inj \" "surj \" by (simp_all add: permutes_inj permutes_surj) then show "(map_of xs \ \) x = map_of (map ?f xs) x" for x by (induct xs) (auto simp: inv_f_f surj_f_inv_f) qed subsection \More lemmas about permutations\ +lemma permutes_in_funpow_image: \<^marker>\contributor \Lars Noschinski\\ + assumes "f permutes S" "x \ S" + shows "(f ^^ n) x \ S" + using assms by (induction n) (auto simp: permutes_in_image) + +lemma permutation_self: \<^marker>\contributor \Lars Noschinski\\ + assumes \permutation p\ + obtains n where \n > 0\ \(p ^^ n) x = x\ +proof (cases \p x = x\) + case True + with that [of 1] show thesis by simp +next + case False + from \permutation p\ have \inj p\ + by (intro permutation_bijective bij_is_inj) + moreover from \p x \ x\ have \(p ^^ Suc n) x \ (p ^^ n) x\ for n + proof (induction n arbitrary: x) + case 0 then show ?case by simp + next + case (Suc n) + have "p (p x) \ p x" + proof (rule notI) + assume "p (p x) = p x" + then show False using \p x \ x\ \inj p\ by (simp add: inj_eq) + qed + have "(p ^^ Suc (Suc n)) x = (p ^^ Suc n) (p x)" + by (simp add: funpow_swap1) + also have "\ \ (p ^^ n) (p x)" + by (rule Suc) fact + also have "(p ^^ n) (p x) = (p ^^ Suc n) x" + by (simp add: funpow_swap1) + finally show ?case by simp + qed + then have "{y. \n. y = (p ^^ n) x} \ {x. p x \ x}" + by auto + then have "finite {y. \n. y = (p ^^ n) x}" + using permutation_finite_support[OF assms] by (rule finite_subset) + ultimately obtain n where \n > 0\ \(p ^^ n) x = x\ + by (rule funpow_inj_finite) + with that [of n] show thesis by blast +qed + text \The following few lemmas were contributed by Lukas Bulwahn.\ lemma count_image_mset_eq_card_vimage: assumes "finite A" shows "count (image_mset f (mset_set A)) b = card {a \ A. f a = b}" using assms proof (induct A) case empty show ?case by simp next case (insert x F) show ?case proof (cases "f x = b") case True with insert.hyps have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})" by auto also from insert.hyps(1,2) have "\ = card (insert x {a \ F. f a = f x})" by simp also from \f x = b\ have "card (insert x {a \ F. f a = f x}) = card {a \ insert x F. f a = b}" by (auto intro: arg_cong[where f="card"]) finally show ?thesis using insert by auto next case False then have "{a \ F. f a = b} = {a \ insert x F. f a = b}" by auto with insert False show ?thesis by simp qed qed \ \Prove \image_mset_eq_implies_permutes\ ...\ lemma image_mset_eq_implies_permutes: fixes f :: "'a \ 'b" assumes "finite A" and mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)" obtains p where "p permutes A" and "\x\A. f x = f' (p x)" proof - from \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for f b by auto have "f ` A = f' ` A" proof - from \finite A\ have "f ` A = f ` (set_mset (mset_set A))" by simp also have "\ = f' ` set_mset (mset_set A)" by (metis mset_eq multiset.set_map) also from \finite A\ have "\ = f' ` A" by simp finally show ?thesis . qed have "\b\(f ` A). \p. bij_betw p {a \ A. f a = b} {a \ A. f' a = b}" proof fix b from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b" by simp with \finite A\ have "card {a \ A. f a = b} = card {a \ A. f' a = b}" by (simp add: count_image_mset_eq_card_vimage) then show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" by (intro finite_same_card_bij) simp_all qed then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) then obtain p where p: "\b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" .. define p' where "p' = (\a. if a \ A then p (f a) a else a)" have "p' permutes A" proof (rule bij_imp_permutes) have "disjoint_family_on (\i. {a \ A. f' a = i}) (f ` A)" by (auto simp: disjoint_family_on_def) moreover have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if "b \ f ` A" for b using p that by (subst bij_betw_cong[where g="p b"]) auto ultimately have "bij_betw (\a. p (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A. f' a = b})" by (rule bij_betw_UNION_disjoint) moreover have "(\b\f ` A. {a \ A. f a = b}) = A" by auto moreover from \f ` A = f' ` A\ have "(\b\f ` A. {a \ A. f' a = b}) = A" by auto ultimately show "bij_betw p' A A" unfolding p'_def by (subst bij_betw_cong[where g="(\a. p (f a) a)"]) auto next show "\x. x \ A \ p' x = x" by (simp add: p'_def) qed moreover from p have "\x\A. f x = f' (p' x)" unfolding p'_def using bij_betwE by fastforce ultimately show ?thesis .. qed \ \... and derive the existing property:\ lemma mset_eq_permutation: fixes xs ys :: "'a list" assumes mset_eq: "mset xs = mset ys" obtains p where "p permutes {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..i \ S. p i \ i" shows "p = id" proof - have "p n = n" for n using assms proof (induct n arbitrary: S rule: less_induct) case (less n) show ?case proof (cases "n \ S") case False with less(2) show ?thesis unfolding permutes_def by metis next case True with less(3) have "p n < n \ p n = n" by auto then show ?thesis proof assume "p n < n" with less have "p (p n) = p n" by metis with permutes_inj[OF less(2)] have "p n = n" unfolding inj_def by blast with \p n < n\ have False by simp then show ?thesis .. qed qed qed then show ?thesis by (auto simp: fun_eq_iff) qed lemma permutes_natset_ge: fixes S :: "'a::wellorder set" assumes p: "p permutes S" and le: "\i \ S. p i \ i" shows "p = id" proof - have "i \ inv p i" if "i \ S" for i proof - from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S" by simp with le have "p (inv p i) \ inv p i" by blast with permutes_inverses[OF p] show ?thesis by simp qed then have "\i\S. inv p i \ i" by blast from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id" by simp then show ?thesis apply (subst permutes_inv_inv[OF p, symmetric]) apply (rule inv_unique_comp) apply simp_all done qed lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto using permutes_inv_inv permutes_inv apply auto apply (rule_tac x="inv x" in exI) apply auto done lemma image_compose_permutations_left: assumes "q permutes S" shows "{q \ p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "inv q \ x" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o) done lemma image_compose_permutations_right: assumes "q permutes S" shows "{p \ q | p. p permutes S} = {p . p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "x \ inv q" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc) done lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} \ 1 \ p i \ p i \ n" by (simp add: permutes_def) metis lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p . p permutes S}" have *: "inj_on inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r" then have "inv (inv q) = inv (inv r)" by simp with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed have **: "inv ` ?S = ?S" using image_inverse_permutations by blast have ***: "?rhs = sum (f \ inv) ?S" by (simp add: o_def) from sum.reindex[OF *, of f] show ?thesis by (simp only: ** ***) qed lemma setum_permutations_compose_left: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(q \ p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ ((\) q)) ?S" by (simp add: o_def) have **: "inj_on ((\) q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "q \ p = q \ r" then have "inv q \ q \ p = inv q \ q \ r" by (simp add: comp_assoc) with permutes_inj[OF q, unfolded inj_iff] show "p = r" by simp qed have "((\) q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed lemma sum_permutations_compose_right: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(p \ q)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ (\p. p \ q)) ?S" by (simp add: o_def) have **: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "p \ q = r \ q" then have "p \ (q \ inv q) = r \ (q \ inv q)" by (simp add: o_assoc) with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed from image_compose_permutations_right[OF q] have "(\p. p \ q) ` ?S = ?S" by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed subsection \Sum over a set of permutations (could generalize to iteration)\ lemma sum_over_permutations_insert: assumes fS: "finite S" and aS: "a \ S" shows "sum f {p. p permutes (insert a S)} = sum (\b. sum (\q. f (Fun.swap a b id \ q)) {p. p permutes S}) (insert a S)" proof - have *: "\f a b. (\(b, p). f (Fun.swap a b id \ p)) = f \ (\(b,p). Fun.swap a b id \ p)" by (simp add: fun_eq_iff) have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q" by blast show ?thesis unfolding * ** sum.cartesian_product permutes_insert proof (rule sum.reindex) let ?f = "(\(b, y). Fun.swap a b id \ y)" let ?P = "{p. p permutes S}" { fix b c p q assume b: "b \ insert a S" assume c: "c \ insert a S" assume p: "p permutes S" assume q: "q permutes S" assume eq: "Fun.swap a b id \ p = Fun.swap a c id \ q" from p q aS have pa: "p a = a" and qa: "q a = a" unfolding permutes_def by metis+ from eq have "(Fun.swap a b id \ p) a = (Fun.swap a c id \ q) a" by simp then have bc: "b = c" by (simp add: permutes_def pa qa o_def fun_upd_def Fun.swap_def id_def cong del: if_weak_cong split: if_split_asm) from eq[unfolded bc] have "(\p. Fun.swap a c id \ p) (Fun.swap a c id \ p) = (\p. Fun.swap a c id \ p) (Fun.swap a c id \ q)" by simp then have "p = q" unfolding o_assoc swap_id_idempotent by simp with bc have "b = c \ p = q" by blast } then show "inj_on ?f (insert a S \ ?P)" unfolding inj_on_def by clarify metis qed qed subsection \Constructing permutations from association lists\ definition list_permutes :: "('a \ 'a) list \ 'a set \ bool" where "list_permutes xs A \ set (map fst xs) \ A \ set (map snd xs) = set (map fst xs) \ distinct (map fst xs) \ distinct (map snd xs)" lemma list_permutesI [simp]: assumes "set (map fst xs) \ A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" shows "list_permutes xs A" proof - from assms(2,3) have "distinct (map snd xs)" by (intro card_distinct) (simp_all add: distinct_card del: set_map) with assms show ?thesis by (simp add: list_permutes_def) qed definition permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)" lemma permutation_of_list_Cons: "permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')" by (simp add: permutation_of_list_def) fun inverse_permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "inverse_permutation_of_list [] x = x" | "inverse_permutation_of_list ((y, x') # xs) x = (if x = x' then y else inverse_permutation_of_list xs x)" declare inverse_permutation_of_list.simps [simp del] lemma inj_on_map_of: assumes "distinct (map snd xs)" shows "inj_on (map_of xs) (set (map fst xs))" proof (rule inj_onI) fix x y assume xy: "x \ set (map fst xs)" "y \ set (map fst xs)" assume eq: "map_of xs x = map_of xs y" from xy obtain x' y' where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'" by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff) moreover from x'y' have *: "(x, x') \ set xs" "(y, y') \ set xs" by (force dest: map_of_SomeD)+ moreover from * eq x'y' have "x' = y'" by simp ultimately show "x = y" using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"]) qed lemma inj_on_the: "None \ A \ inj_on the A" by (auto simp: inj_on_def option.the_def split: option.splits) lemma inj_on_map_of': assumes "distinct (map snd xs)" shows "inj_on (the \ map_of xs) (set (map fst xs))" by (intro comp_inj_on inj_on_map_of assms inj_on_the) (force simp: eq_commute[of None] map_of_eq_None_iff) lemma image_map_of: assumes "distinct (map fst xs)" shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)" using assms by (auto simp: rev_image_eqI) lemma the_Some_image [simp]: "the ` Some ` A = A" by (subst image_image) simp lemma image_map_of': assumes "distinct (map fst xs)" shows "(the \ map_of xs) ` set (map fst xs) = set (map snd xs)" by (simp only: image_comp [symmetric] image_map_of assms the_Some_image) lemma permutation_of_list_permutes [simp]: assumes "list_permutes xs A" shows "permutation_of_list xs permutes A" (is "?f permutes _") proof (rule permutes_subset[OF bij_imp_permutes]) from assms show "set (map fst xs) \ A" by (simp add: list_permutes_def) from assms have "inj_on (the \ map_of xs) (set (map fst xs))" (is ?P) by (intro inj_on_map_of') (simp_all add: list_permutes_def) also have "?P \ inj_on ?f (set (map fst xs))" by (intro inj_on_cong) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))" by (rule inj_on_imp_bij_betw) also from assms have "?f ` set (map fst xs) = (the \ map_of xs) ` set (map fst xs)" by (intro image_cong refl) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) also from assms have "\ = set (map fst xs)" by (subst image_map_of') (simp_all add: list_permutes_def) finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" . qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+ lemma eval_permutation_of_list [simp]: "permutation_of_list [] x = x" "x = x' \ permutation_of_list ((x',y)#xs) x = y" "x \ x' \ permutation_of_list ((x',y')#xs) x = permutation_of_list xs x" by (simp_all add: permutation_of_list_def) lemma eval_inverse_permutation_of_list [simp]: "inverse_permutation_of_list [] x = x" "x = x' \ inverse_permutation_of_list ((y,x')#xs) x = y" "x \ x' \ inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x" by (simp_all add: inverse_permutation_of_list.simps) lemma permutation_of_list_id: "x \ set (map fst xs) \ permutation_of_list xs x = x" by (induct xs) (auto simp: permutation_of_list_Cons) lemma permutation_of_list_unique': "distinct (map fst xs) \ (x, y) \ set xs \ permutation_of_list xs x = y" by (induct xs) (force simp: permutation_of_list_Cons)+ lemma permutation_of_list_unique: "list_permutes xs A \ (x, y) \ set xs \ permutation_of_list xs x = y" by (intro permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_id: "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x" by (induct xs) auto lemma inverse_permutation_of_list_unique': "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x" by (induct xs) (force simp: inverse_permutation_of_list.simps(2))+ lemma inverse_permutation_of_list_unique: "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x" by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_correct: fixes A :: "'a set" assumes "list_permutes xs A" shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)" proof (rule ext, rule sym, subst permutes_inv_eq) from assms show "permutation_of_list xs permutes A" by simp show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x proof (cases "x \ set (map snd xs)") case True then obtain y where "(y, x) \ set xs" by auto with assms show ?thesis by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique) next case False with assms show ?thesis by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id) qed qed end diff --git a/src/HOL/Euclidean_Division.thy b/src/HOL/Euclidean_Division.thy --- a/src/HOL/Euclidean_Division.thy +++ b/src/HOL/Euclidean_Division.thy @@ -1,2086 +1,2099 @@ (* Title: HOL/Euclidean_Division.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Division in euclidean (semi)rings\ theory Euclidean_Division imports Int Lattices_Big begin subsection \Euclidean (semi)rings with explicit division and remainder\ class euclidean_semiring = semidom_modulo + fixes euclidean_size :: "'a \ nat" assumes size_0 [simp]: "euclidean_size 0 = 0" assumes mod_size_less: "b \ 0 \ euclidean_size (a mod b) < euclidean_size b" assumes size_mult_mono: "b \ 0 \ euclidean_size a \ euclidean_size (a * b)" begin lemma euclidean_size_eq_0_iff [simp]: "euclidean_size b = 0 \ b = 0" proof assume "b = 0" then show "euclidean_size b = 0" by simp next assume "euclidean_size b = 0" show "b = 0" proof (rule ccontr) assume "b \ 0" with mod_size_less have "euclidean_size (b mod b) < euclidean_size b" . with \euclidean_size b = 0\ show False by simp qed qed lemma euclidean_size_greater_0_iff [simp]: "euclidean_size b > 0 \ b \ 0" using euclidean_size_eq_0_iff [symmetric, of b] by safe simp lemma size_mult_mono': "b \ 0 \ euclidean_size a \ euclidean_size (b * a)" by (subst mult.commute) (rule size_mult_mono) lemma dvd_euclidean_size_eq_imp_dvd: assumes "a \ 0" and "euclidean_size a = euclidean_size b" and "b dvd a" shows "a dvd b" proof (rule ccontr) assume "\ a dvd b" hence "b mod a \ 0" using mod_0_imp_dvd [of b a] by blast then have "b mod a \ 0" by (simp add: mod_eq_0_iff_dvd) from \b dvd a\ have "b dvd b mod a" by (simp add: dvd_mod_iff) then obtain c where "b mod a = b * c" unfolding dvd_def by blast with \b mod a \ 0\ have "c \ 0" by auto with \b mod a = b * c\ have "euclidean_size (b mod a) \ euclidean_size b" using size_mult_mono by force moreover from \\ a dvd b\ and \a \ 0\ have "euclidean_size (b mod a) < euclidean_size a" using mod_size_less by blast ultimately show False using \euclidean_size a = euclidean_size b\ by simp qed lemma euclidean_size_times_unit: assumes "is_unit a" shows "euclidean_size (a * b) = euclidean_size b" proof (rule antisym) from assms have [simp]: "a \ 0" by auto thus "euclidean_size (a * b) \ euclidean_size b" by (rule size_mult_mono') from assms have "is_unit (1 div a)" by simp hence "1 div a \ 0" by (intro notI) simp_all hence "euclidean_size (a * b) \ euclidean_size ((1 div a) * (a * b))" by (rule size_mult_mono') also from assms have "(1 div a) * (a * b) = b" by (simp add: algebra_simps unit_div_mult_swap) finally show "euclidean_size (a * b) \ euclidean_size b" . qed lemma euclidean_size_unit: "is_unit a \ euclidean_size a = euclidean_size 1" using euclidean_size_times_unit [of a 1] by simp lemma unit_iff_euclidean_size: "is_unit a \ euclidean_size a = euclidean_size 1 \ a \ 0" proof safe assume A: "a \ 0" and B: "euclidean_size a = euclidean_size 1" show "is_unit a" by (rule dvd_euclidean_size_eq_imp_dvd [OF A B]) simp_all qed (auto intro: euclidean_size_unit) lemma euclidean_size_times_nonunit: assumes "a \ 0" "b \ 0" "\ is_unit a" shows "euclidean_size b < euclidean_size (a * b)" proof (rule ccontr) assume "\euclidean_size b < euclidean_size (a * b)" with size_mult_mono'[OF assms(1), of b] have eq: "euclidean_size (a * b) = euclidean_size b" by simp have "a * b dvd b" by (rule dvd_euclidean_size_eq_imp_dvd [OF _ eq]) (insert assms, simp_all) hence "a * b dvd 1 * b" by simp with \b \ 0\ have "is_unit a" by (subst (asm) dvd_times_right_cancel_iff) with assms(3) show False by contradiction qed lemma dvd_imp_size_le: assumes "a dvd b" "b \ 0" shows "euclidean_size a \ euclidean_size b" using assms by (auto elim!: dvdE simp: size_mult_mono) lemma dvd_proper_imp_size_less: assumes "a dvd b" "\ b dvd a" "b \ 0" shows "euclidean_size a < euclidean_size b" proof - from assms(1) obtain c where "b = a * c" by (erule dvdE) hence z: "b = c * a" by (simp add: mult.commute) from z assms have "\is_unit c" by (auto simp: mult.commute mult_unit_dvd_iff) with z assms show ?thesis 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) lemma mod_eq_self_iff_div_eq_0: "a mod b = a \ a div b = 0" (is "?P \ ?Q") proof assume ?P with div_mult_mod_eq [of a b] show ?Q by auto next assume ?Q with div_mult_mod_eq [of a b] show ?P by simp qed lemma coprime_mod_left_iff [simp]: "coprime (a mod b) b \ coprime a b" if "b \ 0" by (rule; rule coprimeI) (use that in \auto dest!: dvd_mod_imp_dvd coprime_common_divisor simp add: dvd_mod_iff\) lemma coprime_mod_right_iff [simp]: "coprime a (b mod a) \ coprime a b" if "a \ 0" using that coprime_mod_left_iff [of a b] by (simp add: ac_simps) end class euclidean_ring = idom_modulo + euclidean_semiring begin lemma dvd_diff_commute [ac_simps]: "a dvd c - b \ a dvd b - c" proof - have "a dvd c - b \ a dvd (c - b) * - 1" by (subst dvd_mult_unit_iff) simp_all then show ?thesis by simp qed end subsection \Euclidean (semi)rings with cancel rules\ class euclidean_semiring_cancel = euclidean_semiring + assumes div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin lemma div_mult_self2 [simp]: assumes "b \ 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 \ 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 \ 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 "\ = (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 \ 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 \ 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 \ 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 "\ = a div b" by (simp only: mod_div_mult_eq) also have "\ = 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 "\ = 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 \c dvd b\ obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: \b = c * k\) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: ac_simps) also have "\ = a mod c" by (simp only: div_mult_mod_eq) finally show ?thesis . qed lemma div_mult_mult2 [simp]: "c \ 0 \ (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 \ k dvd n \ 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 \ (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 \ (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) lemma sum_div_partition: \(\a\A. f a) div b = (\a\A \ {a. b dvd f a}. f a div b) + (\a\A \ {a. \ b dvd f a}. f a) div b\ if \finite A\ proof - have \A = A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}\ by auto then have \(\a\A. f a) = (\a\A \ {a. b dvd f a} \ A \ {a. \ b dvd f a}. f a)\ by simp also have \\ = (\a\A \ {a. b dvd f a}. f a) + (\a\A \ {a. \ b dvd f a}. f a)\ using \finite A\ by (auto intro: sum.union_inter_neutral) finally have *: \sum f A = sum f (A \ {a. b dvd f a}) + sum f (A \ {a. \ b dvd f a})\ . define B where B: \B = A \ {a. b dvd f a}\ with \finite A\ have \finite B\ and \a \ B \ b dvd f a\ for a by simp_all then have \(\a\B. f a) div b = (\a\B. f a div b)\ and \b dvd (\a\B. f a)\ by induction (simp_all add: div_plus_div_distrib_dvd_left) then show ?thesis using * by (simp add: B div_plus_div_distrib_dvd_left) qed named_theorems mod_simps text \Addition respects modular equivalence.\ 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 "\ = (a mod c + b + a div c * c) mod c" by (simp only: ac_simps) also have "\ = (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]: "(\i\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 "(\i\insert i A. f i mod a) mod a = (f i mod a + (\i\A. f i mod a)) mod a" by simp also have "\ = (f i + (\i\A. f i mod a) mod a) mod a" by (simp add: mod_simps) also have "\ = (f i + (\i\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 \Multiplication respects modular equivalence.\ 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 "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (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]: "(\i\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 "(\i\insert i A. f i mod a) mod a = (f i mod a * (\i\A. f i mod a)) mod a" by simp also have "\ = (f i * ((\i\A. f i mod a) mod a)) mod a" by (simp add: mod_simps) also have "\ = (f i * ((\i\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 \Exponentiation respects modular equivalence.\ 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 lemma power_diff_power_eq: \a ^ m div a ^ n = (if n \ m then a ^ (m - n) else 1 div a ^ (n - m))\ if \a \ 0\ proof (cases \n \ m\) case True with that power_diff [symmetric, of a n m] show ?thesis by simp next case False then obtain q where n: \n = m + Suc q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \a ^ m div a ^ n = (a ^ m * 1) div (a ^ m * a ^ Suc q)\ by (simp add: power_add ac_simps) moreover from that have \a ^ m \ 0\ by simp ultimately have \a ^ m div a ^ n = 1 div a ^ Suc q\ by (subst (asm) div_mult_mult1) simp with False n show ?thesis by simp 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 \Negation respects modular equivalence.\ 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 "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp add: ac_simps) also have "\ = (- (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 \Subtraction respects modular equivalence.\ 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 \ c dvd a - b" (is "?P \ ?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 mod_eqE: assumes "a mod c = b mod c" obtains d where "b = a + c * d" proof - from assms have "c dvd a - b" by (simp add: mod_eq_dvd_iff) then obtain d where "a - b = c * d" .. then have "b = a + c * - d" by (simp add: algebra_simps) with that show thesis . qed lemma invertible_coprime: "coprime a c" if "a * b mod c = 1" by (rule coprimeI) (use that dvd_mod_iff [of _ c "a * b"] in auto) end subsection \Uniquely determined division\ class unique_euclidean_semiring = euclidean_semiring + assumes euclidean_size_mult: "euclidean_size (a * b) = euclidean_size a * euclidean_size b" fixes division_segment :: "'a \ 'a" assumes is_unit_division_segment [simp]: "is_unit (division_segment a)" and division_segment_mult: "a \ 0 \ b \ 0 \ division_segment (a * b) = division_segment a * division_segment b" and division_segment_mod: "b \ 0 \ \ b dvd a \ division_segment (a mod b) = division_segment b" assumes div_bounded: "b \ 0 \ division_segment r = division_segment b \ euclidean_size r < euclidean_size b \ (q * b + r) div b = q" begin lemma division_segment_not_0 [simp]: "division_segment a \ 0" using is_unit_division_segment [of a] is_unitE [of "division_segment a"] by blast lemma divmod_cases [case_names divides remainder by0]: obtains (divides) q where "b \ 0" and "a div b = q" and "a mod b = 0" and "a = q * b" | (remainder) q r where "b \ 0" and "division_segment r = division_segment b" and "euclidean_size r < euclidean_size b" and "r \ 0" and "a div b = q" and "a mod b = r" and "a = q * b + r" | (by0) "b = 0" proof (cases "b = 0") case True then show thesis by (rule by0) next case False show thesis proof (cases "b dvd a") case True then obtain q where "a = b * q" .. with \b \ 0\ divides show thesis by (simp add: ac_simps) next case False then have "a mod b \ 0" by (simp add: mod_eq_0_iff_dvd) moreover from \b \ 0\ \\ b dvd a\ have "division_segment (a mod b) = division_segment b" by (rule division_segment_mod) moreover have "euclidean_size (a mod b) < euclidean_size b" using \b \ 0\ by (rule mod_size_less) moreover have "a = a div b * b + a mod b" by (simp add: div_mult_mod_eq) ultimately show thesis using \b \ 0\ by (blast intro!: remainder) qed qed lemma div_eqI: "a div b = q" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "(q * b + r) div b = q" by (auto intro: div_bounded) with that show ?thesis by simp qed lemma mod_eqI: "a mod b = r" if "b \ 0" "division_segment r = division_segment b" "euclidean_size r < euclidean_size b" "q * b + r = a" proof - from that have "a div b = q" by (rule div_eqI) moreover have "a div b * b + a mod b = a" by (fact div_mult_mod_eq) ultimately have "a div b * b + a mod b = a div b * b + r" using \q * b + r = a\ by simp then show ?thesis by simp qed subclass euclidean_semiring_cancel proof show "(a + c * b) div b = c + a div b" if "b \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 with \b \ 0\ 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 \ 0" for a b c proof (cases a b rule: divmod_cases) case by0 then show ?thesis by simp next case (divides q) with \c \ 0\ show ?thesis by (simp add: mult.left_commute [of c]) next case (remainder q r) from \b \ 0\ \c \ 0\ have "b * c \ 0" by simp from remainder \c \ 0\ have "division_segment (r * c) = division_segment (b * c)" and "euclidean_size (r * c) < euclidean_size (b * c)" by (simp_all add: division_segment_mult division_segment_mod euclidean_size_mult) with remainder show ?thesis by (auto intro!: div_eqI [of _ "c * (a mod b)"] simp add: algebra_simps) (use \b * c \ 0\ 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 "\ = (a * (b div c) + q) * c" using divides by (simp add: algebra_simps) finally have "(a * b) div 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 * b = a * (b div c * c + b mod c)" by (simp add: div_mult_mod_eq) also have "\ = 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 "\ = (a div c + b div c) * c + (a mod c + b mod c)" by (simp add: algebra_simps) also have "\ = (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 "\ = 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 lemma div_eq_0_iff: "a div b = 0 \ euclidean_size a < euclidean_size b \ b = 0" (is "_ \ ?P") if "division_segment a = division_segment b" proof assume ?P with that show "a div b = 0" by (cases "b = 0") (auto intro: div_eqI) next assume "a div b = 0" then have "a mod b = a" using div_mult_mod_eq [of a b] by simp with mod_size_less [of b a] show ?P by auto qed end class unique_euclidean_ring = euclidean_ring + unique_euclidean_semiring begin subclass euclidean_ring_cancel .. end subsection \Euclidean division on \<^typ>\nat\\ instantiation nat :: normalization_semidom begin definition normalize_nat :: "nat \ nat" where [simp]: "normalize = (id :: nat \ nat)" definition unit_factor_nat :: "nat \ 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 \ nat \ nat" where "m div n = (if n = 0 then 0 else Max {k::nat. k * n \ m})" instance by standard (auto simp add: divide_nat_def ac_simps unit_factor_nat_def intro: Max_eqI) end lemma coprime_Suc_0_left [simp]: "coprime (Suc 0) n" using coprime_1_left [of n] by simp lemma coprime_Suc_0_right [simp]: "coprime n (Suc 0)" using coprime_1_right [of n] by simp lemma coprime_common_divisor_nat: "coprime a b \ x dvd a \ x dvd b \ x = 1" for a b :: nat by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation nat :: unique_euclidean_semiring begin definition euclidean_size_nat :: "nat \ nat" where [simp]: "euclidean_size_nat = id" definition division_segment_nat :: "nat \ nat" where [simp]: "division_segment_nat n = 1" definition modulo_nat :: "nat \ nat \ nat" where "m mod n = m - (m div n * (n::nat))" instance proof fix m n :: nat have ex: "\k. k * n \ l" for l :: nat by (rule exI [of _ 0]) simp have fin: "finite {k. k * n \ l}" if "n > 0" for l proof - from that have "{k. k * n \ l} \ {k. k \ 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 \ m \ n dvd l}" proof (cases "n = 0") case True moreover have "{l. l = 0 \ l \ m} = {0::nat}" by auto ultimately show ?thesis by simp next case False with ex [of m] fin have "n * Max {k. k * n \ m} = Max (times n ` {k. k * n \ m})" by (auto simp add: nat_mult_max_right intro: hom_Max_commute) also have "times n ` {k. k * n \ m} = {l. l \ m \ 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 \ 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 \ 0" show "euclidean_size (m mod n) < euclidean_size n" proof - have "m < Suc (m div n) * n" proof (rule ccontr) assume "\ m < Suc (m div n) * n" then have "Suc (m div n) * n \ m" by (simp add: not_less) moreover from \n \ 0\ have "Max {k. k * n \ m} < Suc (m div n)" by (simp add: divide_nat_def) with \n \ 0\ ex fin have "\k. k * n \ m \ 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 \n \ 0\ show ?thesis by (simp add: modulo_nat_def) qed show "euclidean_size m \ euclidean_size (m * n)" using \n \ 0\ 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 \ q" if "k * n \ q * n + r" for k proof (rule ccontr) assume "\ k \ q" then have "q < k" by simp then obtain l where "k = Suc (q + l)" by (auto simp add: less_iff_Suc_add) with \r < n\ that show False by (simp add: algebra_simps) qed with \n \ 0\ ex fin show ?thesis by (auto simp add: divide_nat_def Max_eq_iff) qed qed simp_all end text \Tool support\ ML \ 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}) ) \ simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \K Cancel_Div_Mod_Nat.proc\ lemma div_nat_eqI: "m div n = q" if "n * q \ m" and "m < n * Suc q" for m n q :: nat by (rule div_eqI [of _ "m - n * q"]) (use that in \simp_all add: algebra_simps\) lemma mod_nat_eqI: "m mod n = r" if "r < n" and "r \ m" and "n dvd m - r" for m n r :: nat by (rule mod_eqI [of _ _ "(m - r) div n"]) (use that in \simp_all add: algebra_simps\) 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 \ 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 \ 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) \ 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 \ n" using mod_less_divisor [of "Suc n" m] by arith lemma mod_less_eq_dividend [simp]: "m mod n \ 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 \ 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 \ m" for m n :: nat proof - from \n \ m\ obtain q where "m = n + q" by (auto simp add: le_iff_add) with \0 < n\ show ?thesis by (simp add: div_add_self1) qed lemma le_mod_geq: "m mod n = (m - n) mod n" if "n \ m" for m n :: nat proof - from \n \ m\ 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 \ 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 \ m < n \ n = 0" for m n :: nat by (simp add: div_eq_0_iff) lemma div_greater_zero_iff: "m div n > 0 \ n \ m \ 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 \ \ 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 \ 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 \ 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)+ lemma Suc_0_mod_eq [simp]: "Suc 0 mod n = of_bool (n \ Suc 0)" by (cases n) simp_all 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 \n > 0\ have "m mod n < n" by (rule mod_less_divisor) from \q > 0\ 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 \m mod n < n\ 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 \ 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 \ 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 \ n div k" if "m \ 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 \Antimonotonicity of \<^const>\divide\ in second argument\ lemma div_le_mono2: "k div n \ k div m" if "0 < m" and "m \ n" for m n k :: nat using that proof (induct k arbitrary: m rule: less_induct) case (less k) show ?case proof (cases "n \ k") case False then show ?thesis by simp next case True have "(k - n) div n \ (k - m) div n" using less.prems by (blast intro: div_le_mono diff_le_mono2) also have "\ \ (k - m) div m" using \n \ k\ less.prems less.hyps [of "k - m" m] by simp finally show ?thesis using \n \ k\ less.prems by (simp add: le_div_geq) qed qed lemma div_le_dividend [simp]: "m div n \ 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 \ 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 \ 0" by (rule ccontr) (use that P in auto) moreover assume "n \ 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 \ m" by simp then have "m div n * n < i * n" using that by (rule le_less_trans) with \n > 0\ show ?thesis by simp qed text \A fact for the mutilated chess board\ 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 "\ = ?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 \ (1::nat)" by (rule Suc_times_mod_eq) (use that in simp) lemma Suc_div_le_mono [simp]: "m div n \ Suc m div n" by (simp add: div_le_mono) text \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.\ 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 \An ``induction'' law for modulus arithmetic.\ lemma mod_induct [consumes 3, case_names step]: "P m" if "P n" and "n < p" and "m < p" and step: "\n. n < p \ P n \ P (Suc n mod p)" using \m < p\ proof (induct m) case 0 show ?case proof (rule ccontr) assume "\ P 0" from \n < p\ have "0 < p" by simp from \n < p\ obtain m where "0 < m" and "p = n + m" by (blast dest: less_imp_add_positive) with \P n\ have "P (p - m)" by simp moreover have "\ P (p - m)" using \0 < m\ proof (induct m) case 0 then show ?case by simp next case (Suc m) show ?case proof assume P: "P (p - Suc m)" with \\ P 0\ have "Suc m < p" by (auto intro: ccontr) then have "Suc (p - Suc m) = p - m" by arith moreover from \0 < p\ have "p - Suc m < p" by arith with P step have "P ((Suc (p - Suc m)) mod p)" by blast ultimately show False using \\ P 0\ 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 \m < p\ have "P m" by (rule Suc.hyps) with \m < p\ have "P (Suc m mod p)" by (rule step) with mod show ?case by simp qed lemma split_div: "P (m div n) \ (n = 0 \ P 0) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ 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 *: "\i j. j < n \ m = n * i + j \ P i" by simp with False show ?P by (auto intro: * [of "m mod n"]) qed qed lemma split_div': "P (m div n) \ n = 0 \ P 0 \ (\q. (n * q \ m \ m < n * Suc q) \ P q)" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n * q \ m \ m < n * Suc q \ 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) \ (n = 0 \ P m) \ (n \ 0 \ (\i j. j < n \ m = n * i + j \ P j))" (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 *: "\i j. j < n \ m = n * i + j \ P j" by simp with False show ?P by (auto intro: * [of _ "m div n"]) qed qed +lemma funpow_mod_eq: \<^marker>\contributor \Lars Noschinski\\ + \(f ^^ (m mod n)) x = (f ^^ m) x\ if \(f ^^ n) x = x\ +proof - + have \(f ^^ m) x = (f ^^ (m mod n + m div n * n)) x\ + by simp + also have \\ = (f ^^ (m mod n)) (((f ^^ n) ^^ (m div n)) x)\ + by (simp only: funpow_add funpow_mult ac_simps) simp + also have \((f ^^ n) ^^ q) x = x\ for q + by (induction q) (use \(f ^^ n) x = x\ in simp_all) + finally show ?thesis + by simp +qed + subsection \Euclidean division on \<^typ>\int\\ instantiation int :: normalization_semidom begin definition normalize_int :: "int \ int" where [simp]: "normalize = (abs :: int \ int)" definition unit_factor_int :: "int \ int" where [simp]: "unit_factor = (sgn :: int \ int)" definition divide_int :: "int \ int \ int" where "k div l = (if l = 0 then 0 else if sgn k = sgn l then int (nat \k\ div nat \l\) else - int (nat \k\ div nat \l\ + of_bool (\ l dvd k)))" lemma divide_int_unfold: "(sgn k * int m) div (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ n = 0 then 0 else if sgn k = sgn l then int (m div n) else - int (m div n + of_bool (\ n dvd m)))" by (auto simp add: divide_int_def sgn_0_0 sgn_1_pos sgn_mult abs_mult nat_mult_distrib) instance proof fix k :: int show "k div 0 = 0" by (simp add: divide_int_def) next fix k l :: int assume "l \ 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 \l \ 0\ 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 lemma coprime_int_iff [simp]: "coprime (int m) (int n) \ coprime m n" (is "?P \ ?Q") proof assume ?P show ?Q proof (rule coprimeI) fix q assume "q dvd m" "q dvd n" then have "int q dvd int m" "int q dvd int n" by simp_all with \?P\ have "is_unit (int q)" by (rule coprime_common_divisor) then show "is_unit q" by simp qed next assume ?Q show ?P proof (rule coprimeI) fix k assume "k dvd int m" "k dvd int n" then have "nat \k\ dvd m" "nat \k\ dvd n" by simp_all with \?Q\ have "is_unit (nat \k\)" by (rule coprime_common_divisor) then show "is_unit k" by simp qed qed lemma coprime_abs_left_iff [simp]: "coprime \k\ l \ coprime k l" for k l :: int using coprime_normalize_left_iff [of k l] by simp lemma coprime_abs_right_iff [simp]: "coprime k \l\ \ coprime k l" for k l :: int using coprime_abs_left_iff [of l k] by (simp add: ac_simps) lemma coprime_nat_abs_left_iff [simp]: "coprime (nat \k\) n \ coprime k (int n)" proof - define m where "m = nat \k\" then have "\k\ = int m" by simp moreover have "coprime k (int n) \ coprime \k\ (int n)" by simp ultimately show ?thesis by simp qed lemma coprime_nat_abs_right_iff [simp]: "coprime n (nat \k\) \ coprime (int n) k" using coprime_nat_abs_left_iff [of k n] by (simp add: ac_simps) lemma coprime_common_divisor_int: "coprime a b \ x dvd a \ x dvd b \ \x\ = 1" for a b :: int by (drule coprime_common_divisor [of _ _ x]) simp_all instantiation int :: idom_modulo begin definition modulo_int :: "int \ int \ int" where "k mod l = (if l = 0 then k else if sgn k = sgn l then sgn l * int (nat \k\ mod nat \l\) else sgn l * (\l\ * of_bool (\ l dvd k) - int (nat \k\ mod nat \l\)))" lemma modulo_int_unfold: "(sgn k * int m) mod (sgn l * int n) = (if sgn l = 0 \ sgn k = 0 \ 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 (\ 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) 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]) qed end instantiation int :: unique_euclidean_ring begin definition euclidean_size_int :: "int \ nat" where [simp]: "euclidean_size_int = (nat \ abs :: int \ nat)" definition division_segment_int :: "int \ int" where "division_segment_int k = (if k \ 0 then 1 else - 1)" lemma division_segment_eq_sgn: "division_segment k = sgn k" if "k \ 0" for k :: int using that by (simp add: division_segment_int_def) lemma abs_division_segment [simp]: "\division_segment k\ = 1" for k :: int by (simp add: division_segment_int_def) lemma abs_mod_less: "\k mod l\ < \l\" if "l \ 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 \ 0" "\ 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) (simp add: dvd_eq_mod_eq_0) qed instance proof fix k l :: int show "division_segment (k mod l) = division_segment l" if "l \ 0" and "\ l dvd k" using that by (simp add: division_segment_eq_sgn dvd_eq_mod_eq_0 sgn_mod) 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 \l \ 0\ with l have "s \ 0" and "n > 0" by (simp_all add: sgn_0_0) assume "division_segment r = division_segment l" moreover have "r = sgn r * \r\" by (simp add: sgn_mult_abs) moreover define u where "u = nat \r\" ultimately have "r = sgn l * int u" using division_segment_eq_sgn \l \ 0\ by (cases "r = 0") simp_all with l \n > 0\ have r: "r = sgn s * int u" by (simp add: sgn_mult) assume "euclidean_size r < euclidean_size l" with l r \s \ 0\ have "u < n" by (simp add: abs_mult) show "(q * l + r) div l = q" proof (cases "q = 0 \ r = 0") case True then show ?thesis proof assume "q = 0" then show ?thesis using l r \u < n\ by (simp add: divide_int_unfold) next assume "r = 0" from \r = 0\ have *: "q * l + r = sgn (t * s) * int (n * m)" using q l by (simp add: ac_simps sgn_mult) from \s \ 0\ \n > 0\ 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 \ 0" and "m > 0" and "s \ 0" and "u > 0" by (simp_all add: sgn_0_0) moreover from \0 < m\ \u < n\ have "u \ 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 \0 < m\ \u < n\ that by (auto intro: div_nat_eqI simp add: algebra_simps) moreover have "n dvd m * n - u \ n dvd u" using \u \ m * n\ dvd_diffD1 [of n "m * n" u] by auto ultimately show ?thesis using \s \ 0\ \m > 0\ \u > 0\ \u < n\ \u \ m * n\ 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 \auto simp add: division_segment_int_def not_le zero_less_mult_iff mult_less_0_iff abs_mult sgn_mult abs_mod_less sgn_mod nat_mult_distrib\) 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 (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) simp_all moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma neg_mod_bound [simp]: "l < k mod l" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) (simp add: mod_greater_zero_iff_not_dvd) qed lemma pos_mod_sign [simp]: "0 \ k mod l" if "l > 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain n where "l = sgn 1 * int n" by (cases l) auto moreover from this that have "n > 0" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma neg_mod_sign [simp]: "k mod l \ 0" if "l < 0" for k l :: int proof - obtain m and s where "k = sgn s * int m" by (rule int_sgnE) moreover from that obtain q where "l = sgn (- 1) * int (Suc q)" by (cases l) simp_all moreover define n where "n = Suc q" then have "Suc q = n" by simp ultimately show ?thesis by (simp only: modulo_int_unfold) simp qed lemma div_pos_pos_trivial [simp]: "k div l = 0" if "k \ 0" and "k < l" for k l :: int using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_pos_pos_trivial [simp]: "k mod l = k" if "k \ 0" and "k < l" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_neg_neg_trivial [simp]: "k div l = 0" if "k \ 0" and "l < k" for k l :: int using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) lemma mod_neg_neg_trivial [simp]: "k mod l = k" if "k \ 0" and "l < k" for k l :: int using that by (simp add: mod_eq_self_iff_div_eq_0) lemma div_pos_neg_trivial: "k div l = - 1" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule div_eqI [of _ "k + l"]) using False that apply (simp_all add: division_segment_int_def) done qed lemma mod_pos_neg_trivial: "k mod l = k + l" if "0 < k" and "k + l \ 0" for k l :: int proof (cases \l = - k\) case True with that show ?thesis by (simp add: divide_int_def) next case False show ?thesis apply (rule mod_eqI [of _ _ \- 1\]) using False that apply (simp_all add: division_segment_int_def) done qed text \There is neither \div_neg_pos_trivial\ nor \mod_neg_pos_trivial\ because \<^term>\0 div l = 0\ would supersede it.\ text \Distributive laws for function \nat\.\ lemma nat_div_distrib: \nat (x div y) = nat x div nat y\ if \0 \ x\ using that by (simp add: divide_int_def sgn_if) lemma nat_div_distrib': \nat (x div y) = nat x div nat y\ if \0 \ y\ using that by (simp add: divide_int_def sgn_if) lemma nat_mod_distrib: \ \Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't\ \nat (x mod y) = nat x mod nat y\ if \0 \ x\ \0 \ y\ using that by (simp add: modulo_int_def sgn_if) subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?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 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 lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" proof (cases a "of_nat m * of_nat n" rule: divmod_cases) case (divides q) then show ?thesis using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] by (simp add: ac_simps) next case (remainder q r) then have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by simp with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with remainder(6) have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with remainder(1) have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp with remainder(5) remainder(7) show ?thesis using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] by (simp add: ac_simps) next case by0 then show ?thesis by auto qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma exp_mod_exp: \2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ proof - have \(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\ (is \?lhs = ?rhs\) by (auto simp add: not_less monoid_mult_class.power_add dest!: le_Suc_ex) then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod) qed lemma mask_mod_exp: \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\ proof - have \(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\ (is \?lhs = ?rhs\) proof (cases \n \ m\) case True then show ?thesis by (simp add: Suc_le_lessD min.absorb2) next case False then have \m < n\ by simp then obtain q where n: \n = Suc q + m\ by (auto dest: less_imp_Suc_add) then have \min m n = m\ by simp moreover have \(2::nat) ^ m \ 2 * 2 ^ q * 2 ^ m\ using mult_le_mono1 [of 1 \2 * 2 ^ q\ \2 ^ m\] by simp with n have \2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\ by (simp add: monoid_mult_class.power_add algebra_simps) ultimately show ?thesis by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp qed then have \of_nat ?lhs = of_nat ?rhs\ by simp then show ?thesis by (simp add: of_nat_mod of_nat_diff) qed lemma of_bool_half_eq_0 [simp]: \of_bool b div 2 = 0\ by simp end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) instance int :: unique_euclidean_ring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def) subsection \Code generation\ code_identifier code_module Euclidean_Division \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Finite_Set.thy b/src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy +++ b/src/HOL/Finite_Set.thy @@ -1,2346 +1,2347 @@ (* Title: HOL/Finite_Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad Author: Andrei Popescu *) section \Finite sets\ theory Finite_Set imports Product_Type Sum_Type Fields begin subsection \Predicate for finite sets\ context notes [[inductive_internals]] begin inductive finite :: "'a set \ bool" where emptyI [simp, intro!]: "finite {}" | insertI [simp, intro!]: "finite A \ finite (insert a A)" end simproc_setup finite_Collect ("finite (Collect P)") = \K Set_Comprehension_Pointfree.simproc\ declare [[simproc del: finite_Collect]] lemma finite_induct [case_names empty insert, induct set: finite]: \ \Discharging \x \ F\ entails extra work.\ assumes "finite F" assumes "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P F" using \finite F\ proof induct show "P {}" by fact next fix x F assume F: "finite F" and P: "P F" show "P (insert x F)" proof cases assume "x \ F" then have "insert x F = F" by (rule insert_absorb) with P show ?thesis by (simp only:) next assume "x \ F" from F this P show ?thesis by (rule insert) qed qed lemma infinite_finite_induct [case_names infinite empty insert]: assumes infinite: "\A. \ finite A \ P A" and empty: "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P A" proof (cases "finite A") case False with infinite show ?thesis . next case True then show ?thesis by (induct A) (fact empty insert)+ qed subsubsection \Choice principles\ lemma ex_new_if_finite: \ \does not depend on def of finite at all\ assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast then show ?thesis by blast qed text \A finite choice principle. Does not need the SOME choice operator.\ lemma finite_set_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert a A) then obtain f b where f: "\x\A. P x (f x)" and ab: "P a b" by auto show ?case (is "\f. ?P f") proof show "?P (\x. if x = a then b else f x)" using f ab by auto qed qed subsubsection \Finite sets are the images of initial segments of natural numbers\ lemma finite_imp_nat_seg_image_inj_on: assumes "finite A" shows "\(n::nat) f. A = f ` {i. i < n} \ inj_on f {i. i < n}" using assms proof induct case empty show ?case proof show "\f. {} = f ` {i::nat. i < 0} \ inj_on f {i. i < 0}" by simp qed next case (insert a A) have notinA: "a \ A" by fact from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast then have "insert a A = f(n:=a) ` {i. i < Suc n}" and "inj_on (f(n:=a)) {i. i < Suc n}" using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq) then show ?case by blast qed lemma nat_seg_image_imp_finite: "A = f ` {i::nat. i < n} \ finite A" proof (induct n arbitrary: A) case 0 then show ?case by simp next case (Suc n) let ?B = "f ` {i. i < n}" have finB: "finite ?B" by (rule Suc.hyps[OF refl]) show ?case proof (cases "\k (\n f. A = f ` {i::nat. i < n})" by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on) lemma finite_imp_inj_to_nat_seg: assumes "finite A" shows "\f n. f ` A = {i::nat. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on [OF \finite A\] obtain f and n :: nat where bij: "bij_betw f {i. i ?f ` A = {i. i k}" by (simp add: le_eq_less_or_eq Collect_disj_eq) subsection \Finiteness and common set operations\ lemma rev_finite_subset: "finite B \ A \ B \ finite A" proof (induct arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert x F A) have A: "A \ insert x F" and r: "A - {x} \ F \ finite (A - {x})" by fact+ show "finite A" proof cases assume x: "x \ A" with A have "A - {x} \ F" by (simp add: subset_insert_iff) with r have "finite (A - {x})" . then have "finite (insert x (A - {x}))" .. also have "insert x (A - {x}) = A" using x by (rule insert_Diff) finally show ?thesis . next show ?thesis when "A \ F" using that by fact assume "x \ A" with A show "A \ F" by (simp add: subset_insert_iff) qed qed lemma finite_subset: "A \ B \ finite B \ finite A" by (rule rev_finite_subset) lemma finite_UnI: assumes "finite F" and "finite G" shows "finite (F \ G)" using assms by induct simp_all lemma finite_Un [iff]: "finite (F \ G) \ finite F \ finite G" by (blast intro: finite_UnI finite_subset [of _ "F \ G"]) lemma finite_insert [simp]: "finite (insert a A) \ finite A" proof - have "finite {a} \ finite A \ finite A" by simp then have "finite ({a} \ A) \ finite A" by (simp only: finite_Un) then show ?thesis by simp qed lemma finite_Int [simp, intro]: "finite F \ finite G \ finite (F \ G)" by (blast intro: finite_subset) lemma finite_Collect_conjI [simp, intro]: "finite {x. P x} \ finite {x. Q x} \ finite {x. P x \ Q x}" by (simp add: Collect_conj_eq) lemma finite_Collect_disjI [simp]: "finite {x. P x \ Q x} \ finite {x. P x} \ finite {x. Q x}" by (simp add: Collect_disj_eq) lemma finite_Diff [simp, intro]: "finite A \ finite (A - B)" by (rule finite_subset, rule Diff_subset) lemma finite_Diff2 [simp]: assumes "finite B" shows "finite (A - B) \ finite A" proof - have "finite A \ finite ((A - B) \ (A \ B))" by (simp add: Un_Diff_Int) also have "\ \ finite (A - B)" using \finite B\ by simp finally show ?thesis .. qed lemma finite_Diff_insert [iff]: "finite (A - insert a B) \ finite (A - B)" proof - have "finite (A - B) \ finite (A - B - {a})" by simp moreover have "A - insert a B = A - B - {a}" by auto ultimately show ?thesis by simp qed lemma finite_compl [simp]: "finite (A :: 'a set) \ finite (- A) \ finite (UNIV :: 'a set)" by (simp add: Compl_eq_Diff_UNIV) lemma finite_Collect_not [simp]: "finite {x :: 'a. P x} \ finite {x. \ P x} \ finite (UNIV :: 'a set)" by (simp add: Collect_neg_eq) lemma finite_Union [simp, intro]: "finite A \ (\M. M \ A \ finite M) \ finite (\A)" by (induct rule: finite_induct) simp_all lemma finite_UN_I [intro]: "finite A \ (\a. a \ A \ finite (B a)) \ finite (\a\A. B a)" by (induct rule: finite_induct) simp_all lemma finite_UN [simp]: "finite A \ finite (\(B ` A)) \ (\x\A. finite (B x))" by (blast intro: finite_subset) lemma finite_Inter [intro]: "\A\M. finite A \ finite (\M)" by (blast intro: Inter_lower finite_subset) lemma finite_INT [intro]: "\x\I. finite (A x) \ finite (\x\I. A x)" by (blast intro: INT_lower finite_subset) lemma finite_imageI [simp, intro]: "finite F \ finite (h ` F)" by (induct rule: finite_induct) simp_all lemma finite_image_set [simp]: "finite {x. P x} \ finite {f x |x. P x}" by (simp add: image_Collect [symmetric]) lemma finite_image_set2: "finite {x. P x} \ finite {y. Q y} \ finite {f x y |x y. P x \ Q y}" by (rule finite_subset [where B = "\x \ {x. P x}. \y \ {y. Q y}. {f x y}"]) auto lemma finite_imageD: assumes "finite (f ` A)" and "inj_on f A" shows "finite A" using assms proof (induct "f ` A" arbitrary: A) case empty then show ?case by simp next case (insert x B) then have B_A: "insert x B = f ` A" by simp then obtain y where "x = f y" and "y \ A" by blast from B_A \x \ B\ have "B = f ` A - {x}" by blast with B_A \x \ B\ \x = f y\ \inj_on f A\ \y \ A\ have "B = f ` (A - {y})" by (simp add: inj_on_image_set_diff) moreover from \inj_on f A\ have "inj_on f (A - {y})" by (rule inj_on_diff) ultimately have "finite (A - {y})" by (rule insert.hyps) then show "finite A" by simp qed lemma finite_image_iff: "inj_on f A \ finite (f ` A) \ finite A" using finite_imageD by blast lemma finite_surj: "finite A \ B \ f ` A \ finite B" by (erule finite_subset) (rule finite_imageI) lemma finite_range_imageI: "finite (range g) \ finite (range (\x. f (g x)))" by (drule finite_imageI) (simp add: range_composition) lemma finite_subset_image: assumes "finite B" shows "B \ f ` A \ \C\A. finite C \ B = f ` C" using assms proof induct case empty then show ?case by simp next case insert then show ?case by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) blast qed lemma all_subset_image: "(\B. B \ f ` A \ P B) \ (\B. B \ A \ P(f ` B))" by (safe elim!: subset_imageE) (use image_mono in \blast+\) (* slow *) lemma all_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and P: "\B. finite B \ B \ A \ P (f ` B)" show "P B" using finite_subset_image [OF B] P by blast qed blast lemma ex_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and "P B" show "\B. finite B \ B \ A \ P (f ` B)" using finite_subset_image [OF B] \P B\ by blast qed blast lemma finite_vimage_IntI: "finite F \ inj_on h A \ finite (h -` F \ A)" proof (induct rule: finite_induct) case (insert x F) then show ?case by (simp add: vimage_insert [of h x F] finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2) qed simp lemma finite_finite_vimage_IntI: assumes "finite F" and "\y. y \ F \ finite ((h -` {y}) \ A)" shows "finite (h -` F \ A)" proof - have *: "h -` F \ A = (\ y\F. (h -` {y}) \ A)" by blast show ?thesis by (simp only: * assms finite_UN_I) qed lemma finite_vimageI: "finite F \ inj h \ finite (h -` F)" using finite_vimage_IntI[of F h UNIV] by auto lemma finite_vimageD': "finite (f -` A) \ A \ range f \ finite A" by (auto simp add: subset_image_iff intro: finite_subset[rotated]) lemma finite_vimageD: "finite (h -` F) \ surj h \ finite F" by (auto dest: finite_vimageD') lemma finite_vimage_iff: "bij h \ finite (h -` F) \ finite F" unfolding bij_def by (auto elim: finite_vimageD finite_vimageI) lemma finite_Collect_bex [simp]: assumes "finite A" shows "finite {x. \y\A. Q x y} \ (\y\A. finite {x. Q x y})" proof - have "{x. \y\A. Q x y} = (\y\A. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Collect_bounded_ex [simp]: assumes "finite {y. P y}" shows "finite {x. \y. P y \ Q x y} \ (\y. P y \ finite {x. Q x y})" proof - have "{x. \y. P y \ Q x y} = (\y\{y. P y}. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Plus: "finite A \ finite B \ finite (A <+> B)" by (simp add: Plus_def) lemma finite_PlusD: fixes A :: "'a set" and B :: "'b set" assumes fin: "finite (A <+> B)" shows "finite A" "finite B" proof - have "Inl ` A \ A <+> B" by auto then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite A" by (rule finite_imageD) (auto intro: inj_onI) next have "Inr ` B \ A <+> B" by auto then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite B" by (rule finite_imageD) (auto intro: inj_onI) qed lemma finite_Plus_iff [simp]: "finite (A <+> B) \ finite A \ finite B" by (auto intro: finite_PlusD finite_Plus) lemma finite_Plus_UNIV_iff [simp]: "finite (UNIV :: ('a + 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff) lemma finite_SigmaI [simp, intro]: "finite A \ (\a. a\A \ finite (B a)) \ finite (SIGMA a:A. B a)" unfolding Sigma_def by blast lemma finite_SigmaI2: assumes "finite {x\A. B x \ {}}" and "\a. a \ A \ finite (B a)" shows "finite (Sigma A B)" proof - from assms have "finite (Sigma {x\A. B x \ {}} B)" by auto also have "Sigma {x:A. B x \ {}} B = Sigma A B" by auto finally show ?thesis . qed lemma finite_cartesian_product: "finite A \ finite B \ finite (A \ B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: "finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ finite (UNIV :: ('a \ 'b) set)" by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product) lemma finite_cartesian_productD1: assumes "finite (A \ B)" and "B \ {}" shows "finite A" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "fst ` (A \ B) = fst ` f ` {i::nat. i < n}" by simp with \B \ {}\ have "A = (fst \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. A = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_productD2: assumes "finite (A \ B)" and "A \ {}" shows "finite B" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "snd ` (A \ B) = snd ` f ` {i::nat. i < n}" by simp with \A \ {}\ have "B = (snd \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. B = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_product_iff: "finite (A \ B) \ (A = {} \ B = {} \ (finite A \ finite B))" by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product) lemma finite_prod: "finite (UNIV :: ('a \ 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" using finite_cartesian_product_iff[of UNIV UNIV] by simp lemma finite_Pow_iff [iff]: "finite (Pow A) \ finite A" proof assume "finite (Pow A)" then have "finite ((\x. {x}) ` A)" by (blast intro: finite_subset) (* somewhat slow *) then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp next assume "finite A" then show "finite (Pow A)" by induct (simp_all add: Pow_insert) qed corollary finite_Collect_subsets [simp, intro]: "finite A \ finite {B. B \ A}" by (simp add: Pow_def [symmetric]) lemma finite_set: "finite (UNIV :: 'a set set) \ finite (UNIV :: 'a set)" by (simp only: finite_Pow_iff Pow_UNIV[symmetric]) lemma finite_UnionD: "finite (\A) \ finite A" by (blast intro: finite_subset [OF subset_Pow_Union]) lemma finite_bind: assumes "finite S" assumes "\x \ S. finite (f x)" shows "finite (Set.bind S f)" using assms by (simp add: bind_UNION) lemma finite_filter [simp]: "finite S \ finite (Set.filter P S)" unfolding Set.filter_def by simp lemma finite_set_of_finite_funs: assumes "finite A" "finite B" shows "finite {f. \x. (x \ A \ f x \ B) \ (x \ A \ f x = d)}" (is "finite ?S") proof - let ?F = "\f. {(a,b). a \ A \ b = f a}" have "?F ` ?S \ Pow(A \ B)" by auto from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp have 2: "inj_on ?F ?S" by (fastforce simp add: inj_on_def set_eq_iff fun_eq_iff) (* somewhat slow *) show ?thesis by (rule finite_imageD [OF 1 2]) qed lemma not_finite_existsD: assumes "\ finite {a. P a}" shows "\a. P a" proof (rule classical) assume "\ ?thesis" with assms show ?thesis by auto qed subsection \Further induction rules on finite sets\ lemma finite_ne_induct [case_names singleton insert, consumes 2]: assumes "finite F" and "F \ {}" assumes "\x. P {x}" and "\x F. finite F \ F \ {} \ x \ F \ P F \ P (insert x F)" shows "P F" using assms proof induct case empty then show ?case by simp next case (insert x F) then show ?case by cases auto qed lemma finite_subset_induct [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. finite F \ a \ A \ a \ F \ P F \ P (insert a F)" shows "P F" using \finite F\ \F \ A\ proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact qed qed lemma finite_empty_induct: assumes "finite A" and "P A" and remove: "\a A. finite A \ a \ A \ P A \ P (A - {a})" shows "P {}" proof - have "P (A - B)" if "B \ A" for B :: "'a set" proof - from \finite A\ that have "finite B" by (rule rev_finite_subset) from this \B \ A\ show "P (A - B)" proof induct case empty from \P A\ show ?case by simp next case (insert b B) have "P (A - B - {b})" proof (rule remove) from \finite A\ show "finite (A - B)" by induct auto from insert show "b \ A - B" by simp from insert show "P (A - B)" by simp qed also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric]) finally show ?case . qed qed then have "P (A - A)" by blast then show ?thesis by simp qed lemma finite_update_induct [consumes 1, case_names const update]: assumes finite: "finite {a. f a \ c}" and const: "P (\a. c)" and update: "\a b f. finite {a. f a \ c} \ f a = c \ b \ c \ P f \ P (f(a := b))" shows "P f" using finite proof (induct "{a. f a \ c}" arbitrary: f) case empty with const show ?case by simp next case (insert a A) then have "A = {a'. (f(a := c)) a' \ c}" and "f a \ c" by auto with \finite A\ have "finite {a'. (f(a := c)) a' \ c}" by simp have "(f(a := c)) a = c" by simp from insert \A = {a'. (f(a := c)) a' \ c}\ have "P (f(a := c))" by simp with \finite {a'. (f(a := c)) a' \ c}\ \(f(a := c)) a = c\ \f a \ c\ have "P ((f(a := c))(a := f a))" by (rule update) then show ?case by simp qed lemma finite_subset_induct' [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. \finite F; a \ A; F \ A; a \ F; P F \ \ P (insert a F)" shows "P F" using assms(1,2) proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact show "F \ A" by fact qed qed subsection \Class \finite\\ class finite = assumes finite_UNIV: "finite (UNIV :: 'a set)" begin lemma finite [simp]: "finite (A :: 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ lemma finite_code [code]: "finite (A :: 'a set) \ True" by simp end instance prod :: (finite, finite) finite by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite) lemma inj_graph: "inj (\f. {(x, y). y = f x})" by (rule inj_onI) (auto simp add: set_eq_iff fun_eq_iff) instance "fun" :: (finite, finite) finite proof show "finite (UNIV :: ('a \ 'b) set)" proof (rule finite_imageD) let ?graph = "\f::'a \ 'b. {(x, y). y = f x}" have "range ?graph \ Pow UNIV" by simp moreover have "finite (Pow (UNIV :: ('a * 'b) set))" by (simp only: finite_Pow_iff finite) ultimately show "finite (range ?graph)" by (rule finite_subset) show "inj ?graph" by (rule inj_graph) qed qed instance bool :: finite by standard (simp add: UNIV_bool) instance set :: (finite) finite by standard (simp only: Pow_UNIV [symmetric] finite_Pow_iff finite) instance unit :: finite by standard (simp add: UNIV_unit) instance sum :: (finite, finite) finite by standard (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite) subsection \A basic fold functional for finite sets\ text \The intended behaviour is \fold f z {x\<^sub>1, \, x\<^sub>n} = f x\<^sub>1 (\ (f x\<^sub>n z)\)\ if \f\ is ``left-commutative'': \ locale comp_fun_commute = fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma fun_left_comm: "f y (f x z) = f x (f y z)" using comp_fun_commute by (simp add: fun_eq_iff) lemma commute_left_comp: "f y \ (f x \ g) = f x \ (f y \ g)" by (simp add: o_assoc comp_fun_commute) end inductive fold_graph :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_graph f z {} z" | insertI [intro]: "x \ A \ fold_graph f z A y \ fold_graph f z (insert x A) (f x y)" inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x" lemma fold_graph_closed_lemma: "fold_graph f z A x \ x \ B" if "fold_graph g z A x" "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using that(1-3) proof (induction rule: fold_graph.induct) case (insertI x A y) have "fold_graph f z A y" "y \ B" unfolding atomize_conj by (rule insertI.IH) (auto intro: insertI.prems) then have "g x y \ B" and f_eq: "f x y = g x y" by (auto simp: insertI.prems) moreover have "fold_graph f z (insert x A) (f x y)" by (rule fold_graph.insertI; fact) ultimately show ?case by (simp add: f_eq) qed (auto intro!: that) lemma fold_graph_closed_eq: "fold_graph f z A = fold_graph g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using fold_graph_closed_lemma[of f z A _ B g] fold_graph_closed_lemma[of g z A _ B f] that by auto definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)" lemma fold_closed_eq: "fold f z A = fold g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" unfolding Finite_Set.fold_def by (subst fold_graph_closed_eq[where B=B and g=g]) (auto simp: that) text \ A tempting alternative for the definiens is \<^term>\if finite A then THE y. fold_graph f z A y else e\. It allows the removal of finiteness assumptions from the theorems \fold_comm\, \fold_reindex\ and \fold_distrib\. The proofs become ugly. It is not worth the effort. (???) \ lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" by (induct rule: finite_induct) auto subsubsection \From \<^const>\fold_graph\ to \<^term>\fold\\ context comp_fun_commute begin lemma fold_graph_finite: assumes "fold_graph f z A y" shows "finite A" using assms by induct simp_all lemma fold_graph_insertE_aux: "fold_graph f z A y \ a \ A \ \y'. y = f a y' \ fold_graph f z (A - {a}) y'" proof (induct set: fold_graph) case emptyI then show ?case by simp next case (insertI x A y) show ?case proof (cases "x = a") case True with insertI show ?thesis by auto next case False then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'" using insertI by auto have "f x y = f a (f x y')" unfolding y by (rule fun_left_comm) moreover have "fold_graph f z (insert x A - {a}) (f x y')" using y' and \x \ a\ and \x \ A\ by (simp add: insert_Diff_if fold_graph.insertI) ultimately show ?thesis by fast qed qed lemma fold_graph_insertE: assumes "fold_graph f z (insert x A) v" and "x \ A" obtains y where "v = f x y" and "fold_graph f z A y" using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1]) lemma fold_graph_determ: "fold_graph f z A x \ fold_graph f z A y \ y = x" proof (induct arbitrary: y set: fold_graph) case emptyI then show ?case by fast next case (insertI x A y v) from \fold_graph f z (insert x A) v\ and \x \ A\ obtain y' where "v = f x y'" and "fold_graph f z A y'" by (rule fold_graph_insertE) from \fold_graph f z A y'\ have "y' = y" by (rule insertI) with \v = f x y'\ show "v = f x y" by simp qed lemma fold_equality: "fold_graph f z A y \ fold f z A = y" by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite) lemma fold_graph_fold: assumes "finite A" shows "fold_graph f z A (fold f z A)" proof - from assms have "\x. fold_graph f z A x" by (rule finite_imp_fold_graph) moreover note fold_graph_determ ultimately have "\!x. fold_graph f z A x" by (rule ex_ex1I) then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI') with assms show ?thesis by (simp add: fold_def) qed text \The base case for \fold\:\ lemma (in -) fold_infinite [simp]: "\ finite A \ fold f z A = z" by (auto simp: fold_def) lemma (in -) fold_empty [simp]: "fold f z {} = z" by (auto simp: fold_def) text \The various recursion equations for \<^const>\fold\:\ lemma fold_insert [simp]: assumes "finite A" and "x \ A" shows "fold f z (insert x A) = f x (fold f z A)" proof (rule fold_equality) fix z from \finite A\ have "fold_graph f z A (fold f z A)" by (rule fold_graph_fold) with \x \ A\ have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI) then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp qed declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del] \ \No more proofs involve these.\ lemma fold_fun_left_comm: "finite A \ f x (fold f z A) = fold f (f x z) A" proof (induct rule: finite_induct) case empty then show ?case by simp next case insert then show ?case by (simp add: fun_left_comm [of x]) qed lemma fold_insert2: "finite A \ x \ A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) lemma fold_rec: assumes "finite A" and "x \ A" shows "fold f z A = f x (fold f z (A - {x}))" proof - have A: "A = insert x (A - {x})" using \x \ A\ by blast then have "fold f z A = fold f z (insert x (A - {x}))" by simp also have "\ = f x (fold f z (A - {x}))" by (rule fold_insert) (simp add: \finite A\)+ finally show ?thesis . qed lemma fold_insert_remove: assumes "finite A" shows "fold f z (insert x A) = f x (fold f z (A - {x}))" proof - from \finite A\ have "finite (insert x A)" by auto moreover have "x \ insert x A" by auto ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))" by (rule fold_rec) then show ?thesis by simp qed lemma fold_set_union_disj: assumes "finite A" "finite B" "A \ B = {}" shows "Finite_Set.fold f z (A \ B) = Finite_Set.fold f (Finite_Set.fold f z A) B" using assms(2,1,3) by induct simp_all end text \Other properties of \<^const>\fold\:\ lemma fold_image: assumes "inj_on g A" shows "fold f z (g ` A) = fold (f \ g) z A" proof (cases "finite A") case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def) next case True have "fold_graph f z (g ` A) = fold_graph (f \ g) z A" proof fix w show "fold_graph f z (g ` A) w \ fold_graph (f \ g) z A w" (is "?P \ ?Q") proof assume ?P then show ?Q using assms proof (induct "g ` A" w arbitrary: A) case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r B) from \inj_on g B\ \x \ A\ \insert x A = image g B\ obtain x' A' where "x' \ A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'" by (rule inj_img_insertE) from insertI.prems have "fold_graph (f \ g) z A' r" by (auto intro: insertI.hyps) with \x' \ A'\ have "fold_graph (f \ g) z (insert x' A') ((f \ g) x' r)" by (rule fold_graph.insertI) then show ?case by simp qed next assume ?Q then show ?P using assms proof induct case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r) from \x \ A\ insertI.prems have "g x \ g ` A" by auto moreover from insertI have "fold_graph f z (g ` A) r" by simp ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)" by (rule fold_graph.insertI) then show ?case by simp qed qed qed with True assms show ?thesis by (auto simp add: fold_def) qed lemma fold_cong: assumes "comp_fun_commute f" "comp_fun_commute g" and "finite A" and cong: "\x. x \ A \ f x = g x" and "s = t" and "A = B" shows "fold f s A = fold g t B" proof - have "fold f s A = fold g s A" using \finite A\ cong proof (induct A) case empty then show ?case by simp next case insert interpret f: comp_fun_commute f by (fact \comp_fun_commute f\) interpret g: comp_fun_commute g by (fact \comp_fun_commute g\) from insert show ?case by simp qed with assms show ?thesis by simp qed text \A simplified version for idempotent functions:\ locale comp_fun_idem = comp_fun_commute + assumes comp_fun_idem: "f x \ f x = f x" begin lemma fun_left_idem: "f x (f x z) = f x z" using comp_fun_idem by (simp add: fun_eq_iff) lemma fold_insert_idem: assumes fin: "finite A" shows "fold f z (insert x A) = f x (fold f z A)" proof cases assume "x \ A" then obtain B where "A = insert x B" and "x \ B" by (rule set_insert) then show ?thesis using assms by (simp add: comp_fun_idem fun_left_idem) next assume "x \ A" then show ?thesis using assms by simp qed declare fold_insert [simp del] fold_insert_idem [simp] lemma fold_insert_idem2: "finite A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) end subsubsection \Liftings to \comp_fun_commute\ etc.\ lemma (in comp_fun_commute) comp_comp_fun_commute: "comp_fun_commute (f \ g)" by standard (simp_all add: comp_fun_commute) lemma (in comp_fun_idem) comp_comp_fun_idem: "comp_fun_idem (f \ g)" by (rule comp_fun_idem.intro, rule comp_comp_fun_commute, unfold_locales) (simp_all add: comp_fun_idem) lemma (in comp_fun_commute) comp_fun_commute_funpow: "comp_fun_commute (\x. f x ^^ g x)" proof show "f y ^^ g y \ f x ^^ g x = f x ^^ g x \ f y ^^ g y" for x y proof (cases "x = y") case True then show ?thesis by simp next case False show ?thesis proof (induct "g x" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) have hyp1: "f y ^^ g y \ f x = f x \ f y ^^ g y" proof (induct "g y" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) define h where "h z = g z - 1" for z with Suc have "n = h y" by simp with Suc have hyp: "f y ^^ h y \ f x = f x \ f y ^^ h y" by auto from Suc h_def have "g y = Suc (h y)" by simp then show ?case by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute) qed define h where "h z = (if z = x then g x - 1 else g z)" for z with Suc have "n = h x" by simp with Suc have "f y ^^ h y \ f x ^^ h x = f x ^^ h x \ f y ^^ h y" by auto with False h_def have hyp2: "f y ^^ g y \ f x ^^ h x = f x ^^ h x \ f y ^^ g y" by simp from Suc h_def have "g x = Suc (h x)" by simp then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2) (simp add: comp_assoc hyp1) qed qed qed subsubsection \Expressing set operations via \<^const>\fold\\ lemma comp_fun_commute_const: "comp_fun_commute (\_. f)" by standard rule lemma comp_fun_idem_insert: "comp_fun_idem insert" by standard auto lemma comp_fun_idem_remove: "comp_fun_idem Set.remove" by standard auto lemma (in semilattice_inf) comp_fun_idem_inf: "comp_fun_idem inf" by standard (auto simp add: inf_left_commute) lemma (in semilattice_sup) comp_fun_idem_sup: "comp_fun_idem sup" by standard (auto simp add: sup_left_commute) lemma union_fold_insert: assumes "finite A" shows "A \ B = fold insert B A" proof - interpret comp_fun_idem insert by (fact comp_fun_idem_insert) from \finite A\ show ?thesis by (induct A arbitrary: B) simp_all qed lemma minus_fold_remove: assumes "finite A" shows "B - A = fold Set.remove B A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) from \finite A\ have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto (* slow *) then show ?thesis .. qed lemma comp_fun_commute_filter_fold: "comp_fun_commute (\x A'. if P x then Set.insert x A' else A')" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff) qed lemma Set_filter_fold: assumes "finite A" shows "Set.filter P A = fold (\x A'. if P x then Set.insert x A' else A') {} A" using assms by induct (auto simp add: Set.filter_def comp_fun_commute.fold_insert[OF comp_fun_commute_filter_fold]) lemma inter_Set_filter: assumes "finite B" shows "A \ B = Set.filter (\x. x \ A) B" using assms by induct (auto simp: Set.filter_def) lemma image_fold_insert: assumes "finite A" shows "image f A = fold (\k A. Set.insert (f k) A) {} A" proof - interpret comp_fun_commute "\k A. Set.insert (f k) A" by standard auto show ?thesis using assms by (induct A) auto qed lemma Ball_fold: assumes "finite A" shows "Ball A P = fold (\k s. s \ P k) True A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma Bex_fold: assumes "finite A" shows "Bex A P = fold (\k s. s \ P k) False A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma comp_fun_commute_Pow_fold: "comp_fun_commute (\x A. A \ Set.insert x ` A)" by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast (* somewhat slow *) lemma Pow_fold: assumes "finite A" shows "Pow A = fold (\x A. A \ Set.insert x ` A) {{}} A" proof - interpret comp_fun_commute "\x A. A \ Set.insert x ` A" by (rule comp_fun_commute_Pow_fold) show ?thesis using assms by (induct A) (auto simp: Pow_insert) qed lemma fold_union_pair: assumes "finite B" shows "(\y\B. {(x, y)}) \ A = fold (\y. Set.insert (x, y)) A B" proof - interpret comp_fun_commute "\y. Set.insert (x, y)" by standard auto show ?thesis using assms by (induct arbitrary: A) simp_all qed lemma comp_fun_commute_product_fold: "finite B \ comp_fun_commute (\x z. fold (\y. Set.insert (x, y)) z B)" by standard (auto simp: fold_union_pair [symmetric]) lemma product_fold: assumes "finite A" "finite B" shows "A \ B = fold (\x z. fold (\y. Set.insert (x, y)) z B) {} A" using assms unfolding Sigma_def by (induct A) (simp_all add: comp_fun_commute.fold_insert[OF comp_fun_commute_product_fold] fold_union_pair) context complete_lattice begin lemma inf_Inf_fold_inf: assumes "finite A" shows "inf (Inf A) B = fold inf B A" proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: inf_commute fun_eq_iff) qed lemma sup_Sup_fold_sup: assumes "finite A" shows "sup (Sup A) B = fold sup B A" proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: sup_commute fun_eq_iff) qed lemma Inf_fold_inf: "finite A \ Inf A = fold inf top A" using inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2) lemma Sup_fold_sup: "finite A \ Sup A = fold sup bot A" using sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2) lemma inf_INF_fold_inf: assumes "finite A" shows "inf B (\(f ` A)) = fold (inf \ f) B A" (is "?inf = ?fold") proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) interpret comp_fun_idem "inf \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?inf" by (induct A arbitrary: B) (simp_all add: inf_left_commute) then show ?thesis .. qed lemma sup_SUP_fold_sup: assumes "finite A" shows "sup B (\(f ` A)) = fold (sup \ f) B A" (is "?sup = ?fold") proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) interpret comp_fun_idem "sup \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?sup" by (induct A arbitrary: B) (simp_all add: sup_left_commute) then show ?thesis .. qed lemma INF_fold_inf: "finite A \ \(f ` A) = fold (inf \ f) top A" using inf_INF_fold_inf [of A top] by simp lemma SUP_fold_sup: "finite A \ \(f ` A) = fold (sup \ f) bot A" using sup_SUP_fold_sup [of A bot] by simp lemma finite_Inf_in: assumes "finite A" "A\{}" and inf: "\x y. \x \ A; y \ A\ \ inf x y \ A" shows "Inf A \ A" proof - have "Inf B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use inf in \force+\) then show ?thesis by (simp add: assms) qed lemma finite_Sup_in: assumes "finite A" "A\{}" and sup: "\x y. \x \ A; y \ A\ \ sup x y \ A" shows "Sup A \ A" proof - have "Sup B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use sup in \force+\) then show ?thesis by (simp add: assms) qed end subsection \Locales as mini-packages for fold operations\ subsubsection \The natural case\ locale folding = fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin interpretation fold?: comp_fun_commute f by standard (use comp_fun_commute in \simp add: fun_eq_iff\) definition F :: "'a set \ 'b" where eq_fold: "F A = fold f z A" lemma empty [simp]:"F {} = z" by (simp add: eq_fold) lemma infinite [simp]: "\ finite A \ F A = z" by (simp add: eq_fold) lemma insert [simp]: assumes "finite A" and "x \ A" shows "F (insert x A) = f x (F A)" proof - from fold_insert assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed lemma remove: assumes "finite A" and "x \ A" shows "F A = f x (F (A - {x}))" proof - from \x \ A\ obtain B where A: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ A have "finite B" by simp ultimately show ?thesis by simp qed lemma insert_remove: "finite A \ F (insert x A) = f x (F (A - {x}))" by (cases "x \ A") (simp_all add: remove insert_absorb) end subsubsection \With idempotency\ locale folding_idem = folding + assumes comp_fun_idem: "f x \ f x = f x" begin declare insert [simp del] interpretation fold?: comp_fun_idem f by standard (insert comp_fun_commute comp_fun_idem, simp add: fun_eq_iff) lemma insert_idem [simp]: assumes "finite A" shows "F (insert x A) = f x (F A)" proof - from fold_insert_idem assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed end subsection \Finite cardinality\ text \ The traditional definition \<^prop>\card A \ LEAST n. \f. A = {f i |i. i < n}\ is ugly to work with. But now that we have \<^const>\fold\ things are easy: \ global_interpretation card: folding "\_. Suc" 0 defines card = "folding.F (\_. Suc) 0" by standard rule lemma card_insert_disjoint: "finite A \ x \ A \ card (insert x A) = Suc (card A)" by (fact card.insert) lemma card_insert_if: "finite A \ card (insert x A) = (if x \ A then card A else Suc (card A))" by auto (simp add: card.insert_remove card.remove) lemma card_ge_0_finite: "card A > 0 \ finite A" by (rule ccontr) simp lemma card_0_eq [simp]: "finite A \ card A = 0 \ A = {}" by (auto dest: mk_disjoint_insert) lemma finite_UNIV_card_ge_0: "finite (UNIV :: 'a set) \ card (UNIV :: 'a set) > 0" by (rule ccontr) simp lemma card_eq_0_iff: "card A = 0 \ A = {} \ \ finite A" by auto lemma card_range_greater_zero: "finite (range f) \ card (range f) > 0" by (rule ccontr) (simp add: card_eq_0_iff) lemma card_gt_0_iff: "0 < card A \ A \ {} \ finite A" by (simp add: neq0_conv [symmetric] card_eq_0_iff) lemma card_Suc_Diff1: assumes "finite A" "x \ A" shows "Suc (card (A - {x})) = card A" proof - have "Suc (card (A - {x})) = card (insert x (A - {x}))" using assms by (simp add: card.insert_remove) also have "... = card A" using assms by (simp add: card_insert_if) finally show ?thesis . qed lemma card_insert_le_m1: assumes "n > 0" "card y \ n - 1" shows "card (insert x y) \ n" using assms by (cases "finite y") (auto simp: card_insert_if) lemma card_Diff_singleton: "finite A \ x \ A \ card (A - {x}) = card A - 1" by (simp add: card_Suc_Diff1 [symmetric]) lemma card_Diff_singleton_if: "finite A \ card (A - {x}) = (if x \ A then card A - 1 else card A)" by (simp add: card_Diff_singleton) lemma card_Diff_insert[simp]: assumes "finite A" and "a \ A" and "a \ B" shows "card (A - insert a B) = card (A - B) - 1" proof - have "A - insert a B = (A - B) - {a}" using assms by blast then show ?thesis using assms by (simp add: card_Diff_singleton) qed lemma card_insert_le: "finite A \ card A \ card (insert x A)" by (simp add: card_insert_if) lemma card_Collect_less_nat[simp]: "card {i::nat. i < n} = n" by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq) lemma card_Collect_le_nat[simp]: "card {i::nat. i \ n} = Suc n" using card_Collect_less_nat[of "Suc n"] by (simp add: less_Suc_eq_le) lemma card_mono: assumes "finite B" and "A \ B" shows "card A \ card B" proof - from assms have "finite A" by (auto intro: finite_subset) then show ?thesis using assms proof (induct A arbitrary: B) case empty then show ?case by simp next case (insert x A) then have "x \ B" by simp from insert have "A \ B - {x}" and "finite (B - {x})" by auto with insert.hyps have "card A \ card (B - {x})" by auto with \finite A\ \x \ A\ \finite B\ \x \ B\ show ?case by simp (simp only: card.remove) qed qed lemma card_seteq: assumes "finite B" and A: "A \ B" "card B \ card A" shows "A = B" using assms proof (induction arbitrary: A rule: finite_induct) case (insert b B) then have A: "finite A" "A - {b} \ B" by force+ then have "card B \ card (A - {b})" using insert by (auto simp add: card_Diff_singleton_if) then have "A - {b} = B" using A insert.IH by auto then show ?case using insert.hyps insert.prems by auto qed auto lemma psubset_card_mono: "finite B \ A < B \ card A < card B" using card_seteq [of B A] by (auto simp add: psubset_eq) lemma card_Un_Int: assumes "finite A" "finite B" shows "card A + card B = card (A \ B) + card (A \ B)" using assms proof (induct A) case empty then show ?case by simp next case insert then show ?case by (auto simp add: insert_absorb Int_insert_left) qed lemma card_Un_disjoint: "finite A \ finite B \ A \ B = {} \ card (A \ B) = card A + card B" using card_Un_Int [of A B] by simp lemma card_Un_disjnt: "\finite A; finite B; disjnt A B\ \ card (A \ B) = card A + card B" by (simp add: card_Un_disjoint disjnt_def) lemma card_Un_le: "card (A \ B) \ card A + card B" proof (cases "finite A \ finite B") case True then show ?thesis using le_iff_add card_Un_Int [of A B] by auto qed auto lemma card_Diff_subset: assumes "finite B" and "B \ A" shows "card (A - B) = card A - card B" using assms proof (cases "finite A") case False with assms show ?thesis by simp next case True with assms show ?thesis by (induct B arbitrary: A) simp_all qed lemma card_Diff_subset_Int: assumes "finite (A \ B)" shows "card (A - B) = card A - card (A \ B)" proof - have "A - B = A - A \ B" by auto with assms show ?thesis by (simp add: card_Diff_subset) qed lemma diff_card_le_card_Diff: assumes "finite B" shows "card A - card B \ card (A - B)" proof - have "card A - card B \ card A - card (A \ B)" using card_mono[OF assms Int_lower2, of A] by arith also have "\ = card (A - B)" using assms by (simp add: card_Diff_subset_Int) finally show ?thesis . qed lemma card_le_sym_Diff: assumes "finite A" "finite B" "card A \ card B" shows "card(A - B) \ card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ \ card B - card (A \ B)" using assms(3) by linarith also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_less_sym_Diff: assumes "finite A" "finite B" "card A < card B" shows "card(A - B) < card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ < card B - card (A \ B)" using assms(1,3) by (simp add: card_mono diff_less_mono) also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_Diff1_less_iff: "card (A - {x}) < card A \ finite A \ x \ A" proof (cases "finite A \ x \ A") case True then show ?thesis by (auto simp: card_gt_0_iff intro: diff_less) qed auto lemma card_Diff1_less: "finite A \ x \ A \ card (A - {x}) < card A" unfolding card_Diff1_less_iff by auto lemma card_Diff2_less: assumes "finite A" "x \ A" "y \ A" shows "card (A - {x} - {y}) < card A" proof (cases "x = y") case True with assms show ?thesis by (simp add: card_Diff1_less del: card_Diff_insert) next case False then have "card (A - {x} - {y}) < card (A - {x})" "card (A - {x}) < card A" using assms by (intro card_Diff1_less; simp)+ then show ?thesis by (blast intro: less_trans) qed lemma card_Diff1_le: "finite A \ card (A - {x}) \ card A" by (cases "x \ A") (simp_all add: card_Diff1_less less_imp_le) lemma card_psubset: "finite B \ A \ B \ card A < card B \ A < B" by (erule psubsetI) blast lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" and c: "card A \ card B" shows "\f. f ` A \ B \ inj_on f A" using fA fB c proof (induct arbitrary: B rule: finite_induct) case empty then show ?case by simp next case (insert x s t) then show ?case proof (induct rule: finite_induct [OF insert.prems(1)]) case 1 then show ?case by simp next case (2 y t) from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \ card t" by simp from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where "f ` s \ t" "inj_on f s" by blast with "2.prems"(2) "2.hyps"(2) show ?case unfolding inj_on_def by (rule_tac x = "\z. if z = x then y else f z" in exI) auto qed qed lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and c: "card A = card B" shows "A = B" proof - from fB AB have fA: "finite A" by (auto intro: finite_subset) from fA fB have fBA: "finite (B - A)" by auto have e: "A \ (B - A) = {}" by blast have eq: "A \ (B - A) = B" using AB by blast from card_Un_disjoint[OF fA fBA e, unfolded eq c] have "card (B - A) = 0" by arith then have "B - A = {}" unfolding card_eq_0_iff using fA fB by simp with AB show "A = B" by blast qed lemma insert_partition: "x \ F \ \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ x \ \F = {}" by auto (* somewhat slow *) lemma finite_psubset_induct [consumes 1, case_names psubset]: assumes finite: "finite A" and major: "\A. finite A \ (\B. B \ A \ P B) \ P A" shows "P A" using finite proof (induct A taking: card rule: measure_induct_rule) case (less A) have fin: "finite A" by fact have ih: "card B < card A \ finite B \ P B" for B by fact have "P B" if "B \ A" for B proof - from that have "card B < card A" using psubset_card_mono fin by blast moreover from that have "B \ A" by auto then have "finite B" using fin finite_subset by blast ultimately show ?thesis using ih by simp qed with fin show "P A" using major by blast qed lemma finite_induct_select [consumes 1, case_names empty select]: assumes "finite S" and "P {}" and select: "\T. T \ S \ P T \ \s\S - T. P (insert s T)" shows "P S" proof - have "0 \ card S" by simp then have "\T \ S. card T = card S \ P T" proof (induct rule: dec_induct) case base with \P {}\ show ?case by (intro exI[of _ "{}"]) auto next case (step n) then obtain T where T: "T \ S" "card T = n" "P T" by auto with \n < card S\ have "T \ S" "P T" by auto with select[of T] obtain s where "s \ S" "s \ T" "P (insert s T)" by auto with step(2) T \finite S\ show ?case by (intro exI[of _ "insert s T"]) (auto dest: finite_subset) qed with \finite S\ show "P S" by (auto dest: card_subset_eq) qed lemma remove_induct [case_names empty infinite remove]: assumes empty: "P ({} :: 'a set)" and infinite: "\ finite B \ P B" and remove: "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" shows "P B" proof (cases "finite B") case False then show ?thesis by (rule infinite) next case True define A where "A = B" with True have "finite A" "A \ B" by simp_all then show "P A" proof (induct "card A" arbitrary: A) case 0 then have "A = {}" by auto with empty show ?case by simp next case (Suc n A) from \A \ B\ and \finite B\ have "finite A" by (rule finite_subset) moreover from Suc.hyps have "A \ {}" by auto moreover note \A \ B\ moreover have "P (A - {x})" if x: "x \ A" for x using x Suc.prems \Suc n = card A\ by (intro Suc) auto ultimately show ?case by (rule remove) qed qed lemma finite_remove_induct [consumes 1, case_names empty remove]: fixes P :: "'a set \ bool" assumes "finite B" and "P {}" and "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" defines "B' \ B" shows "P B'" by (induct B' rule: remove_induct) (simp_all add: assms) text \Main cardinality theorem.\ lemma card_partition [rule_format]: "finite C \ finite (\C) \ (\c\C. card c = k) \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}) \ k * card C = card (\C)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: card_Un_disjoint insert_partition finite_subset [of _ "\(insert _ _)"]) qed lemma card_eq_UNIV_imp_eq_UNIV: assumes fin: "finite (UNIV :: 'a set)" and card: "card A = card (UNIV :: 'a set)" shows "A = (UNIV :: 'a set)" proof show "A \ UNIV" by simp show "UNIV \ A" proof show "x \ A" for x proof (rule ccontr) assume "x \ A" then have "A \ UNIV" by auto with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono) with card show False by simp qed qed qed text \The form of a finite set of given cardinality\ lemma card_eq_SucD: assumes "card A = Suc k" shows "\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {})" proof - have fin: "finite A" using assms by (auto intro: ccontr) moreover have "card A \ 0" using assms by auto ultimately obtain b where b: "b \ A" by auto show ?thesis proof (intro exI conjI) show "A = insert b (A - {b})" using b by blast show "b \ A - {b}" by blast show "card (A - {b}) = k" and "k = 0 \ A - {b} = {}" using assms b fin by (fastforce dest: mk_disjoint_insert)+ qed qed lemma card_Suc_eq: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {}))" by (auto simp: card_insert_if card_gt_0_iff elim!: card_eq_SucD) lemma card_1_singletonE: assumes "card A = 1" obtains x where "A = {x}" using assms by (auto simp: card_Suc_eq) lemma is_singleton_altdef: "is_singleton A \ card A = 1" unfolding is_singleton_def by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def) lemma card_1_singleton_iff: "card A = Suc 0 \ (\x. A = {x})" by (simp add: card_Suc_eq) lemma card_le_Suc0_iff_eq: assumes "finite A" shows "card A \ Suc 0 \ (\a1 \ A. \a2 \ A. a1 = a2)" (is "?C = ?A") proof assume ?C thus ?A using assms by (auto simp: le_Suc_eq dest: card_eq_SucD) next assume ?A show ?C proof cases assume "A = {}" thus ?C using \?A\ by simp next assume "A \ {}" then obtain a where "A = {a}" using \?A\ by blast thus ?C by simp qed qed lemma card_le_Suc_iff: "Suc n \ card A = (\a B. A = insert a B \ a \ B \ n \ card B \ finite B)" proof (cases "finite A") case True then show ?thesis by (fastforce simp: card_Suc_eq less_eq_nat.simps split: nat.splits) qed auto lemma finite_fun_UNIVD2: assumes fin: "finite (UNIV :: ('a \ 'b) set)" shows "finite (UNIV :: 'b set)" proof - from fin have "finite (range (\f :: 'a \ 'b. f arbitrary))" for arbitrary by (rule finite_imageI) moreover have "UNIV = range (\f :: 'a \ 'b. f arbitrary)" for arbitrary by (rule UNIV_eq_I) auto ultimately show "finite (UNIV :: 'b set)" by simp qed lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1" unfolding UNIV_unit by simp lemma infinite_arbitrarily_large: assumes "\ finite A" shows "\B. finite B \ card B = n \ B \ A" proof (induction n) case 0 show ?case by (intro exI[of _ "{}"]) auto next case (Suc n) then obtain B where B: "finite B \ card B = n \ B \ A" .. with \\ finite A\ have "A \ B" by auto with B have "B \ A" by auto then have "\x. x \ A - B" by (elim psubset_imp_ex_mem) then obtain x where x: "x \ A - B" .. with B have "finite (insert x B) \ card (insert x B) = Suc n \ insert x B \ A" by auto then show "\B. finite B \ card B = Suc n \ B \ A" .. qed text \Sometimes, to prove that a set is finite, it is convenient to work with finite subsets and to show that their cardinalities are uniformly bounded. This possibility is formalized in the next criterion.\ lemma finite_if_finite_subsets_card_bdd: assumes "\G. G \ F \ finite G \ card G \ C" shows "finite F \ card F \ C" proof (cases "finite F") case False obtain n::nat where n: "n > max C 0" by auto obtain G where G: "G \ F" "card G = n" using infinite_arbitrarily_large[OF False] by auto hence "finite G" using \n > max C 0\ using card.infinite gr_implies_not0 by blast hence False using assms G n not_less by auto thus ?thesis .. next case True thus ?thesis using assms[of F] by auto qed subsubsection \Cardinality of image\ lemma card_image_le: "finite A \ card (f ` A) \ card A" by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if) lemma card_image: "inj_on f A \ card (f ` A) = card A" proof (induct A rule: infinite_finite_induct) case (infinite A) then have "\ finite (f ` A)" by (auto dest: finite_imageD) with infinite show ?case by simp qed simp_all lemma bij_betw_same_card: "bij_betw f A B \ card A = card B" by (auto simp: card_image bij_betw_def) lemma endo_inj_surj: "finite A \ f ` A \ A \ inj_on f A \ f ` A = A" by (simp add: card_seteq card_image) lemma eq_card_imp_inj_on: assumes "finite A" "card(f ` A) = card A" shows "inj_on f A" using assms proof (induct rule:finite_induct) case empty show ?case by simp next case (insert x A) then show ?case using card_image_le [of A f] by (simp add: card_insert_if split: if_splits) qed lemma inj_on_iff_eq_card: "finite A \ inj_on f A \ card (f ` A) = card A" by (blast intro: card_image eq_card_imp_inj_on) lemma card_inj_on_le: assumes "inj_on f A" "f ` A \ B" "finite B" shows "card A \ card B" proof - have "finite A" using assms by (blast intro: finite_imageD dest: finite_subset) then show ?thesis using assms by (force intro: card_mono simp: card_image [symmetric]) qed lemma inj_on_iff_card_le: "\ finite A; finite B \ \ (\f. inj_on f A \ f ` A \ B) = (card A \ card B)" using card_inj_on_le[of _ A B] card_le_inj[of A B] by blast lemma surj_card_le: "finite A \ B \ f ` A \ card B \ card A" by (blast intro: card_image_le card_mono le_trans) lemma card_bij_eq: "inj_on f A \ f ` A \ B \ inj_on g B \ g ` B \ A \ finite A \ finite B \ card A = card B" by (auto intro: le_antisym card_inj_on_le) lemma bij_betw_finite: "bij_betw f A B \ finite A \ finite B" unfolding bij_betw_def using finite_imageD [of f A] by auto lemma inj_on_finite: "inj_on f A \ f ` A \ B \ finite B \ finite A" using finite_imageD finite_subset by blast lemma card_vimage_inj: "inj f \ A \ range f \ card (f -` A) = card A" by (auto 4 3 simp: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on]) subsubsection \Pigeonhole Principles\ lemma pigeonhole: "card A > card (f ` A) \ \ inj_on f A " by (auto dest: card_image less_irrefl_nat) lemma pigeonhole_infinite: assumes "\ finite A" and "finite (f`A)" shows "\a0\A. \ finite {a\A. f a = f a0}" using assms(2,1) proof (induct "f`A" arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert b F) show ?case proof (cases "finite {a\A. f a = b}") case True with \\ finite A\ have "\ finite (A - {a\A. f a = b})" by simp also have "A - {a\A. f a = b} = {a\A. f a \ b}" by blast finally have "\ finite {a\A. f a \ b}" . from insert(3)[OF _ this] insert(2,4) show ?thesis by simp (blast intro: rev_finite_subset) next case False then have "{a \ A. f a = b} \ {}" by force with False show ?thesis by blast qed qed lemma pigeonhole_infinite_rel: assumes "\ finite A" and "finite B" and "\a\A. \b\B. R a b" shows "\b\B. \ finite {a:A. R a b}" proof - let ?F = "\a. {b\B. R a b}" from finite_Pow_iff[THEN iffD2, OF \finite B\] have "finite (?F ` A)" by (blast intro: rev_finite_subset) from pigeonhole_infinite [where f = ?F, OF assms(1) this] obtain a0 where "a0 \ A" and infinite: "\ finite {a\A. ?F a = ?F a0}" .. obtain b0 where "b0 \ B" and "R a0 b0" using \a0 \ A\ assms(3) by blast have "finite {a\A. ?F a = ?F a0}" if "finite {a\A. R a b0}" using \b0 \ B\ \R a0 b0\ that by (blast intro: rev_finite_subset) with infinite \b0 \ B\ show ?thesis by blast qed subsubsection \Cardinality of sums\ lemma card_Plus: assumes "finite A" "finite B" shows "card (A <+> B) = card A + card B" proof - have "Inl`A \ Inr`B = {}" by fast with assms show ?thesis by (simp add: Plus_def card_Un_disjoint card_image) qed lemma card_Plus_conv_if: "card (A <+> B) = (if finite A \ finite B then card A + card B else 0)" by (auto simp add: card_Plus) text \Relates to equivalence classes. Based on a theorem of F. Kammüller.\ lemma dvd_partition: assumes f: "finite (\C)" and "\c\C. k dvd card c" "\c1\C. \c2\C. c1 \ c2 \ c1 \ c2 = {}" shows "k dvd card (\C)" proof - have "finite C" by (rule finite_UnionD [OF f]) then show ?thesis using assms proof (induct rule: finite_induct) case empty show ?case by simp next case (insert c C) then have "c \ \C = {}" by auto with insert show ?case by (simp add: card_Un_disjoint) qed qed subsubsection \Finite orders\ context order begin lemma finite_has_maximal: "\ finite A; A \ {} \ \ \ m \ A. \ b \ A. m \ b \ m = b" proof (induction rule: finite_psubset_induct) case (psubset A) from \A \ {}\ obtain a where "a \ A" by auto let ?B = "{b \ A. a < b}" show ?case proof cases assume "?B = {}" hence "\ b \ A. a \ b \ a = b" using le_neq_trans by blast thus ?thesis using \a \ A\ by blast next assume "?B \ {}" have "a \ ?B" by auto hence "?B \ A" using \a \ A\ by blast from psubset.IH[OF this \?B \ {}\] show ?thesis using order.strict_trans2 by blast qed qed lemma finite_has_maximal2: "\ finite A; a \ A \ \ \ m \ A. a \ m \ (\ b \ A. m \ b \ m = b)" using finite_has_maximal[of "{b \ A. a \ b}"] by fastforce lemma finite_has_minimal: "\ finite A; A \ {} \ \ \ m \ A. \ b \ A. b \ m \ m = b" proof (induction rule: finite_psubset_induct) case (psubset A) from \A \ {}\ obtain a where "a \ A" by auto let ?B = "{b \ A. b < a}" show ?case proof cases assume "?B = {}" hence "\ b \ A. b \ a \ a = b" using le_neq_trans by blast thus ?thesis using \a \ A\ by blast next assume "?B \ {}" have "a \ ?B" by auto hence "?B \ A" using \a \ A\ by blast from psubset.IH[OF this \?B \ {}\] show ?thesis using order.strict_trans1 by blast qed qed lemma finite_has_minimal2: "\ finite A; a \ A \ \ \ m \ A. m \ a \ (\ b \ A. b \ m \ m = b)" using finite_has_minimal[of "{b \ A. b \ a}"] by fastforce end subsubsection \Relating injectivity and surjectivity\ lemma finite_surj_inj: assumes "finite A" "A \ f ` A" shows "inj_on f A" proof - have "f ` A = A" by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le) then show ?thesis using assms by (simp add: eq_card_imp_inj_on) qed lemma finite_UNIV_surj_inj: "finite(UNIV:: 'a set) \ surj f \ inj f" for f :: "'a \ 'a" by (blast intro: finite_surj_inj subset_UNIV) lemma finite_UNIV_inj_surj: "finite(UNIV:: 'a set) \ inj f \ surj f" for f :: "'a \ 'a" by (fastforce simp:surj_def dest!: endo_inj_surj) lemma surjective_iff_injective_gen: assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" and ST: "f ` S \ T" shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") proof assume h: "?lhs" { fix x y assume x: "x \ S" assume y: "y \ S" assume f: "f x = f y" from x fS have S0: "card S \ 0" by auto have "x = y" proof (rule ccontr) assume xy: "\ ?thesis" have th: "card S \ card (f ` (S - {y}))" unfolding c proof (rule card_mono) show "finite (f ` (S - {y}))" by (simp add: fS) have "\x \ y; x \ S; z \ S; f x = f y\ \ \x \ S. x \ y \ f z = f x" for z by (case_tac "z = y \ z = x") auto then show "T \ f ` (S - {y})" using h xy x y f by fastforce qed also have " \ \ card (S - {y})" by (simp add: card_image_le fS) also have "\ \ card S - 1" using y fS by simp finally show False using S0 by arith qed } then show ?rhs unfolding inj_on_def by blast next assume h: ?rhs have "f ` S = T" by (simp add: ST c card_image card_subset_eq fT h) then show ?lhs by blast qed hide_const (open) Finite_Set.fold subsection \Infinite Sets\ text \ Some elementary facts about infinite sets, mostly by Stephan Merz. Beware! Because "infinite" merely abbreviates a negation, these lemmas may not work well with \blast\. \ abbreviation infinite :: "'a set \ bool" where "infinite S \ \ finite S" text \ Infinite sets are non-empty, and if we remove some elements from an infinite set, the result is still infinite. \ lemma infinite_UNIV_nat [iff]: "infinite (UNIV :: nat set)" proof assume "finite (UNIV :: nat set)" with finite_UNIV_inj_surj [of Suc] show False by simp (blast dest: Suc_neq_Zero surjD) qed lemma infinite_UNIV_char_0: "infinite (UNIV :: 'a::semiring_char_0 set)" proof assume "finite (UNIV :: 'a set)" with subset_UNIV have "finite (range of_nat :: 'a set)" by (rule finite_subset) moreover have "inj (of_nat :: nat \ 'a)" by (simp add: inj_on_def) ultimately have "finite (UNIV :: nat set)" by (rule finite_imageD) then show False by simp qed lemma infinite_imp_nonempty: "infinite S \ S \ {}" by auto lemma infinite_remove: "infinite S \ infinite (S - {a})" by simp lemma Diff_infinite_finite: assumes "finite T" "infinite S" shows "infinite (S - T)" using \finite T\ proof induct from \infinite S\ show "infinite (S - {})" by auto next fix T x assume ih: "infinite (S - T)" have "S - (insert x T) = (S - T) - {x}" by (rule Diff_insert) with ih show "infinite (S - (insert x T))" by (simp add: infinite_remove) qed lemma Un_infinite: "infinite S \ infinite (S \ T)" by simp lemma infinite_Un: "infinite (S \ T) \ infinite S \ infinite T" by simp lemma infinite_super: assumes "S \ T" and "infinite S" shows "infinite T" proof assume "finite T" with \S \ T\ have "finite S" by (simp add: finite_subset) with \infinite S\ show False by simp qed proposition infinite_coinduct [consumes 1, case_names infinite]: assumes "X A" and step: "\A. X A \ \x\A. X (A - {x}) \ infinite (A - {x})" shows "infinite A" proof assume "finite A" then show False using \X A\ proof (induction rule: finite_psubset_induct) case (psubset A) then obtain x where "x \ A" "X (A - {x}) \ infinite (A - {x})" using local.step psubset.prems by blast then have "X (A - {x})" using psubset.hyps by blast show False proof (rule psubset.IH [where B = "A - {x}"]) show "A - {x} \ A" using \x \ A\ by blast qed fact qed qed text \ For any function with infinite domain and finite range there is some element that is the image of infinitely many domain elements. In particular, any infinite sequence of elements from a finite set contains some element that occurs infinitely often. \ lemma inf_img_fin_dom': assumes img: "finite (f ` A)" and dom: "infinite A" shows "\y \ f ` A. infinite (f -` {y} \ A)" proof (rule ccontr) have "A \ (\y\f ` A. f -` {y} \ A)" by auto moreover assume "\ ?thesis" with img have "finite (\y\f ` A. f -` {y} \ A)" by blast ultimately have "finite A" by (rule finite_subset) with dom show False by contradiction qed lemma inf_img_fin_domE': assumes "finite (f ` A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y} \ A)" using assms by (blast dest: inf_img_fin_dom') lemma inf_img_fin_dom: assumes img: "finite (f`A)" and dom: "infinite A" shows "\y \ f`A. infinite (f -` {y})" using inf_img_fin_dom'[OF assms] by auto lemma inf_img_fin_domE: assumes "finite (f`A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) proposition finite_image_absD: "finite (abs ` S) \ finite S" for S :: "'a::linordered_ring set" by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) + subsection \The finite powerset operator\ definition Fpow :: "'a set \ 'a set set" where "Fpow A \ {X. X \ A \ finite X}" lemma Fpow_mono: "A \ B \ Fpow A \ Fpow B" unfolding Fpow_def by auto lemma empty_in_Fpow: "{} \ Fpow A" unfolding Fpow_def by auto lemma Fpow_not_empty: "Fpow A \ {}" using empty_in_Fpow by blast lemma Fpow_subset_Pow: "Fpow A \ Pow A" unfolding Fpow_def by auto lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}" unfolding Fpow_def Pow_def by blast lemma inj_on_image_Fpow: assumes "inj_on f A" shows "inj_on (image f) (Fpow A)" using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"] inj_on_image_Pow by blast lemma image_Fpow_mono: assumes "f ` A \ B" shows "(image f) ` (Fpow A) \ Fpow B" using assms by(unfold Fpow_def, auto) end diff --git a/src/HOL/Hilbert_Choice.thy b/src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy +++ b/src/HOL/Hilbert_Choice.thy @@ -1,1231 +1,1275 @@ (* Title: HOL/Hilbert_Choice.thy Author: Lawrence C Paulson, Tobias Nipkow Author: Viorel Preoteasa (Results about complete distributive lattices) Copyright 2001 University of Cambridge *) section \Hilbert's Epsilon-Operator and the Axiom of Choice\ theory Hilbert_Choice imports Wellfounded keywords "specification" :: thy_goal_defn begin subsection \Hilbert's epsilon\ axiomatization Eps :: "('a \ bool) \ 'a" where someI: "P x \ P (Eps P)" syntax (epsilon) "_Eps" :: "pttrn \ bool \ 'a" ("(3\_./ _)" [0, 10] 10) syntax (input) "_Eps" :: "pttrn \ bool \ 'a" ("(3@ _./ _)" [0, 10] 10) syntax "_Eps" :: "pttrn \ bool \ 'a" ("(3SOME _./ _)" [0, 10] 10) translations "SOME x. P" \ "CONST Eps (\x. P)" print_translation \ [(\<^const_syntax>\Eps\, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_Eps\ $ x $ t end)] \ \ \to avoid eta-contraction of body\ definition inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" where "inv_into A f = (\x. SOME y. y \ A \ f y = x)" lemma inv_into_def2: "inv_into A f x = (SOME y. y \ A \ f y = x)" by(simp add: inv_into_def) abbreviation inv :: "('a \ 'b) \ ('b \ 'a)" where "inv \ inv_into UNIV" subsection \Hilbert's Epsilon-operator\ lemma Eps_cong: assumes "\x. P x = Q x" shows "Eps P = Eps Q" using ext[of P Q, OF assms] by simp text \ Easier to use than \someI\ if the witness comes from an existential formula. \ lemma someI_ex [elim?]: "\x. P x \ P (SOME x. P x)" by (elim exE someI) lemma some_eq_imp: assumes "Eps P = a" "P b" shows "P a" using assms someI_ex by force text \ Easier to use than \someI\ because the conclusion has only one occurrence of \<^term>\P\. \ lemma someI2: "P a \ (\x. P x \ Q x) \ Q (SOME x. P x)" by (blast intro: someI) text \ Easier to use than \someI2\ if the witness comes from an existential formula. \ lemma someI2_ex: "\a. P a \ (\x. P x \ Q x) \ Q (SOME x. P x)" by (blast intro: someI2) lemma someI2_bex: "\a\A. P a \ (\x. x \ A \ P x \ Q x) \ Q (SOME x. x \ A \ P x)" by (blast intro: someI2) lemma some_equality [intro]: "P a \ (\x. P x \ x = a) \ (SOME x. P x) = a" by (blast intro: someI2) lemma some1_equality: "\!x. P x \ P a \ (SOME x. P x) = a" by blast lemma some_eq_ex: "P (SOME x. P x) \ (\x. P x)" by (blast intro: someI) lemma some_in_eq: "(SOME x. x \ A) \ A \ A \ {}" unfolding ex_in_conv[symmetric] by (rule some_eq_ex) lemma some_eq_trivial [simp]: "(SOME y. y = x) = x" by (rule some_equality) (rule refl) lemma some_sym_eq_trivial [simp]: "(SOME y. x = y) = x" by (iprover intro: some_equality) subsection \Axiom of Choice, Proved Using the Description Operator\ lemma choice: "\x. \y. Q x y \ \f. \x. Q x (f x)" by (fast elim: someI) lemma bchoice: "\x\S. \y. Q x y \ \f. \x\S. Q x (f x)" by (fast elim: someI) lemma choice_iff: "(\x. \y. Q x y) \ (\f. \x. Q x (f x))" by (fast elim: someI) lemma choice_iff': "(\x. P x \ (\y. Q x y)) \ (\f. \x. P x \ Q x (f x))" by (fast elim: someI) lemma bchoice_iff: "(\x\S. \y. Q x y) \ (\f. \x\S. Q x (f x))" by (fast elim: someI) lemma bchoice_iff': "(\x\S. P x \ (\y. Q x y)) \ (\f. \x\S. P x \ Q x (f x))" by (fast elim: someI) lemma dependent_nat_choice: assumes 1: "\x. P 0 x" and 2: "\x n. P n x \ \y. P (Suc n) y \ Q n x y" shows "\f. \n. P n (f n) \ Q n (f n) (f (Suc n))" proof (intro exI allI conjI) fix n define f where "f = rec_nat (SOME x. P 0 x) (\n x. SOME y. P (Suc n) y \ Q n x y)" then have "P 0 (f 0)" "\n. P n (f n) \ P (Suc n) (f (Suc n)) \ Q n (f n) (f (Suc n))" using someI_ex[OF 1] someI_ex[OF 2] by simp_all then show "P n (f n)" "Q n (f n) (f (Suc n))" by (induct n) auto qed lemma finite_subset_Union: assumes "finite A" "A \ \\" obtains \ where "finite \" "\ \ \" "A \ \\" proof - have "\x\A. \B\\. x\B" using assms by blast then obtain f where f: "\x. x \ A \ f x \ \ \ x \ f x" by (auto simp add: bchoice_iff Bex_def) show thesis proof show "finite (f ` A)" using assms by auto qed (use f in auto) qed subsection \Function Inverse\ lemma inv_def: "inv f = (\y. SOME x. f x = y)" by (simp add: inv_into_def) lemma inv_into_into: "x \ f ` A \ inv_into A f x \ A" by (simp add: inv_into_def) (fast intro: someI2) lemma inv_identity [simp]: "inv (\a. a) = (\a. a)" by (simp add: inv_def) lemma inv_id [simp]: "inv id = id" by (simp add: id_def) lemma inv_into_f_f [simp]: "inj_on f A \ x \ A \ inv_into A f (f x) = x" by (simp add: inv_into_def inj_on_def) (blast intro: someI2) lemma inv_f_f: "inj f \ inv f (f x) = x" by simp lemma f_inv_into_f: "y \ f`A \ f (inv_into A f y) = y" by (simp add: inv_into_def) (fast intro: someI2) lemma inv_into_f_eq: "inj_on f A \ x \ A \ f x = y \ inv_into A f y = x" by (erule subst) (fast intro: inv_into_f_f) lemma inv_f_eq: "inj f \ f x = y \ inv f y = x" by (simp add:inv_into_f_eq) lemma inj_imp_inv_eq: "inj f \ \x. f (g x) = x \ inv f = g" by (blast intro: inv_into_f_eq) text \But is it useful?\ lemma inj_transfer: assumes inj: "inj f" and minor: "\y. y \ range f \ P (inv f y)" shows "P x" proof - have "f x \ range f" by auto then have "P(inv f (f x))" by (rule minor) then show "P x" by (simp add: inv_into_f_f [OF inj]) qed lemma inj_iff: "inj f \ inv f \ f = id" by (simp add: o_def fun_eq_iff) (blast intro: inj_on_inverseI inv_into_f_f) lemma inv_o_cancel[simp]: "inj f \ inv f \ f = id" by (simp add: inj_iff) lemma o_inv_o_cancel[simp]: "inj f \ g \ inv f \ f = g" by (simp add: comp_assoc) lemma inv_into_image_cancel[simp]: "inj_on f A \ S \ A \ inv_into A f ` f ` S = S" by (fastforce simp: image_def) lemma inj_imp_surj_inv: "inj f \ surj (inv f)" by (blast intro!: surjI inv_into_f_f) lemma surj_f_inv_f: "surj f \ f (inv f y) = y" by (simp add: f_inv_into_f) lemma bij_inv_eq_iff: "bij p \ x = inv p y \ p x = y" using surj_f_inv_f[of p] by (auto simp add: bij_def) lemma inv_into_injective: assumes eq: "inv_into A f x = inv_into A f y" and x: "x \ f`A" and y: "y \ f`A" shows "x = y" proof - from eq have "f (inv_into A f x) = f (inv_into A f y)" by simp with x y show ?thesis by (simp add: f_inv_into_f) qed lemma inj_on_inv_into: "B \ f`A \ inj_on (inv_into A f) B" by (blast intro: inj_onI dest: inv_into_injective injD) lemma inj_imp_bij_betw_inv: "inj f \ bij_betw (inv f) (f ` M) M" by (simp add: bij_betw_def image_subsetI inj_on_inv_into) lemma bij_betw_inv_into: "bij_betw f A B \ bij_betw (inv_into A f) B A" by (auto simp add: bij_betw_def inj_on_inv_into) lemma surj_imp_inj_inv: "surj f \ inj (inv f)" by (simp add: inj_on_inv_into) lemma surj_iff: "surj f \ f \ inv f = id" by (auto intro!: surjI simp: surj_f_inv_f fun_eq_iff[where 'b='a]) lemma surj_iff_all: "surj f \ (\x. f (inv f x) = x)" by (simp add: o_def surj_iff fun_eq_iff) lemma surj_imp_inv_eq: assumes "surj f" and gf: "\x. g (f x) = x" shows "inv f = g" proof (rule ext) fix x have "g (f (inv f x)) = inv f x" by (rule gf) then show "inv f x = g x" by (simp add: surj_f_inv_f \surj f\) qed lemma bij_imp_bij_inv: "bij f \ bij (inv f)" by (simp add: bij_def inj_imp_surj_inv surj_imp_inj_inv) lemma inv_equality: "(\x. g (f x) = x) \ (\y. f (g y) = y) \ inv f = g" by (rule ext) (auto simp add: inv_into_def) lemma inv_inv_eq: "bij f \ inv (inv f) = f" by (rule inv_equality) (auto simp add: bij_def surj_f_inv_f) text \ \bij (inv f)\ implies little about \f\. Consider \f :: bool \ bool\ such that \f True = f False = True\. Then it ia consistent with axiom \someI\ that \inv f\ could be any function at all, including the identity function. If \inv f = id\ then \inv f\ is a bijection, but \inj f\, \surj f\ and \inv (inv f) = f\ all fail. \ lemma inv_into_comp: "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ inv_into A (f \ g) x = (inv_into A g \ inv_into (g ` A) f) x" by (auto simp: f_inv_into_f inv_into_into intro: inv_into_f_eq comp_inj_on) lemma o_inv_distrib: "bij f \ bij g \ inv (f \ g) = inv g \ inv f" by (rule inv_equality) (auto simp add: bij_def surj_f_inv_f) lemma image_f_inv_f: "surj f \ f ` (inv f ` A) = A" by (simp add: surj_f_inv_f image_comp comp_def) lemma image_inv_f_f: "inj f \ inv f ` (f ` A) = A" by simp lemma bij_image_Collect_eq: assumes "bij f" shows "f ` Collect P = {y. P (inv f y)}" proof show "f ` Collect P \ {y. P (inv f y)}" using assms by (force simp add: bij_is_inj) show "{y. P (inv f y)} \ f ` Collect P" using assms by (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric]) qed lemma bij_vimage_eq_inv_image: assumes "bij f" shows "f -` A = inv f ` A" proof show "f -` A \ inv f ` A" using assms by (blast intro: bij_is_inj [THEN inv_into_f_f, symmetric]) show "inv f ` A \ f -` A" using assms by (auto simp add: bij_is_surj [THEN surj_f_inv_f]) qed lemma inv_fn_o_fn_is_id: fixes f::"'a \ 'a" assumes "bij f" shows "((inv f)^^n) o (f^^n) = (\x. x)" proof - have "((inv f)^^n)((f^^n) x) = x" for x n proof (induction n) case (Suc n) have *: "(inv f) (f y) = y" for y by (simp add: assms bij_is_inj) have "(inv f ^^ Suc n) ((f ^^ Suc n) x) = (inv f^^n) (inv f (f ((f^^n) x)))" by (simp add: funpow_swap1) also have "... = (inv f^^n) ((f^^n) x)" using * by auto also have "... = x" using Suc.IH by auto finally show ?case by simp qed (auto) then show ?thesis unfolding o_def by blast qed lemma fn_o_inv_fn_is_id: fixes f::"'a \ 'a" assumes "bij f" shows "(f^^n) o ((inv f)^^n) = (\x. x)" proof - have "(f^^n) (((inv f)^^n) x) = x" for x n proof (induction n) case (Suc n) have *: "f(inv f y) = y" for y using bij_inv_eq_iff[OF assms] by auto have "(f ^^ Suc n) ((inv f ^^ Suc n) x) = (f^^n) (f (inv f ((inv f^^n) x)))" by (simp add: funpow_swap1) also have "... = (f^^n) ((inv f^^n) x)" using * by auto also have "... = x" using Suc.IH by auto finally show ?case by simp qed (auto) then show ?thesis unfolding o_def by blast qed lemma inv_fn: fixes f::"'a \ 'a" assumes "bij f" shows "inv (f^^n) = ((inv f)^^n)" proof - have "inv (f^^n) x = ((inv f)^^n) x" for x proof (rule inv_into_f_eq) show "inj (f ^^ n)" by (simp add: inj_fn[OF bij_is_inj [OF assms]]) show "(f ^^ n) ((inv f ^^ n) x) = x" using fn_o_inv_fn_is_id[OF assms, THEN fun_cong] by force qed auto then show ?thesis by auto qed +lemma funpow_inj_finite: \<^marker>\contributor \Lars Noschinski\\ + assumes \inj p\ \finite {y. \n. y = (p ^^ n) x}\ + obtains n where \n > 0\ \(p ^^ n) x = x\ +proof - + have \infinite (UNIV :: nat set)\ + by simp + moreover have \{y. \n. y = (p ^^ n) x} = (\n. (p ^^ n) x) ` UNIV\ + by auto + with assms have \finite \\ + by simp + ultimately have "\n \ UNIV. \ finite {m \ UNIV. (p ^^ m) x = (p ^^ n) x}" + by (rule pigeonhole_infinite) + then obtain n where "infinite {m. (p ^^ m) x = (p ^^ n) x}" by auto + then have "infinite ({m. (p ^^ m) x = (p ^^ n) x} - {n})" by auto + then have "({m. (p ^^ m) x = (p ^^ n) x} - {n}) \ {}" + by (auto simp add: subset_singleton_iff) + then obtain m where m: "(p ^^ m) x = (p ^^ n) x" "m \ n" by auto + + { fix m n assume "(p ^^ n) x = (p ^^ m) x" "m < n" + have "(p ^^ (n - m)) x = inv (p ^^ m) ((p ^^ m) ((p ^^ (n - m)) x))" + using \inj p\ by (simp add: inv_f_f) + also have "((p ^^ m) ((p ^^ (n - m)) x)) = (p ^^ n) x" + using \m < n\ funpow_add [of m \n - m\ p] by simp + also have "inv (p ^^ m) \ = x" + using \inj p\ by (simp add: \(p ^^ n) x = _\) + finally have "(p ^^ (n - m)) x = x" "0 < n - m" + using \m < n\ by auto } + note general = this + + show thesis + proof (cases m n rule: linorder_cases) + case less + then have \n - m > 0\ \(p ^^ (n - m)) x = x\ + using general [of n m] m by simp_all + then show thesis by (blast intro: that) + next + case equal + then show thesis using m by simp + next + case greater + then have \m - n > 0\ \(p ^^ (m - n)) x = x\ + using general [of m n] m by simp_all + then show thesis by (blast intro: that) + qed +qed + lemma mono_inv: fixes f::"'a::linorder \ 'b::linorder" assumes "mono f" "bij f" shows "mono (inv f)" proof fix x y::'b assume "x \ y" from \bij f\ obtain a b where x: "x = f a" and y: "y = f b" by(fastforce simp: bij_def surj_def) show "inv f x \ inv f y" proof (rule le_cases) assume "a \ b" thus ?thesis using \bij f\ x y by(simp add: bij_def inv_f_f) next assume "b \ a" hence "f b \ f a" by(rule monoD[OF \mono f\]) hence "y \ x" using x y by simp hence "x = y" using \x \ y\ by auto thus ?thesis by simp qed qed lemma strict_mono_inv_on_range: fixes f :: "'a::linorder \ 'b::order" assumes "strict_mono f" shows "strict_mono_on (inv f) (range f)" proof (clarsimp simp: strict_mono_on_def) fix x y assume "f x < f y" then show "inv f (f x) < inv f (f y)" using assms strict_mono_imp_inj_on strict_mono_less by fastforce qed lemma mono_bij_Inf: fixes f :: "'a::complete_linorder \ 'b::complete_linorder" assumes "mono f" "bij f" shows "f (Inf A) = Inf (f`A)" proof - have "surj f" using \bij f\ by (auto simp: bij_betw_def) have *: "(inv f) (Inf (f`A)) \ Inf ((inv f)`(f`A))" using mono_Inf[OF mono_inv[OF assms], of "f`A"] by simp have "Inf (f`A) \ f (Inf ((inv f)`(f`A)))" using monoD[OF \mono f\ *] by(simp add: surj_f_inv_f[OF \surj f\]) also have "... = f(Inf A)" using assms by (simp add: bij_is_inj) finally show ?thesis using mono_Inf[OF assms(1), of A] by auto qed lemma finite_fun_UNIVD1: assumes fin: "finite (UNIV :: ('a \ 'b) set)" and card: "card (UNIV :: 'b set) \ Suc 0" shows "finite (UNIV :: 'a set)" proof - let ?UNIV_b = "UNIV :: 'b set" from fin have "finite ?UNIV_b" by (rule finite_fun_UNIVD2) with card have "card ?UNIV_b \ Suc (Suc 0)" by (cases "card ?UNIV_b") (auto simp: card_eq_0_iff) then have "card ?UNIV_b = Suc (Suc (card ?UNIV_b - Suc (Suc 0)))" by simp then obtain b1 b2 :: 'b where b1b2: "b1 \ b2" by (auto simp: card_Suc_eq) from fin have fin': "finite (range (\f :: 'a \ 'b. inv f b1))" by (rule finite_imageI) have "UNIV = range (\f :: 'a \ 'b. inv f b1)" proof (rule UNIV_eq_I) fix x :: 'a from b1b2 have "x = inv (\y. if y = x then b1 else b2) b1" by (simp add: inv_into_def) then show "x \ range (\f::'a \ 'b. inv f b1)" by blast qed with fin' show ?thesis by simp qed text \ Every infinite set contains a countable subset. More precisely we show that a set \S\ is infinite if and only if there exists an injective function from the naturals into \S\. The ``only if'' direction is harder because it requires the construction of a sequence of pairwise different elements of an infinite set \S\. The idea is to construct a sequence of non-empty and infinite subsets of \S\ obtained by successively removing elements of \S\. \ lemma infinite_countable_subset: assumes inf: "\ finite S" shows "\f::nat \ 'a. inj f \ range f \ S" \ \Courtesy of Stephan Merz\ proof - define Sseq where "Sseq = rec_nat S (\n T. T - {SOME e. e \ T})" define pick where "pick n = (SOME e. e \ Sseq n)" for n have *: "Sseq n \ S" "\ finite (Sseq n)" for n by (induct n) (auto simp: Sseq_def inf) then have **: "\n. pick n \ Sseq n" unfolding pick_def by (subst (asm) finite.simps) (auto simp add: ex_in_conv intro: someI_ex) with * have "range pick \ S" by auto moreover have "pick n \ pick (n + Suc m)" for m n proof - have "pick n \ Sseq (n + Suc m)" by (induct m) (auto simp add: Sseq_def pick_def) with ** show ?thesis by auto qed then have "inj pick" by (intro linorder_injI) (auto simp add: less_iff_Suc_add) ultimately show ?thesis by blast qed lemma infinite_iff_countable_subset: "\ finite S \ (\f::nat \ 'a. inj f \ range f \ S)" \ \Courtesy of Stephan Merz\ using finite_imageD finite_subset infinite_UNIV_char_0 infinite_countable_subset by auto lemma image_inv_into_cancel: assumes surj: "f`A = A'" and sub: "B' \ A'" shows "f `((inv_into A f)`B') = B'" using assms proof (auto simp: f_inv_into_f) let ?f' = "inv_into A f" fix a' assume *: "a' \ B'" with sub have "a' \ A'" by auto with surj have "a' = f (?f' a')" by (auto simp: f_inv_into_f) with * show "a' \ f ` (?f' ` B')" by blast qed lemma inv_into_inv_into_eq: assumes "bij_betw f A A'" and a: "a \ A" shows "inv_into A' (inv_into A f) a = f a" proof - let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'" from assms have *: "bij_betw ?f' A' A" by (auto simp: bij_betw_inv_into) with a obtain a' where a': "a' \ A'" "?f' a' = a" unfolding bij_betw_def by force with a * have "?f'' a = a'" by (auto simp: f_inv_into_f bij_betw_def) moreover from assms a' have "f a = a'" by (auto simp: bij_betw_def) ultimately show "?f'' a = f a" by simp qed lemma inj_on_iff_surj: assumes "A \ {}" shows "(\f. inj_on f A \ f ` A \ A') \ (\g. g ` A' = A)" proof safe fix f assume inj: "inj_on f A" and incl: "f ` A \ A'" let ?phi = "\a' a. a \ A \ f a = a'" let ?csi = "\a. a \ A" let ?g = "\a'. if a' \ f ` A then (SOME a. ?phi a' a) else (SOME a. ?csi a)" have "?g ` A' = A" proof show "?g ` A' \ A" proof clarify fix a' assume *: "a' \ A'" show "?g a' \ A" proof (cases "a' \ f ` A") case True then obtain a where "?phi a' a" by blast then have "?phi a' (SOME a. ?phi a' a)" using someI[of "?phi a'" a] by blast with True show ?thesis by auto next case False with assms have "?csi (SOME a. ?csi a)" using someI_ex[of ?csi] by blast with False show ?thesis by auto qed qed next show "A \ ?g ` A'" proof - have "?g (f a) = a \ f a \ A'" if a: "a \ A" for a proof - let ?b = "SOME aa. ?phi (f a) aa" from a have "?phi (f a) a" by auto then have *: "?phi (f a) ?b" using someI[of "?phi(f a)" a] by blast then have "?g (f a) = ?b" using a by auto moreover from inj * a have "a = ?b" by (auto simp add: inj_on_def) ultimately have "?g(f a) = a" by simp with incl a show ?thesis by auto qed then show ?thesis by force qed qed then show "\g. g ` A' = A" by blast next fix g let ?f = "inv_into A' g" have "inj_on ?f (g ` A')" by (auto simp: inj_on_inv_into) moreover have "?f (g a') \ A'" if a': "a' \ A'" for a' proof - let ?phi = "\ b'. b' \ A' \ g b' = g a'" from a' have "?phi a'" by auto then have "?phi (SOME b'. ?phi b')" using someI[of ?phi] by blast then show ?thesis by (auto simp: inv_into_def) qed ultimately show "\f. inj_on f (g ` A') \ f ` g ` A' \ A'" by auto qed lemma Ex_inj_on_UNION_Sigma: "\f. (inj_on f (\i \ I. A i) \ f ` (\i \ I. A i) \ (SIGMA i : I. A i))" proof let ?phi = "\a i. i \ I \ a \ A i" let ?sm = "\a. SOME i. ?phi a i" let ?f = "\a. (?sm a, a)" have "inj_on ?f (\i \ I. A i)" by (auto simp: inj_on_def) moreover have "?sm a \ I \ a \ A(?sm a)" if "i \ I" and "a \ A i" for i a using that someI[of "?phi a" i] by auto then have "?f ` (\i \ I. A i) \ (SIGMA i : I. A i)" by auto ultimately show "inj_on ?f (\i \ I. A i) \ ?f ` (\i \ I. A i) \ (SIGMA i : I. A i)" by auto qed lemma inv_unique_comp: assumes fg: "f \ g = id" and gf: "g \ f = id" shows "inv f = g" using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff) lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" by (rule inv_unique_comp) simp_all lemma bij_swap_comp: assumes "bij p" shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) lemma subset_image_inj: "S \ f ` T \ (\U. U \ T \ inj_on f U \ S = f ` U)" proof safe show "\U\T. inj_on f U \ S = f ` U" if "S \ f ` T" proof - from that [unfolded subset_image_iff subset_iff] obtain g where g: "\x. x \ S \ g x \ T \ x = f (g x)" by (auto simp add: image_iff Bex_def choice_iff') show ?thesis proof (intro exI conjI) show "g ` S \ T" by (simp add: g image_subsetI) show "inj_on f (g ` S)" using g by (auto simp: inj_on_def) show "S = f ` (g ` S)" using g image_subset_iff by auto qed qed qed blast subsection \Other Consequences of Hilbert's Epsilon\ text \Hilbert's Epsilon and the \<^term>\split\ Operator\ text \Looping simprule!\ lemma split_paired_Eps: "(SOME x. P x) = (SOME (a, b). P (a, b))" by simp lemma Eps_case_prod: "Eps (case_prod P) = (SOME xy. P (fst xy) (snd xy))" by (simp add: split_def) lemma Eps_case_prod_eq [simp]: "(SOME (x', y'). x = x' \ y = y') = (x, y)" by blast text \A relation is wellfounded iff it has no infinite descending chain.\ lemma wf_iff_no_infinite_down_chain: "wf r \ (\f. \i. (f (Suc i), f i) \ r)" (is "_ \ \ ?ex") proof assume "wf r" show "\ ?ex" proof assume ?ex then obtain f where f: "(f (Suc i), f i) \ r" for i by blast from \wf r\ have minimal: "x \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" for x Q by (auto simp: wf_eq_minimal) let ?Q = "{w. \i. w = f i}" fix n have "f n \ ?Q" by blast from minimal [OF this] obtain j where "(y, f j) \ r \ y \ ?Q" for y by blast with this [OF \(f (Suc j), f j) \ r\] have "f (Suc j) \ ?Q" by simp then show False by blast qed next assume "\ ?ex" then show "wf r" proof (rule contrapos_np) assume "\ wf r" then obtain Q x where x: "x \ Q" and rec: "z \ Q \ \y. (y, z) \ r \ y \ Q" for z by (auto simp add: wf_eq_minimal) obtain descend :: "nat \ 'a" where descend_0: "descend 0 = x" and descend_Suc: "descend (Suc n) = (SOME y. y \ Q \ (y, descend n) \ r)" for n by (rule that [of "rec_nat x (\_ rec. (SOME y. y \ Q \ (y, rec) \ r))"]) simp_all have descend_Q: "descend n \ Q" for n proof (induct n) case 0 with x show ?case by (simp only: descend_0) next case Suc then show ?case by (simp only: descend_Suc) (rule someI2_ex; use rec in blast) qed have "(descend (Suc i), descend i) \ r" for i by (simp only: descend_Suc) (rule someI2_ex; use descend_Q rec in blast) then show "\f. \i. (f (Suc i), f i) \ r" by blast qed qed lemma wf_no_infinite_down_chainE: assumes "wf r" obtains k where "(f (Suc k), f k) \ r" using assms wf_iff_no_infinite_down_chain[of r] by blast text \A dynamically-scoped fact for TFL\ lemma tfl_some: "\P x. P x \ P (Eps P)" by (blast intro: someI) subsection \An aside: bounded accessible part\ text \Finite monotone eventually stable sequences\ lemma finite_mono_remains_stable_implies_strict_prefix: fixes f :: "nat \ 'a::order" assumes S: "finite (range f)" "mono f" and eq: "\n. f n = f (Suc n) \ f (Suc n) = f (Suc (Suc n))" shows "\N. (\n\N. \m\N. m < n \ f m < f n) \ (\n\N. f N = f n)" using assms proof - have "\n. f n = f (Suc n)" proof (rule ccontr) assume "\ ?thesis" then have "\n. f n \ f (Suc n)" by auto with \mono f\ have "\n. f n < f (Suc n)" by (auto simp: le_less mono_iff_le_Suc) with lift_Suc_mono_less_iff[of f] have *: "\n m. n < m \ f n < f m" by auto have "inj f" proof (intro injI) fix x y assume "f x = f y" then show "x = y" by (cases x y rule: linorder_cases) (auto dest: *) qed with \finite (range f)\ have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed then obtain n where n: "f n = f (Suc n)" .. define N where "N = (LEAST n. f n = f (Suc n))" have N: "f N = f (Suc N)" unfolding N_def using n by (rule LeastI) show ?thesis proof (intro exI[of _ N] conjI allI impI) fix n assume "N \ n" then have "\m. N \ m \ m \ n \ f m = f N" proof (induct rule: dec_induct) case base then show ?case by simp next case (step n) then show ?case using eq [rule_format, of "n - 1"] N by (cases n) (auto simp add: le_Suc_eq) qed from this[of n] \N \ n\ show "f N = f n" by auto next fix n m :: nat assume "m < n" "n \ N" then show "f m < f n" proof (induct rule: less_Suc_induct) case (1 i) then have "i < N" by simp then have "f i \ f (Suc i)" unfolding N_def by (rule not_less_Least) with \mono f\ show ?case by (simp add: mono_iff_le_Suc less_le) next case 2 then show ?case by simp qed qed qed lemma finite_mono_strict_prefix_implies_finite_fixpoint: fixes f :: "nat \ 'a set" assumes S: "\i. f i \ S" "finite S" and ex: "\N. (\n\N. \m\N. m < n \ f m \ f n) \ (\n\N. f N = f n)" shows "f (card S) = (\n. f n)" proof - from ex obtain N where inj: "\n m. n \ N \ m \ N \ m < n \ f m \ f n" and eq: "\n\N. f N = f n" by atomize auto have "i \ N \ i \ card (f i)" for i proof (induct i) case 0 then show ?case by simp next case (Suc i) with inj [of "Suc i" i] have "(f i) \ (f (Suc i))" by auto moreover have "finite (f (Suc i))" using S by (rule finite_subset) ultimately have "card (f i) < card (f (Suc i))" by (intro psubset_card_mono) with Suc inj show ?case by auto qed then have "N \ card (f N)" by simp also have "\ \ card S" using S by (intro card_mono) finally have \
: "f (card S) = f N" using eq by auto moreover have "\ (range f) \ f N" proof clarify fix x n assume "x \ f n" with eq inj [of N] show "x \ f N" by (cases "n < N") (auto simp: not_less) qed ultimately show ?thesis by auto qed subsection \More on injections, bijections, and inverses\ locale bijection = fixes f :: "'a \ 'a" assumes bij: "bij f" begin lemma bij_inv: "bij (inv f)" using bij by (rule bij_imp_bij_inv) lemma surj [simp]: "surj f" using bij by (rule bij_is_surj) lemma inj: "inj f" using bij by (rule bij_is_inj) lemma surj_inv [simp]: "surj (inv f)" using inj by (rule inj_imp_surj_inv) lemma inj_inv: "inj (inv f)" using surj by (rule surj_imp_inj_inv) lemma eqI: "f a = f b \ a = b" using inj by (rule injD) lemma eq_iff [simp]: "f a = f b \ a = b" by (auto intro: eqI) lemma eq_invI: "inv f a = inv f b \ a = b" using inj_inv by (rule injD) lemma eq_inv_iff [simp]: "inv f a = inv f b \ a = b" by (auto intro: eq_invI) lemma inv_left [simp]: "inv f (f a) = a" using inj by (simp add: inv_f_eq) lemma inv_comp_left [simp]: "inv f \ f = id" by (simp add: fun_eq_iff) lemma inv_right [simp]: "f (inv f a) = a" using surj by (simp add: surj_f_inv_f) lemma inv_comp_right [simp]: "f \ inv f = id" by (simp add: fun_eq_iff) lemma inv_left_eq_iff [simp]: "inv f a = b \ f b = a" by auto lemma inv_right_eq_iff [simp]: "b = inv f a \ f b = a" by auto end lemma infinite_imp_bij_betw: assumes infinite: "\ finite A" shows "\h. bij_betw h A (A - {a})" proof (cases "a \ A") case False then have "A - {a} = A" by blast then show ?thesis using bij_betw_id[of A] by auto next case True with infinite have "\ finite (A - {a})" by auto with infinite_iff_countable_subset[of "A - {a}"] obtain f :: "nat \ 'a" where "inj f" and f: "f ` UNIV \ A - {a}" by blast define g where "g n = (if n = 0 then a else f (Suc n))" for n define A' where "A' = g ` UNIV" have *: "\y. f y \ a" using f by blast have 3: "inj_on g UNIV \ g ` UNIV \ A \ a \ g ` UNIV" using \inj f\ f * unfolding inj_on_def g_def by (auto simp add: True image_subset_iff) then have 4: "bij_betw g UNIV A' \ a \ A' \ A' \ A" using inj_on_imp_bij_betw[of g] by (auto simp: A'_def) then have 5: "bij_betw (inv g) A' UNIV" by (auto simp add: bij_betw_inv_into) from 3 obtain n where n: "g n = a" by auto have 6: "bij_betw g (UNIV - {n}) (A' - {a})" by (rule bij_betw_subset) (use 3 4 n in \auto simp: image_set_diff A'_def\) define v where "v m = (if m < n then m else Suc m)" for m have "m < n \ m = n" if "\k. k < n \ m \ Suc k" for m using that [of "m-1"] by auto then have 7: "bij_betw v UNIV (UNIV - {n})" unfolding bij_betw_def inj_on_def v_def by auto define h' where "h' = g \ v \ (inv g)" with 5 6 7 have 8: "bij_betw h' A' (A' - {a})" by (auto simp add: bij_betw_trans) define h where "h b = (if b \ A' then h' b else b)" for b with 8 have "bij_betw h A' (A' - {a})" using bij_betw_cong[of A' h] by auto moreover have "\b \ A - A'. h b = b" by (auto simp: h_def) then have "bij_betw h (A - A') (A - A')" using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto moreover from 4 have "(A' \ (A - A') = {} \ A' \ (A - A') = A) \ ((A' - {a}) \ (A - A') = {} \ (A' - {a}) \ (A - A') = A - {a})" by blast ultimately have "bij_betw h A (A - {a})" using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp then show ?thesis by blast qed lemma infinite_imp_bij_betw2: assumes "\ finite A" shows "\h. bij_betw h A (A \ {a})" proof (cases "a \ A") case True then have "A \ {a} = A" by blast then show ?thesis using bij_betw_id[of A] by auto next case False let ?A' = "A \ {a}" from False have "A = ?A' - {a}" by blast moreover from assms have "\ finite ?A'" by auto ultimately obtain f where "bij_betw f ?A' A" using infinite_imp_bij_betw[of ?A' a] by auto then have "bij_betw (inv_into ?A' f) A ?A'" by (rule bij_betw_inv_into) then show ?thesis by auto qed lemma bij_betw_inv_into_left: "bij_betw f A A' \ a \ A \ inv_into A f (f a) = a" unfolding bij_betw_def by clarify (rule inv_into_f_f) lemma bij_betw_inv_into_right: "bij_betw f A A' \ a' \ A' \ f (inv_into A f a') = a'" unfolding bij_betw_def using f_inv_into_f by force lemma bij_betw_inv_into_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw (inv_into A f) B' B" by (auto simp: bij_betw_def intro: inj_on_inv_into) subsection \Specification package -- Hilbertized version\ lemma exE_some: "Ex P \ c \ Eps P \ P c" by (simp only: someI_ex) ML_file \Tools/choice_specification.ML\ subsection \Complete Distributive Lattices -- Properties depending on Hilbert Choice\ context complete_distrib_lattice begin lemma Sup_Inf: "\ (Inf ` A) = \ (Sup ` {f ` A |f. \B\A. f B \ B})" proof (rule order.antisym) show "\ (Inf ` A) \ \ (Sup ` {f ` A |f. \B\A. f B \ B})" using Inf_lower2 Sup_upper by (fastforce simp add: intro: Sup_least INF_greatest) next show "\ (Sup ` {f ` A |f. \B\A. f B \ B}) \ \ (Inf ` A)" proof (simp add: Inf_Sup, rule SUP_least, simp, safe) fix f assume "\Y. (\f. Y = f ` A \ (\Y\A. f Y \ Y)) \ f Y \ Y" then have B: "\ F . (\ Y \ A . F Y \ Y) \ \ Z \ A . f (F ` A) = F Z" by auto show "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ \(Inf ` A)" proof (cases "\ Z \ A . \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ Inf Z") case True from this obtain Z where [simp]: "Z \ A" and A: "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ Inf Z" by blast have B: "... \ \(Inf ` A)" by (simp add: SUP_upper) from A and B show ?thesis by simp next case False then have X: "\ Z . Z \ A \ \ x . x \ Z \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ x" using Inf_greatest by blast define F where "F = (\ Z . SOME x . x \ Z \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ x)" have C: "\Y. Y \ A \ F Y \ Y" using X by (simp add: F_def, rule someI2_ex, auto) have E: "\Y. Y \ A \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ F Y" using X by (simp add: F_def, rule someI2_ex, auto) from C and B obtain Z where D: "Z \ A " and Y: "f (F ` A) = F Z" by blast from E and D have W: "\ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ F Z" by simp have "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ f (F ` A)" using C by (blast intro: INF_lower) with W Y show ?thesis by simp qed qed qed lemma dual_complete_distrib_lattice: "class.complete_distrib_lattice Sup Inf sup (\) (>) inf \ \" by (simp add: class.complete_distrib_lattice.intro [OF dual_complete_lattice] class.complete_distrib_lattice_axioms_def Sup_Inf) lemma sup_Inf: "a \ \B = \((\) a ` B)" proof (rule order.antisym) show "a \ \B \ \((\) a ` B)" using Inf_lower sup.mono by (fastforce intro: INF_greatest) next have "\((\) a ` B) \ \(Sup ` {{f {a}, f B} |f. f {a} = a \ f B \ B})" by (rule INF_greatest, auto simp add: INF_lower) also have "... = \(Inf ` {{a}, B})" by (unfold Sup_Inf, simp) finally show "\((\) a ` B) \ a \ \B" by simp qed lemma inf_Sup: "a \ \B = \((\) a ` B)" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.sup_Inf) lemma INF_SUP: "(\y. \x. P x y) = (\f. \x. P (f x) x)" proof (rule order.antisym) show "(SUP x. INF y. P (x y) y) \ (INF y. SUP x. P x y)" by (rule SUP_least, rule INF_greatest, rule SUP_upper2, simp_all, rule INF_lower2, simp, blast) next have "(INF y. SUP x. ((P x y))) \ Inf (Sup ` {{P x y | x . True} | y . True })" (is "?A \ ?B") proof (rule INF_greatest, clarsimp) fix y have "?A \ (SUP x. P x y)" by (rule INF_lower, simp) also have "... \ Sup {uu. \x. uu = P x y}" by (simp add: full_SetCompr_eq) finally show "?A \ Sup {uu. \x. uu = P x y}" by simp qed also have "... \ (SUP x. INF y. P (x y) y)" proof (subst Inf_Sup, rule SUP_least, clarsimp) fix f assume A: "\Y. (\y. Y = {uu. \x. uu = P x y}) \ f Y \ Y" have " \(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ (\y. P (SOME x. f {P x y |x. True} = P x y) y)" proof (rule INF_greatest, clarsimp) fix y have "(INF x\{uu. \y. uu = {uu. \x. uu = P x y}}. f x) \ f {uu. \x. uu = P x y}" by (rule INF_lower, blast) also have "... \ P (SOME x. f {uu . \x. uu = P x y} = P x y) y" by (rule someI2_ex) (use A in auto) finally show "\(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ P (SOME x. f {uu. \x. uu = P x y} = P x y) y" by simp qed also have "... \ (SUP x. INF y. P (x y) y)" by (rule SUP_upper, simp) finally show "\(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ (\x. \y. P (x y) y)" by simp qed finally show "(INF y. SUP x. P x y) \ (SUP x. INF y. P (x y) y)" by simp qed lemma INF_SUP_set: "(\B\A. \(g ` B)) = (\B\{f ` A |f. \C\A. f C \ C}. \(g ` B))" (is "_ = (\B\?F. _)") proof (rule order.antisym) have "\ ((g \ f) ` A) \ \ (g ` B)" if "\B. B \ A \ f B \ B" "B \ A" for f B using that by (auto intro: SUP_upper2 INF_lower2) then show "(\x\?F. \a\x. g a) \ (\x\A. \a\x. g a)" by (auto intro!: SUP_least INF_greatest simp add: image_comp) next show "(\x\A. \a\x. g a) \ (\x\?F. \a\x. g a)" proof (cases "{} \ A") case True then show ?thesis by (rule INF_lower2) simp_all next case False {fix x have "(\x\A. \x\x. g x) \ (\u. if x \ A then if u \ x then g u else \ else \)" proof (cases "x \ A") case True then show ?thesis by (intro INF_lower2 SUP_least SUP_upper2) auto qed auto } then have "(\Y\A. \a\Y. g a) \ (\Y. \y. if Y \ A then if y \ Y then g y else \ else \)" by (rule INF_greatest) also have "... = (\x. \Y. if Y \ A then if x Y \ Y then g (x Y) else \ else \)" by (simp only: INF_SUP) also have "... \ (\x\?F. \a\x. g a)" proof (rule SUP_least) show "(\B. if B \ A then if x B \ B then g (x B) else \ else \) \ (\x\?F. \x\x. g x)" for x proof - define G where "G \ \Y. if x Y \ Y then x Y else (SOME x. x \Y)" have "\Y\A. G Y \ Y" using False some_in_eq G_def by auto then have A: "G ` A \ ?F" by blast show "(\Y. if Y \ A then if x Y \ Y then g (x Y) else \ else \) \ (\x\?F. \x\x. g x)" by (fastforce simp: G_def intro: SUP_upper2 [OF A] INF_greatest INF_lower2) qed qed finally show ?thesis by simp qed qed lemma SUP_INF: "(\y. \x. P x y) = (\x. \y. P (x y) y)" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP) lemma SUP_INF_set: "(\x\A. \ (g ` x)) = (\x\{f ` A |f. \Y\A. f Y \ Y}. \ (g ` x))" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP_set) end (*properties of the former complete_distrib_lattice*) context complete_distrib_lattice begin lemma sup_INF: "a \ (\b\B. f b) = (\b\B. a \ f b)" by (simp add: sup_Inf image_comp) lemma inf_SUP: "a \ (\b\B. f b) = (\b\B. a \ f b)" by (simp add: inf_Sup image_comp) lemma Inf_sup: "\B \ a = (\b\B. b \ a)" by (simp add: sup_Inf sup_commute) lemma Sup_inf: "\B \ a = (\b\B. b \ a)" by (simp add: inf_Sup inf_commute) lemma INF_sup: "(\b\B. f b) \ a = (\b\B. f b \ a)" by (simp add: sup_INF sup_commute) lemma SUP_inf: "(\b\B. f b) \ a = (\b\B. f b \ a)" by (simp add: inf_SUP inf_commute) lemma Inf_sup_eq_top_iff: "(\B \ a = \) \ (\b\B. b \ a = \)" by (simp only: Inf_sup INF_top_conv) lemma Sup_inf_eq_bot_iff: "(\B \ a = \) \ (\b\B. b \ a = \)" by (simp only: Sup_inf SUP_bot_conv) lemma INF_sup_distrib2: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" by (subst INF_commute) (simp add: sup_INF INF_sup) lemma SUP_inf_distrib2: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" by (subst SUP_commute) (simp add: inf_SUP SUP_inf) end context complete_boolean_algebra begin lemma dual_complete_boolean_algebra: "class.complete_boolean_algebra Sup Inf sup (\) (>) inf \ \ (\x y. x \ - y) uminus" by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) end instantiation set :: (type) complete_distrib_lattice begin instance proof (standard, clarsimp) fix A :: "(('a set) set) set" fix x::'a assume A: "\\\A. \X\\. x \ X" define F where "F \ \Y. SOME X. Y \ A \ X \ Y \ x \ X" have "(\S \ F ` A. x \ S)" using A unfolding F_def by (fastforce intro: someI2_ex) moreover have "\Y\A. F Y \ Y" using A unfolding F_def by (fastforce intro: someI2_ex) then have "\f. F ` A = f ` A \ (\Y\A. f Y \ Y)" by blast ultimately show "\X. (\f. X = f ` A \ (\Y\A. f Y \ Y)) \ (\S\X. x \ S)" by auto qed end instance set :: (type) complete_boolean_algebra .. instantiation "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice begin instance by standard (simp add: le_fun_def INF_SUP_set image_comp) end instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. context complete_linorder begin subclass complete_distrib_lattice proof (standard, rule ccontr) fix A :: "'a set set" let ?F = "{f ` A |f. \Y\A. f Y \ Y}" assume "\ \(Sup ` A) \ \(Inf ` ?F)" then have C: "\(Sup ` A) > \(Inf ` ?F)" by (simp add: not_le) show False proof (cases "\ z . \(Sup ` A) > z \ z > \(Inf ` ?F)") case True then obtain z where A: "z < \(Sup ` A)" and X: "z > \(Inf ` ?F)" by blast then have B: "\Y. Y \ A \ \k \Y . z < k" using local.less_Sup_iff by(force dest: less_INF_D) define G where "G \ \Y. SOME k . k \ Y \ z < k" have E: "\Y. Y \ A \ G Y \ Y" using B unfolding G_def by (fastforce intro: someI2_ex) have "z \ Inf (G ` A)" proof (rule INF_greatest) show "\Y. Y \ A \ z \ G Y" using B unfolding G_def by (fastforce intro: someI2_ex) qed also have "... \ \(Inf ` ?F)" by (rule SUP_upper) (use E in blast) finally have "z \ \(Inf ` ?F)" by simp with X show ?thesis using local.not_less by blast next case False have B: "\Y. Y \ A \ \ k \Y . \(Inf ` ?F) < k" using C local.less_Sup_iff by(force dest: less_INF_D) define G where "G \ \ Y . SOME k . k \ Y \ \(Inf ` ?F) < k" have E: "\Y. Y \ A \ G Y \ Y" using B unfolding G_def by (fastforce intro: someI2_ex) have "\Y. Y \ A \ \(Sup ` A) \ G Y" using B False local.leI unfolding G_def by (fastforce intro: someI2_ex) then have "\(Sup ` A) \ Inf (G ` A)" by (simp add: local.INF_greatest) also have "Inf (G ` A) \ \(Inf ` ?F)" by (rule SUP_upper) (use E in blast) finally have "\(Sup ` A) \ \(Inf ` ?F)" by simp with C show ?thesis using not_less by blast qed qed end - - end diff --git a/src/HOL/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2548 +1,2560 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" proof - have "P (Abs_Nat (Rep_Nat n))" using assms unfolding Zero_nat_def Suc_def by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) then show ?thesis by (simp add: Rep_Nat_inverse) qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" apply atomize_elim apply (rename_tac n, induct_tac n rule: nat_induct0, auto) apply (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) apply (simp only: Suc_not_Zero) done \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" by (simp add: eq_commute flip: mult_eq_1_iff) lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P proof (rule major [THEN lessE]) show "Suc n = Suc m \ P" using eq by blast show "\j. \m < j; Suc n = Suc j\ \ P" by (blast intro: less) qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P proof (rule major [THEN lessE]) show "k = Suc (Suc i) \ P" using lessI minor by iprover show "\j. \Suc i < j; k = Suc j\ \ P" using Suc_lessD minor by iprover qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ lemma add_less_mono: fixes i j k l :: nat assumes "i < j" "k < l" shows "i + k < j + l" proof - have "i + k < j + k" by (simp add: add_less_mono1 assms) also have "... < j + l" using \i < j\ by (induction j) (auto simp: assms) finally show ?thesis . qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ global_interpretation bot_nat_0: ordering_top \(\)\ \(>)\ \0::nat\ by standard simp global_interpretation max_nat: semilattice_neutr_order max \0::nat\ \(\)\ \(>)\ by standard (simp add: max_def) lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat by (fact max_nat.left_neutral) lemma max_0R [simp]: "max n 0 = n" for n :: nat by (fact max_nat.right_neutral) lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) lemma Least_Suc: assumes "P n" "\ P 0" shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" proof (cases n) case (Suc m) show ?thesis proof (rule antisym) show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" using assms Suc by (force intro: LeastI Least_le) have \
: "P (LEAST x. P x)" by (blast intro: LeastI assms) show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" proof (cases "(LEAST n. P n)") case 0 then show ?thesis using \
by (simp add: assms) next case Suc with \
show ?thesis by (auto simp: Least_le) qed qed qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp lemma ex_least_nat_le: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\k\n. (\i P i) \ P k" proof (cases n) case (Suc m) with assms show ?thesis by (blast intro: Least_le LeastI_ex dest: not_less_Least) qed (use assms in auto) lemma ex_least_nat_less: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\ki\k. \ P i) \ P (Suc k)" proof (cases n) case (Suc m) then obtain k where k: "k \ n" "\i P i" "P k" using ex_least_nat_le [OF assms] by blast show ?thesis by (cases k) (use assms k less_eq_Suc_le in auto) qed (use assms in auto) lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" proof (rule infinite_descent) show "\n. \ P n \ \m P m" using assms by (case_tac "n > 0") auto qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent, auto) show "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat proof (intro iffI conjI) assume m: "m * k < n * k" then show "0 < k" by (cases k) auto show "m < n" proof (cases k) case 0 then show ?thesis using m by auto next case (Suc k') then show ?thesis using m by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) qed next assume "0 < k \ m < n" then show "m * k < n * k" by (blast intro: mult_less_mono1) qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ lemma mult_eq_self_implies_10: fixes m n :: nat assumes "m = m * n" shows "n = 1 \ m = 0" proof (rule disjCI) assume "m \ 0" show "n = 1" proof (cases n "1::nat" rule: linorder_cases) case greater show ?thesis using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto qed (use assms \m \ 0\ in auto) qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ -definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" +context +begin + +qualified definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) -hide_const (open) funpow +end lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) +lemma bij_betw_funpow: \<^marker>\contributor \Lars Noschinski\\ + assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" +proof (induct n) + case 0 then show ?case by (auto simp: id_def[symmetric]) +next + case (Suc n) + then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) +qed + subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case 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] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: order.eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) 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" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" using of_nat_0 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" using of_nat_1 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" unfolding Nats_def using of_nat_add [symmetric] by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" unfolding Nats_def using of_nat_mult [symmetric] by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = \fn phi => try o Nat_Arith.cancel_eq_conv\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = \fn phi => try o Nat_Arith.cancel_less_conv\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = \fn phi => try o Nat_Arith.cancel_le_conv\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = \fn phi => try o Nat_Arith.cancel_diff_conv\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed lemma fixes k::nat assumes "P k" and minor: "\y. P y \ y \ b" shows GreatestI_nat: "P (Greatest P)" and Greatest_le_nat: "k \ Greatest P" proof - obtain x where "P x" "\y. P y \ y \ x" using assms ex_has_greatest_nat by blast with \P k\ show "P (Greatest P)" "k \ Greatest P" using GreatestI2_order by blast+ qed lemma GreatestI_ex_nat: "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) lemma dvd_diffD: fixes k m n :: nat assumes "k dvd m - n" "k dvd n" "n \ m" shows "k dvd m" proof - have "k dvd n + (m - n)" using assms by (blast intro: dvd_add) with assms show ?thesis by simp qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed lemma dvd_mult_cancel1: fixes m n :: nat assumes "0 < m" shows "m * n dvd m \ n = 1" proof assume "m * n dvd m" then have "m * n dvd m * 1" by simp then have "n dvd 1" by (iprover intro: assms dvd_mult_cancel) then show "n = 1" by auto qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') (*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*) lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" for i j k :: nat by (fact le_diff_conv2) lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def lemma size_neq_size_imp_neq: "size x \ size y \ x \ y" by (erule contrapos_nn) (rule arg_cong) subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2569 +1,2589 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval imports Divides begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by (simp add: order_class.order.eq_iff) (auto intro: order_trans) lemma atLeastAtMost_singleton_iff[simp]: "{a .. b} = {c} \ a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" apply (auto simp:subset_eq Ball_def not_le) apply(frule_tac x=a in spec) apply(erule_tac x=d in allE) apply (auto simp: ) done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: \{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d\ (is \?P \ ?Q\) proof assume ?Q then show ?P by auto next assume ?P then have \a < x \ x \ b \ c < x \ x \ d\ for x by (simp add: set_eq_iff) from this [of a] this [of b] this [of c] this [of d] show ?Q by auto qed lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" proof (rule UN_finite_subset, rule) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: "(\n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) apply auto apply (force simp add: atLeastLessThan_add_Un [of 0])+ done subsubsection \Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" proof (auto simp: inj_on_def) fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed +lemma inj_on_funpow_least: \<^marker>\contributor \Lars Noschinski\\ + \inj_on (\k. (f ^^ k) s) {0.. + if \(f ^^ n) s = s\ \\m. 0 < m \ m < n \ (f ^^ m) s \ s\ +proof - + { fix k l assume A: "k < n" "l < n" "k \ l" "(f ^^ k) s = (f ^^ l) s" + define k' l' where "k' = min k l" and "l' = max k l" + with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" + by (auto simp: min_def max_def) + + have "s = (f ^^ ((n - l') + l')) s" using that \l' < n\ by simp + also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) + also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') + also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) + finally have "(f ^^ (n - l' + k')) s = s" by simp + moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ + ultimately have False using that(2) by auto + } + then show ?thesis by (intro inj_onI) auto +qed + subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ lemma image_atLeastZeroLessThan_int: "0 \ u ==> {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed lemma card_le_Suc_Max: "finite S \ card S \ Suc (Max S)" proof (rule classical) assume "finite S" and "\ Suc (Max S) \ card S" then have "Suc (Max S) < card S" by simp with `finite S` have "S \ {0..Max S}" by auto hence "card S \ card {0..Max S}" by (intro card_mono; auto) thus "card S \ Suc (Max S)" by simp qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] apply (auto simp: linorder_not_le) apply (force intro: leI)+ done lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" by (metis assms le_add_diff_inverse) with that show thesis proof (induct n' arbitrary: S) case 0 then show ?case by (cases "finite S") auto next case Suc then show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) qed qed subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" by rule simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_int_lessThan_int_shift: "F g {int m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma last_plus: fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" by (simp add: commute last_plus) lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma card_sum_le_nat_sum: "\ {0.. \ S" proof (cases "finite S") case True then show ?thesis proof (induction "card S" arbitrary: S) case (Suc x) then have "Max S \ x" using card_le_Suc_Max by fastforce let ?S' = "S - {Max S}" from Suc have "Max S \ S" by (auto intro: Max_in) hence cards: "card S = Suc (card ?S')" using `finite S` by (intro card.remove; auto) hence "\ {0.. \ ?S'" using Suc by (intro Suc; auto) hence "\ {0.. \ ?S' + Max S" using `Max S \ x` by simp also have "... = \ S" using sum.remove[OF `finite S` `Max S \ S`, where g="\x. x"] by simp finally show ?case using cards Suc by auto qed simp qed simp lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxk = 0..k = 0..k = Suc 0..k = Suc 0..k = 0..Shifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" using sum_gp_multiplied [of m n x] apply auto by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end