diff --git a/src/HOL/Library/Ramsey.thy b/src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy +++ b/src/HOL/Library/Ramsey.thy @@ -1,878 +1,907 @@ (* Title: HOL/Library/Ramsey.thy Author: Tom Ridge. Full finite version by L C Paulson. *) section \Ramsey's Theorem\ theory Ramsey imports Infinite_Set FuncSet begin subsection \Preliminary definitions\ subsubsection \The $n$-element subsets of a set $A$\ definition nsets :: "['a set, nat] \ 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999) where "nsets A n \ {N. N \ A \ finite N \ card N = n}" lemma nsets_mono: "A \ B \ nsets A n \ nsets B n" by (auto simp: nsets_def) lemma nsets_2_eq: "nsets A 2 = (\x\A. \y\A - {x}. {{x, y}})" by (auto simp: nsets_def card_2_iff) +lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})" + by (auto simp: nsets_2_eq) + +lemma doubleton_in_nsets_2 [simp]: "{x,y} \ [A]\<^bsup>2\<^esup> \ x \ A \ y \ A \ x \ y" + by (auto simp: nsets_2_eq Set.doubleton_eq_iff) + lemma binomial_eq_nsets: "n choose k = card (nsets {0.. finite A \ card A < r" unfolding nsets_def proof (intro iffI conjI) assume that: "{N. N \ A \ finite N \ card N = r} = {}" show "finite A" using infinite_arbitrarily_large that by auto then have "\ r \ card A" using that by (simp add: set_eq_iff) (metis finite_subset get_smaller_card [of A r]) then show "card A < r" using not_less by blast next show "{N. N \ A \ finite N \ card N = r} = {}" if "finite A \ card A < r" using that card_mono leD by auto qed -lemma nsets_eq_empty: "n < r \ nsets {..finite A; card A < r\ \ nsets A r = {}" by (simp add: nsets_eq_empty_iff) lemma nsets_empty_iff: "nsets {} r = (if r=0 then {{}} else {})" by (auto simp: nsets_def) lemma nsets_singleton_iff: "nsets {a} r = (if r=0 then {{}} else if r=1 then {{a}} else {})" by (auto simp: nsets_def card_gt_0_iff subset_singleton_iff) lemma nsets_self [simp]: "nsets {..x. {x}) ` A" using card_eq_SucD by (force simp: nsets_def) lemma inj_on_nsets: assumes "inj_on f A" shows "inj_on (\X. f ` X) ([A]\<^bsup>n\<^esup>)" using assms unfolding nsets_def by (metis (no_types, lifting) inj_on_inverseI inv_into_image_cancel mem_Collect_eq) lemma bij_betw_nsets: assumes "bij_betw f A B" shows "bij_betw (\X. f ` X) ([A]\<^bsup>n\<^esup>) ([B]\<^bsup>n\<^esup>)" proof - have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>" using assms apply (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset) by (metis card_image inj_on_finite order_refl subset_image_inj) with assms show ?thesis by (auto simp: bij_betw_def inj_on_nsets) qed +lemma nset_image_obtains: + assumes "X \ [f`A]\<^bsup>k\<^esup>" "inj_on f A" + obtains Y where "Y \ [A]\<^bsup>k\<^esup>" "X = f ` Y" + using assms + apply (clarsimp simp add: nsets_def subset_image_iff) + by (metis card_image finite_imageD inj_on_subset) + +lemma nsets_image_funcset: + assumes "g \ S \ T" and "inj_on g S" + shows "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" + using assms + by (fastforce simp add: nsets_def card_image inj_on_subset subset_iff simp flip: image_subset_iff_funcset) + +lemma nsets_compose_image_funcset: + assumes f: "f \ [T]\<^bsup>k\<^esup> \ D" and "g \ S \ T" and "inj_on g S" + shows "f \ (\X. g ` X) \ [S]\<^bsup>k\<^esup> \ D" +proof - + have "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" + using assms by (simp add: nsets_image_funcset) + then show ?thesis + using f by fastforce +qed + subsubsection \Partition predicates\ definition partn :: "'a set \ nat \ nat \ 'b set \ bool" where "partn \ \ \ \ \ \f \ nsets \ \ \ \. \H \ nsets \ \. \\\\. f ` (nsets H \) \ {\}" definition partn_lst :: "'a set \ nat list \ nat \ bool" where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. \i < length \. \H \ nsets \ (\!i). f ` (nsets H \) \ {i}" lemma partn_lst_greater_resource: fixes M::nat assumes M: "partn_lst {.. \" and "M \ N" shows "partn_lst {.. \" proof (clarsimp simp: partn_lst_def) fix f assume "f \ nsets {.. \ {..}" then have "f \ nsets {.. \ {..}" by (meson Pi_anti_mono \M \ N\ lessThan_subset_iff nsets_mono subsetD) then obtain i H where i: "i < length \" and H: "H \ nsets {.. ! i)" and subi: "f ` nsets H \ \ {i}" using M partn_lst_def by blast have "H \ nsets {.. ! i)" using \M \ N\ H by (auto simp: nsets_def subset_iff) then show "\i. \H\nsets {.. ! i). f ` nsets H \ \ {i}" using i subi by blast qed lemma partn_lst_fewer_colours: assumes major: "partn_lst \ (n#\) \" and "n \ \" shows "partn_lst \ \ \" proof (clarsimp simp: partn_lst_def) fix f :: "'a set \ nat" assume f: "f \ [\]\<^bsup>\\<^esup> \ {..}" then obtain i H where i: "i < Suc (length \)" and H: "H \ [\]\<^bsup>((n # \) ! i)\<^esup>" and hom: "\x\[H]\<^bsup>\\<^esup>. Suc (f x) = i" using \n \ \\ major [unfolded partn_lst_def, rule_format, of "Suc o f"] by (fastforce simp: image_subset_iff nsets_eq_empty_iff) show "\i. \H\nsets \ (\ ! i). f ` [H]\<^bsup>\\<^esup> \ {i}" proof (cases i) case 0 then have "[H]\<^bsup>\\<^esup> = {}" using hom by blast then show ?thesis using 0 H \n \ \\ by (simp add: nsets_eq_empty_iff) (simp add: nsets_def) next case (Suc i') then show ?thesis using i H hom by auto qed qed lemma partn_lst_eq_partn: "partn_lst {..Finite versions of Ramsey's theorem\ text \ To distinguish the finite and infinite ones, lower and upper case names are used. \ subsubsection \Trivial cases\ text \Vacuous, since we are dealing with 0-sets!\ lemma ramsey0: "\N::nat. partn_lst {..Just the pigeon hole principle, since we are dealing with 1-sets\ lemma ramsey1: "\N::nat. partn_lst {..iH\nsets {.. {i}" if "f \ nsets {.. {.. \i. {q. q \ q0+q1 \ f {q} = i}" have "A 0 \ A 1 = {..q0 + q1}" using that by (auto simp: A_def PiE_iff nsets_one lessThan_Suc_atMost le_Suc_eq) moreover have "A 0 \ A 1 = {}" by (auto simp: A_def) ultimately have "q0 + q1 \ card (A 0) + card (A 1)" by (metis card_Un_le card_atMost eq_imp_le le_SucI le_trans) then consider "card (A 0) \ q0" | "card (A 1) \ q1" by linarith then obtain i where "i < Suc (Suc 0)" "card (A i) \ [q0, q1] ! i" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) then obtain B where "B \ A i" "card B = [q0, q1] ! i" "finite B" by (meson finite_subset get_smaller_card infinite_arbitrarily_large) then have "B \ nsets {.. f ` nsets B (Suc 0) \ {i}" by (auto simp: A_def nsets_def card_1_singleton_iff) then show ?thesis using \i < Suc (Suc 0)\ by auto qed then show ?thesis by (clarsimp simp: partn_lst_def) blast qed subsubsection \Ramsey's theorem with two colours and arbitrary exponents (hypergraph version)\ proposition ramsey2_full: "\N::nat. partn_lst {.. 0" by simp show ?thesis using Suc.prems proof (induct k \ "q1 + q2" arbitrary: q1 q2) case 0 show ?case proof show "partn_lst {..<1::nat} [q1, q2] (Suc r)" using nsets_empty_iff subset_insert 0 by (fastforce simp: partn_lst_def funcset_to_empty_iff nsets_eq_empty image_subset_iff) qed next case (Suc k) consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto then show ?case proof cases case 1 then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)" unfolding partn_lst_def using \r > 0\ by (fastforce simp add: nsets_empty_iff nsets_singleton_iff lessThan_Suc) then show ?thesis by blast next case 2 with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto then obtain p1 p2::nat where p1: "partn_lst {..iH\nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \ {i}" if f: "f \ nsets {..p} (Suc r) \ {.. \R. f (insert p R)" have "f (insert p i) \ {.. nsets {.. nsets {.. {.. {i}" and U: "U \ nsets {.. {.. nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q1" unfolding nsets_def using \q1 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q1" by (simp add: izero inq1) ultimately show ?thesis by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) next case jone then have "u ` V \ nsets {..p} q2" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD) ultimately show ?thesis using jone not_less_eq by fastforce qed next case ione then have "U \ nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q2" unfolding nsets_def using \q2 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q2" by (simp add: ione inq1) ultimately show ?thesis by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc) next case jzero then have "u ` V \ nsets {..p} q1" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj apply (clarsimp simp add: h_def image_subset_iff nsets_def) by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj) ultimately show ?thesis using jzero not_less_eq by fastforce qed qed qed then show "partn_lst {..Full Ramsey's theorem with multiple colours and arbitrary exponents\ theorem ramsey_full: "\N::nat. partn_lst {.. "length qs" arbitrary: qs) case 0 then show ?case by (rule_tac x=" r" in exI) (simp add: partn_lst_def) next case (Suc k) note IH = this show ?case proof (cases k) case 0 with Suc obtain q where "qs = [q]" by (metis length_0_conv length_Suc_conv) then show ?thesis by (rule_tac x=q in exI) (auto simp: partn_lst_def funcset_to_empty_iff) next case (Suc k') then obtain q1 q2 l where qs: "qs = q1#q2#l" by (metis Suc.hyps(2) length_Suc_conv) then obtain q::nat where q: "partn_lst {..qs = q1 # q2 # l\ by fastforce have keq: "Suc (length l) = k" using IH qs by auto show ?thesis proof (intro exI conjI) show "partn_lst {.. nsets {.. {.. \X. if f X < Suc (Suc 0) then 0 else f X - Suc 0" have "g \ nsets {.. {.. {i}" and U: "U \ nsets {..iH\nsets {.. {i}" proof (cases "i = 0") case True then have "U \ nsets {.. {0, Suc 0}" using U gi unfolding g_def by (auto simp: image_subset_iff) then obtain u where u: "bij_betw u {.. {..U \ nsets {.. mem_Collect_eq nsets_def) have u_nsets: "u ` X \ nsets {.. nsets {.. \X. f (u ` X)" have "f (u ` X) < Suc (Suc 0)" if "X \ nsets {.. nsets U r" using u u_nsets that by (auto simp: nsets_def bij_betwE subset_eq) then show ?thesis using f01 by auto qed then have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. (\x. (u ` x)) ` nsets V r" apply (clarsimp simp add: nsets_def image_iff) by (metis card_eq_0_iff card_image image_is_empty subset_image_inj) then have "f ` nsets (u ` V) r \ h ` nsets V r" by (auto simp: h_def) then show "f ` nsets (u ` V) r \ {j}" using hj by auto show "(u ` V) \ nsets {.. {Suc i}" using i gi False apply (auto simp: g_def image_subset_iff) by (metis Suc_lessD Suc_pred g_def gi image_subset_iff not_less_eq singleton_iff) show "U \ nsets {..Simple graph version\ text \This is the most basic version in terms of cliques and independent sets, i.e. the version for graphs and 2 colours. \ definition "clique V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" definition "indep V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" lemma ramsey2: "\r\1. \(V::'a set) (E::'a set set). finite V \ card V \ r \ (\R \ V. card R = m \ clique R E \ card R = n \ indep R E)" proof - obtain N where "N \ Suc 0" and N: "partn_lst {..R\V. card R = m \ clique R E \ card R = n \ indep R E" if "finite V" "N \ card V" for V :: "'a set" and E :: "'a set set" proof - from that obtain v where u: "inj_on v {.. V" by (metis card_le_inj card_lessThan finite_lessThan) define f where "f \ \e. if v ` e \ E then 0 else Suc 0" have f: "f \ nsets {.. {.. {i}" and U: "U \ nsets {.. V" using U u by (auto simp: image_subset_iff nsets_def) show "card (v ` U) = m \ clique (v ` U) E \ card (v ` U) = n \ indep (v ` U) E" using i unfolding numeral_2_eq_2 using gi U u apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq) apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm) done qed qed then show ?thesis using \Suc 0 \ N\ by auto qed subsection \Preliminaries\ subsubsection \``Axiom'' of Dependent Choice\ primrec choice :: "('a \ bool) \ ('a \ 'a) set \ nat \ 'a" where \ \An integer-indexed chain of choices\ choice_0: "choice P r 0 = (SOME x. P x)" | choice_Suc: "choice P r (Suc n) = (SOME y. P y \ (choice P r n, y) \ r)" lemma choice_n: assumes P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" shows "P (choice P r n)" proof (induct n) case 0 show ?case by (force intro: someI P0) next case Suc then show ?case by (auto intro: someI2_ex [OF Pstep]) qed lemma dependent_choice: assumes trans: "trans r" and P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" obtains f :: "nat \ 'a" where "\n. P (f n)" and "\n m. n < m \ (f n, f m) \ r" proof fix n show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) next fix n m :: nat assume "n < m" from Pstep [OF choice_n [OF P0 Pstep]] have "(choice P r k, choice P r (Suc k)) \ r" for k by (auto intro: someI2_ex) then show "(choice P r n, choice P r m) \ r" by (auto intro: less_Suc_induct [OF \n < m\] transD [OF trans]) qed subsubsection \Partition functions\ definition part_fn :: "nat \ nat \ 'a set \ ('a set \ nat) \ bool" \ \the function \<^term>\f\ partitions the \<^term>\r\-subsets of the typically infinite set \<^term>\Y\ into \<^term>\s\ distinct categories.\ where "part_fn r s Y f \ (f \ nsets Y r \ {..For induction, we decrease the value of \<^term>\r\ in partitions.\ lemma part_fn_Suc_imp_part_fn: "\infinite Y; part_fn (Suc r) s Y f; y \ Y\ \ part_fn r s (Y - {y}) (\u. f (insert y u))" by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert) lemma part_fn_subset: "part_fn r s YY f \ Y \ YY \ part_fn r s Y f" unfolding part_fn_def nsets_def by blast subsection \Ramsey's Theorem: Infinitary Version\ lemma Ramsey_induction: fixes s r :: nat and YY :: "'a set" and f :: "'a set \ nat" assumes "infinite YY" "part_fn r s YY f" shows "\Y' t'. Y' \ YY \ infinite Y' \ t' < s \ (\X. X \ Y' \ finite X \ card X = r \ f X = t')" using assms proof (induct r arbitrary: YY f) case 0 then show ?case by (auto simp add: part_fn_def card_eq_0_iff cong: conj_cong) next case (Suc r) show ?case proof - from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \ YY" by blast let ?ramr = "{((y, Y, t), (y', Y', t')). y' \ Y \ Y' \ Y}" let ?propr = "\(y, Y, t). y \ YY \ y \ Y \ Y \ YY \ infinite Y \ t < s \ (\X. X\Y \ finite X \ card X = r \ (f \ insert y) X = t)" from Suc.prems have infYY': "infinite (YY - {yy})" by auto from Suc.prems have partf': "part_fn r s (YY - {yy}) (f \ insert yy)" by (simp add: o_def part_fn_Suc_imp_part_fn yy) have transr: "trans ?ramr" by (force simp add: trans_def) from Suc.hyps [OF infYY' partf'] obtain Y0 and t0 where "Y0 \ YY - {yy}" "infinite Y0" "t0 < s" "X \ Y0 \ finite X \ card X = r \ (f \ insert yy) X = t0" for X by blast with yy have propr0: "?propr(yy, Y0, t0)" by blast have proprstep: "\y. ?propr y \ (x, y) \ ?ramr" if x: "?propr x" for x proof (cases x) case (fields yx Yx tx) with x obtain yx' where yx': "yx' \ Yx" by (blast dest: infinite_imp_nonempty) from fields x have infYx': "infinite (Yx - {yx'})" by auto with fields x yx' Suc.prems have partfx': "part_fn r s (Yx - {yx'}) (f \ insert yx')" by (simp add: o_def part_fn_Suc_imp_part_fn part_fn_subset [where YY=YY and Y=Yx]) from Suc.hyps [OF infYx' partfx'] obtain Y' and t' where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" "X \ Y' \ finite X \ card X = r \ (f \ insert yx') X = t'" for X by blast from fields x Y' yx' have "?propr (yx', Y', t') \ (x, (yx', Y', t')) \ ?ramr" by blast then show ?thesis .. qed from dependent_choice [OF transr propr0 proprstep] obtain g where pg: "?propr (g n)" and rg: "n < m \ (g n, g m) \ ?ramr" for n m :: nat by blast let ?gy = "fst \ g" let ?gt = "snd \ snd \ g" have rangeg: "\k. range ?gt \ {.. range ?gt" then obtain n where "x = ?gt n" .. with pg [of n] show "x \ {.. ?gy m'" by (cases "g m", cases "g m'") auto qed show ?thesis proof (intro exI conjI) from pg show "?gy ` {n. ?gt n = s'} \ YY" by (auto simp add: Let_def split_beta) from infeqs' show "infinite (?gy ` {n. ?gt n = s'})" by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) show "s' < s" by (rule less') show "\X. X \ ?gy ` {n. ?gt n = s'} \ finite X \ card X = Suc r \ f X = s'" proof - have "f X = s'" if X: "X \ ?gy ` {n. ?gt n = s'}" and cardX: "finite X" "card X = Suc r" for X proof - from X obtain AA where AA: "AA \ {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" by (auto simp add: subset_image_iff) with cardX have "AA \ {}" by auto then have AAleast: "(LEAST x. x \ AA) \ AA" by (auto intro: LeastI_ex) show ?thesis proof (cases "g (LEAST x. x \ AA)") case (fields ya Ya ta) with AAleast Xeq have ya: "ya \ X" by (force intro!: rev_image_eqI) then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb) also have "\ = ta" proof - have *: "X - {ya} \ Ya" proof fix x assume x: "x \ X - {ya}" then obtain a' where xeq: "x = ?gy a'" and a': "a' \ AA" by (auto simp add: Xeq) with fields x have "a' \ (LEAST x. x \ AA)" by auto with Least_le [of "\x. x \ AA", OF a'] have "(LEAST x. x \ AA) < a'" by arith from xeq fields rg [OF this] show "x \ Ya" by auto qed have "card (X - {ya}) = r" by (simp add: cardX ya) with pg [of "LEAST x. x \ AA"] fields cardX * show ?thesis by (auto simp del: insert_Diff_single) qed also from AA AAleast fields have "\ = s'" by auto finally show ?thesis . qed qed then show ?thesis by blast qed qed qed qed theorem Ramsey: fixes s r :: nat and Z :: "'a set" and f :: "'a set \ nat" shows "\infinite Z; \X. X \ Z \ finite X \ card X = r \ f X < s\ \ \Y t. Y \ Z \ infinite Y \ t < s \ (\X. X \ Y \ finite X \ card X = r \ f X = t)" by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def]) corollary Ramsey2: fixes s :: nat and Z :: "'a set" and f :: "'a set \ nat" assumes infZ: "infinite Z" and part: "\x\Z. \y\Z. x \ y \ f {x, y} < s" shows "\Y t. Y \ Z \ infinite Y \ t < s \ (\x\Y. \y\Y. x\y \ f {x, y} = t)" proof - from part have part2: "\X. X \ Z \ finite X \ card X = 2 \ f X < s" by (fastforce simp add: eval_nat_numeral card_Suc_eq) obtain Y t where *: "Y \ Z" "infinite Y" "t < s" "(\X. X \ Y \ finite X \ card X = 2 \ f X = t)" by (insert Ramsey [OF infZ part2]) auto then have "\x\Y. \y\Y. x \ y \ f {x, y} = t" by auto with * show ?thesis by iprover qed subsection \Disjunctive Well-Foundedness\ text \ An application of Ramsey's theorem to program termination. See @{cite "Podelski-Rybalchenko"}. \ definition disj_wf :: "('a \ 'a) set \ bool" where "disj_wf r \ (\T. \n::nat. (\i r = (\i 'a) \ (nat \ ('a \ 'a) set) \ nat set \ nat" where "transition_idx s T A = (LEAST k. \i j. A = {i, j} \ i < j \ (s j, s i) \ T k)" lemma transition_idx_less: assumes "i < j" "(s j, s i) \ T k" "k < n" shows "transition_idx s T {i, j} < n" proof - from assms(1,2) have "transition_idx s T {i, j} \ k" by (simp add: transition_idx_def, blast intro: Least_le) with assms(3) show ?thesis by simp qed lemma transition_idx_in: assumes "i < j" "(s j, s i) \ T k" shows "(s j, s i) \ T (transition_idx s T {i, j})" using assms by (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR cong: conj_cong) (erule LeastI) text \To be equal to the union of some well-founded relations is equivalent to being the subset of such a union.\ lemma disj_wf: "disj_wf r \ (\T. \n::nat. (\i r \ (\iT n. \\i \ (T ` {.. \ (\i r)) \ r = (\i r)" by (force simp add: wf_Int1) show ?thesis unfolding disj_wf_def by auto (metis "*") qed theorem trans_disj_wf_implies_wf: assumes "trans r" and "disj_wf r" shows "wf r" proof (simp only: wf_iff_no_infinite_down_chain, rule notI) assume "\s. \i. (s (Suc i), s i) \ r" then obtain s where sSuc: "\i. (s (Suc i), s i) \ r" .. from \disj_wf r\ obtain T and n :: nat where wfT: "\kkk. (s j, s i) \ T k \ ki < j\ have "(s j, s i) \ r" proof (induct rule: less_Suc_induct) case 1 then show ?case by (simp add: sSuc) next case 2 with \trans r\ show ?case unfolding trans_def by blast qed then show ?thesis by (auto simp add: r) qed have "i < j \ transition_idx s T {i, j} < n" for i j using s_in_T transition_idx_less by blast then have trless: "i \ j \ transition_idx s T {i, j} < n" for i j by (metis doubleton_eq_iff less_linear) have "\K k. K \ UNIV \ infinite K \ k < n \ (\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k)" by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat) then obtain K and k where infK: "infinite K" and "k < n" and allk: "\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k" by auto have "(s (enumerate K (Suc m)), s (enumerate K m)) \ T k" for m :: nat proof - let ?j = "enumerate K (Suc m)" let ?i = "enumerate K m" have ij: "?i < ?j" by (simp add: enumerate_step infK) have "?j \ K" "?i \ K" by (simp_all add: enumerate_in_set infK) with ij have k: "k = transition_idx s T {?i, ?j}" by (simp add: allk) from s_in_T [OF ij] obtain k' where "(s ?j, s ?i) \ T k'" "k' T k" by (simp add: k transition_idx_in ij) qed then have "\ wf (T k)" by (meson wf_iff_no_infinite_down_chain) with wfT \k < n\ show False by blast qed end diff --git a/src/HOL/Num.thy b/src/HOL/Num.thy --- a/src/HOL/Num.thy +++ b/src/HOL/Num.thy @@ -1,1435 +1,1438 @@ (* Title: HOL/Num.thy Author: Florian Haftmann Author: Brian Huffman *) section \Binary Numerals\ theory Num imports BNF_Least_Fixpoint Transfer begin subsection \The \num\ type\ datatype num = One | Bit0 num | Bit1 num text \Increment function for type \<^typ>\num\\ primrec inc :: "num \ num" where "inc One = Bit0 One" | "inc (Bit0 x) = Bit1 x" | "inc (Bit1 x) = Bit0 (inc x)" text \Converting between type \<^typ>\num\ and type \<^typ>\nat\\ primrec nat_of_num :: "num \ nat" where "nat_of_num One = Suc 0" | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" primrec num_of_nat :: "nat \ num" where "num_of_nat 0 = One" | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" lemma nat_of_num_pos: "0 < nat_of_num x" by (induct x) simp_all lemma nat_of_num_neq_0: " nat_of_num x \ 0" by (induct x) simp_all lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" by (induct x) simp_all lemma num_of_nat_double: "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" by (induct n) simp_all text \Type \<^typ>\num\ is isomorphic to the strictly positive natural numbers.\ lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" by (induct n) (simp_all add: nat_of_num_inc) lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" apply safe apply (drule arg_cong [where f=num_of_nat]) apply (simp add: nat_of_num_inverse) done lemma num_induct [case_names One inc]: fixes P :: "num \ bool" assumes One: "P One" and inc: "\x. P x \ P (inc x)" shows "P x" proof - obtain n where n: "Suc n = nat_of_num x" by (cases "nat_of_num x") (simp_all add: nat_of_num_neq_0) have "P (num_of_nat (Suc n))" proof (induct n) case 0 from One show ?case by simp next case (Suc n) then have "P (inc (num_of_nat (Suc n)))" by (rule inc) then show "P (num_of_nat (Suc (Suc n)))" by simp qed with n show "P x" by (simp add: nat_of_num_inverse) qed text \ From now on, there are two possible models for \<^typ>\num\: as positive naturals (rule \num_induct\) and as digit representation (rules \num.induct\, \num.cases\). \ subsection \Numeral operations\ instantiation num :: "{plus,times,linorder}" begin definition [code del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)" definition [code del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)" definition [code del]: "m \ n \ nat_of_num m \ nat_of_num n" definition [code del]: "m < n \ nat_of_num m < nat_of_num n" instance by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) end lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" unfolding plus_num_def by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" unfolding times_num_def by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) lemma add_num_simps [simp, code]: "One + One = Bit0 One" "One + Bit0 n = Bit1 n" "One + Bit1 n = Bit0 (n + One)" "Bit0 m + One = Bit1 m" "Bit0 m + Bit0 n = Bit0 (m + n)" "Bit0 m + Bit1 n = Bit1 (m + n)" "Bit1 m + One = Bit0 (m + One)" "Bit1 m + Bit0 n = Bit1 (m + n)" "Bit1 m + Bit1 n = Bit0 (m + n + One)" by (simp_all add: num_eq_iff nat_of_num_add) lemma mult_num_simps [simp, code]: "m * One = m" "One * n = n" "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult distrib_right distrib_left) lemma eq_num_simps: "One = One \ True" "One = Bit0 n \ False" "One = Bit1 n \ False" "Bit0 m = One \ False" "Bit1 m = One \ False" "Bit0 m = Bit0 n \ m = n" "Bit0 m = Bit1 n \ False" "Bit1 m = Bit0 n \ False" "Bit1 m = Bit1 n \ m = n" by simp_all lemma le_num_simps [simp, code]: "One \ n \ True" "Bit0 m \ One \ False" "Bit1 m \ One \ False" "Bit0 m \ Bit0 n \ m \ n" "Bit0 m \ Bit1 n \ m \ n" "Bit1 m \ Bit1 n \ m \ n" "Bit1 m \ Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma less_num_simps [simp, code]: "m < One \ False" "One < Bit0 n \ True" "One < Bit1 n \ True" "Bit0 m < Bit0 n \ m < n" "Bit0 m < Bit1 n \ m \ n" "Bit1 m < Bit1 n \ m < n" "Bit1 m < Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma le_num_One_iff: "x \ num.One \ x = num.One" by (simp add: antisym_conv) text \Rules using \One\ and \inc\ as constructors.\ lemma add_One: "x + One = inc x" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma add_One_commute: "One + n = n + One" by (induct n) simp_all lemma add_inc: "x + inc y = inc (x + y)" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma mult_inc: "x * inc y = x * y + x" by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) text \The \<^const>\num_of_nat\ conversion.\ lemma num_of_nat_One: "n \ 1 \ num_of_nat n = One" by (cases n) simp_all lemma num_of_nat_plus_distrib: "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" by (induct n) (auto simp add: add_One add_One_commute add_inc) text \A double-and-decrement function.\ primrec BitM :: "num \ num" where "BitM One = One" | "BitM (Bit0 n) = Bit1 (BitM n)" | "BitM (Bit1 n) = Bit1 (Bit0 n)" lemma BitM_plus_one: "BitM n + One = Bit0 n" by (induct n) simp_all lemma one_plus_BitM: "One + BitM n = Bit0 n" unfolding add_One_commute BitM_plus_one .. text \Squaring and exponentiation.\ primrec sqr :: "num \ num" where "sqr One = One" | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" primrec pow :: "num \ num \ num" where "pow x One = x" | "pow x (Bit0 y) = sqr (pow x y)" | "pow x (Bit1 y) = sqr (pow x y) * x" lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" by (induct x) (simp_all add: algebra_simps nat_of_num_add) lemma sqr_conv_mult: "sqr x = x * x" by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) lemma num_double [simp]: "num.Bit0 num.One * n = num.Bit0 n" by (simp add: num_eq_iff nat_of_num_mult) subsection \Binary numerals\ text \ We embed binary representations into a generic algebraic structure using \numeral\. \ class numeral = one + semigroup_add begin primrec numeral :: "num \ 'a" where numeral_One: "numeral One = 1" | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" lemma numeral_code [code]: "numeral One = 1" "numeral (Bit0 n) = (let m = numeral n in m + m)" "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" by (simp_all add: Let_def) lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) next case Bit1 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) qed lemma numeral_inc: "numeral (inc x) = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by simp next case (Bit1 x) have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case by (simp add: add.assoc) qed declare numeral.simps [simp del] abbreviation "Numeral1 \ numeral One" declare numeral_One [code_post] end text \Numeral syntax.\ syntax "_Numeral" :: "num_const \ 'a" ("_") ML_file \Tools/numeral.ML\ parse_translation \ let fun numeral_tr [(c as Const (\<^syntax_const>\_constrain\, _)) $ t $ u] = c $ numeral_tr [t] $ u | numeral_tr [Const (num, _)] = (Numeral.mk_number_syntax o #value o Lexicon.read_num) num | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_Numeral\, K numeral_tr)] end \ typed_print_translation \ let fun num_tr' ctxt T [n] = let val k = Numeral.dest_num_syntax n; val t' = Syntax.const \<^syntax_const>\_Numeral\ $ Syntax.free (string_of_int k); in (case T of Type (\<^type_name>\fun\, [_, T']) => if Printer.type_emphasis ctxt T' then Syntax.const \<^syntax_const>\_constrain\ $ t' $ Syntax_Phases.term_of_typ ctxt T' else t' | _ => if T = dummyT then t' else raise Match) end; in [(\<^const_syntax>\numeral\, num_tr')] end \ subsection \Class-specific numeral rules\ text \\<^const>\numeral\ is a morphism.\ subsubsection \Structures with addition: class \numeral\\ context numeral begin lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" using numeral_add [of n One] by (simp add: numeral_One) lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" using numeral_add [of One n] by (simp add: numeral_One) lemma one_add_one: "1 + 1 = 2" using numeral_add [of One One] by (simp add: numeral_One) lemmas add_numeral_special = numeral_plus_one one_plus_numeral one_add_one end subsubsection \Structures with negation: class \neg_numeral\\ class neg_numeral = numeral + group_add begin lemma uminus_numeral_One: "- Numeral1 = - 1" by (simp add: numeral_One) text \Numerals form an abelian subgroup.\ inductive is_num :: "'a \ bool" where "is_num 1" | "is_num x \ is_num (- x)" | "is_num x \ is_num y \ is_num (x + y)" lemma is_num_numeral: "is_num (numeral k)" by (induct k) (simp_all add: numeral.simps is_num.intros) lemma is_num_add_commute: "is_num x \ is_num y \ x + y = y + x" apply (induct x rule: is_num.induct) apply (induct y rule: is_num.induct) apply simp apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) apply (simp add: add.assoc) apply (simp add: add.assoc [symmetric]) apply (simp add: add.assoc) apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) apply (simp add: add.assoc) apply (simp add: add.assoc) apply (simp add: add.assoc [symmetric]) done lemma is_num_add_left_commute: "is_num x \ is_num y \ x + (y + z) = y + (x + z)" by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = add.assoc is_num_add_commute is_num_add_left_commute is_num.intros is_num_numeral minus_add definition dbl :: "'a \ 'a" where "dbl x = x + x" definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" definition sub :: "num \ num \ 'a" where "sub k l = numeral k - numeral l" lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) lemma dbl_simps [simp]: "dbl (- numeral k) = - dbl (numeral k)" "dbl 0 = 0" "dbl 1 = 2" "dbl (- 1) = - 2" "dbl (numeral k) = numeral (Bit0 k)" by (simp_all add: dbl_def numeral.simps minus_add) lemma dbl_inc_simps [simp]: "dbl_inc (- numeral k) = - dbl_dec (numeral k)" "dbl_inc 0 = 1" "dbl_inc 1 = 3" "dbl_inc (- 1) = - 1" "dbl_inc (numeral k) = numeral (Bit1 k)" by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) lemma dbl_dec_simps [simp]: "dbl_dec (- numeral k) = - dbl_inc (numeral k)" "dbl_dec 0 = - 1" "dbl_dec 1 = 1" "dbl_dec (- 1) = - 3" "dbl_dec (numeral k) = numeral (BitM k)" by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize) lemma sub_num_simps [simp]: "sub One One = 0" "sub One (Bit0 l) = - numeral (BitM l)" "sub One (Bit1 l) = - numeral (Bit0 l)" "sub (Bit0 k) One = numeral (BitM k)" "sub (Bit1 k) One = numeral (Bit0 k)" "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_simps: "numeral m + - numeral n = sub m n" "- numeral m + numeral n = sub n m" "- numeral m + - numeral n = - (numeral m + numeral n)" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_special: "1 + - numeral m = sub One m" "- numeral m + 1 = sub One m" "numeral m + - 1 = sub m One" "- 1 + numeral n = sub n One" "- 1 + - numeral n = - numeral (inc n)" "- numeral m + - 1 = - numeral (inc m)" "1 + - 1 = 0" "- 1 + 1 = 0" "- 1 + - 1 = - 2" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_simps: "numeral m - numeral n = sub m n" "numeral m - - numeral n = numeral (m + n)" "- numeral m - numeral n = - numeral (m + n)" "- numeral m - - numeral n = sub n m" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_special: "1 - numeral n = sub One n" "numeral m - 1 = sub m One" "1 - - numeral n = numeral (One + n)" "- numeral m - 1 = - numeral (m + One)" "- 1 - numeral n = - numeral (inc n)" "numeral m - - 1 = numeral (inc m)" "- 1 - - numeral n = sub n One" "- numeral m - - 1 = sub One m" "1 - 1 = 0" "- 1 - 1 = - 2" "1 - - 1 = 2" "- 1 - - 1 = 0" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) end subsubsection \Structures with multiplication: class \semiring_numeral\\ class semiring_numeral = semiring + monoid_mult begin subclass numeral .. lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" by (induct n rule: num_induct) (simp_all add: numeral_One mult_inc numeral_inc numeral_add distrib_left) lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" by (rule numeral_mult [symmetric]) lemma mult_2: "2 * z = z + z" by (simp add: one_add_one [symmetric] distrib_right) lemma mult_2_right: "z * 2 = z + z" by (simp add: one_add_one [symmetric] distrib_left) lemma left_add_twice: "a + (a + b) = 2 * a + b" by (simp add: mult_2 ac_simps) end subsubsection \Structures with a zero: class \semiring_1\\ context semiring_1 begin subclass semiring_numeral .. lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) lemma numeral_unfold_funpow: "numeral k = ((+) 1 ^^ numeral k) 0" unfolding of_nat_def [symmetric] by simp end context includes lifting_syntax begin lemma transfer_rule_numeral: "((=) ===> R) numeral numeral" if [transfer_rule]: "R 0 0" "R 1 1" "(R ===> R ===> R) (+) (+)" for R :: "'a::semiring_1 \ 'b::semiring_1 \ bool" proof - have "((=) ===> R) (\k. ((+) 1 ^^ numeral k) 0) (\k. ((+) 1 ^^ numeral k) 0)" by transfer_prover then show ?thesis by (simp flip: numeral_unfold_funpow [abs_def]) qed end lemma nat_of_num_numeral [code_abbrev]: "nat_of_num = numeral" proof fix n have "numeral n = nat_of_num n" by (induct n) (simp_all add: numeral.simps) then show "nat_of_num n = numeral n" by simp qed lemma nat_of_num_code [code]: "nat_of_num One = 1" "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)" "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" by (simp_all add: Let_def) subsubsection \Equality: class \semiring_char_0\\ context semiring_char_0 begin lemma numeral_eq_iff: "numeral m = numeral n \ m = n" by (simp only: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] of_nat_eq_iff num_eq_iff) lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" by (rule numeral_eq_iff [of n One, unfolded numeral_One]) lemma one_eq_numeral_iff: "1 = numeral n \ One = n" by (rule numeral_eq_iff [of One n, unfolded numeral_One]) lemma numeral_neq_zero: "numeral n \ 0" by (simp add: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] nat_of_num_pos) lemma zero_neq_numeral: "0 \ numeral n" unfolding eq_commute [of 0] by (rule numeral_neq_zero) lemmas eq_numeral_simps [simp] = numeral_eq_iff numeral_eq_one_iff one_eq_numeral_iff numeral_neq_zero zero_neq_numeral end subsubsection \Comparisons: class \linordered_nonzero_semiring\\ context linordered_nonzero_semiring begin lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" proof - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) then show ?thesis by simp qed lemma one_le_numeral: "1 \ numeral n" using numeral_le_iff [of num.One n] by (simp add: numeral_One) lemma numeral_le_one_iff: "numeral n \ 1 \ n \ num.One" using numeral_le_iff [of n num.One] by (simp add: numeral_One) lemma numeral_less_iff: "numeral m < numeral n \ m < n" proof - have "of_nat (numeral m) < of_nat (numeral n) \ m < n" unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. then show ?thesis by simp qed lemma not_numeral_less_one: "\ numeral n < 1" using numeral_less_iff [of n num.One] by (simp add: numeral_One) lemma one_less_numeral_iff: "1 < numeral n \ num.One < n" using numeral_less_iff [of num.One n] by (simp add: numeral_One) lemma zero_le_numeral: "0 \ numeral n" using dual_order.trans one_le_numeral zero_le_one by blast lemma zero_less_numeral: "0 < numeral n" using less_linear not_numeral_less_one order.strict_trans zero_less_one by blast lemma not_numeral_le_zero: "\ numeral n \ 0" by (simp add: not_le zero_less_numeral) lemma not_numeral_less_zero: "\ numeral n < 0" by (simp add: not_less zero_le_numeral) lemmas le_numeral_extra = zero_le_one not_one_le_zero order_refl [of 0] order_refl [of 1] lemmas less_numeral_extra = zero_less_one not_one_less_zero less_irrefl [of 0] less_irrefl [of 1] lemmas le_numeral_simps [simp] = numeral_le_iff one_le_numeral numeral_le_one_iff zero_le_numeral not_numeral_le_zero lemmas less_numeral_simps [simp] = numeral_less_iff one_less_numeral_iff not_numeral_less_one zero_less_numeral not_numeral_less_zero lemma min_0_1 [simp]: fixes min' :: "'a \ 'a \ 'a" defines "min' \ min" shows "min' 0 1 = 0" "min' 1 0 = 0" "min' 0 (numeral x) = 0" "min' (numeral x) 0 = 0" "min' 1 (numeral x) = 1" "min' (numeral x) 1 = 1" by (simp_all add: min'_def min_def le_num_One_iff) lemma max_0_1 [simp]: fixes max' :: "'a \ 'a \ 'a" defines "max' \ max" shows "max' 0 1 = 1" "max' 1 0 = 1" "max' 0 (numeral x) = numeral x" "max' (numeral x) 0 = numeral x" "max' 1 (numeral x) = numeral x" "max' (numeral x) 1 = numeral x" by (simp_all add: max'_def max_def le_num_One_iff) end text \Unfold \min\ and \max\ on numerals.\ lemmas max_number_of [simp] = max_def [of "numeral u" "numeral v"] max_def [of "numeral u" "- numeral v"] max_def [of "- numeral u" "numeral v"] max_def [of "- numeral u" "- numeral v"] for u v lemmas min_number_of [simp] = min_def [of "numeral u" "numeral v"] min_def [of "numeral u" "- numeral v"] min_def [of "- numeral u" "numeral v"] min_def [of "- numeral u" "- numeral v"] for u v subsubsection \Multiplication and negation: class \ring_1\\ context ring_1 begin subclass neg_numeral .. lemma mult_neg_numeral_simps: "- numeral m * - numeral n = numeral (m * n)" "- numeral m * numeral n = - numeral (m * n)" "numeral m * - numeral n = - numeral (m * n)" by (simp_all only: mult_minus_left mult_minus_right minus_minus numeral_mult) lemma mult_minus1 [simp]: "- 1 * z = - z" by (simp add: numeral.simps) lemma mult_minus1_right [simp]: "z * - 1 = - z" by (simp add: numeral.simps) end subsubsection \Equality using \iszero\ for rings with non-zero characteristic\ context ring_1 begin definition iszero :: "'a \ bool" where "iszero z \ z = 0" lemma iszero_0 [simp]: "iszero 0" by (simp add: iszero_def) lemma not_iszero_1 [simp]: "\ iszero 1" by (simp add: iszero_def) lemma not_iszero_Numeral1: "\ iszero Numeral1" by (simp add: numeral_One) lemma not_iszero_neg_1 [simp]: "\ iszero (- 1)" by (simp add: iszero_def) lemma not_iszero_neg_Numeral1: "\ iszero (- Numeral1)" by (simp add: numeral_One) lemma iszero_neg_numeral [simp]: "iszero (- numeral w) \ iszero (numeral w)" unfolding iszero_def by (rule neg_equal_0_iff_equal) lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" unfolding iszero_def by (rule eq_iff_diff_eq_0) text \ The \eq_numeral_iff_iszero\ lemmas are not declared \[simp]\ by default, because for rings of characteristic zero, better simp rules are possible. For a type like integers mod \n\, type-instantiated versions of these rules should be added to the simplifier, along with a type-specific rule for deciding propositions of the form \iszero (numeral w)\. bh: Maybe it would not be so bad to just declare these as simp rules anyway? I should test whether these rules take precedence over the \ring_char_0\ rules in the simplifier. \ lemma eq_numeral_iff_iszero: "numeral x = numeral y \ iszero (sub x y)" "numeral x = - numeral y \ iszero (numeral (x + y))" "- numeral x = numeral y \ iszero (numeral (x + y))" "- numeral x = - numeral y \ iszero (sub y x)" "numeral x = 1 \ iszero (sub x One)" "1 = numeral y \ iszero (sub One y)" "- numeral x = 1 \ iszero (numeral (x + One))" "1 = - numeral y \ iszero (numeral (One + y))" "numeral x = 0 \ iszero (numeral x)" "0 = numeral y \ iszero (numeral y)" "- numeral x = 0 \ iszero (numeral x)" "0 = - numeral y \ iszero (numeral y)" unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special by simp_all end subsubsection \Equality and negation: class \ring_char_0\\ context ring_char_0 begin lemma not_iszero_numeral [simp]: "\ iszero (numeral w)" by (simp add: iszero_def) lemma neg_numeral_eq_iff: "- numeral m = - numeral n \ m = n" by simp lemma numeral_neq_neg_numeral: "numeral m \ - numeral n" by (simp add: eq_neg_iff_add_eq_0 numeral_plus_numeral) lemma neg_numeral_neq_numeral: "- numeral m \ numeral n" by (rule numeral_neq_neg_numeral [symmetric]) lemma zero_neq_neg_numeral: "0 \ - numeral n" by simp lemma neg_numeral_neq_zero: "- numeral n \ 0" by simp lemma one_neq_neg_numeral: "1 \ - numeral n" using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) lemma neg_numeral_neq_one: "- numeral n \ 1" using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) lemma neg_one_neq_numeral: "- 1 \ numeral n" using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) lemma numeral_neq_neg_one: "numeral n \ - 1" using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) lemma neg_one_eq_numeral_iff: "- 1 = - numeral n \ n = One" using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) lemma numeral_eq_neg_one_iff: "- numeral n = - 1 \ n = One" using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) lemma neg_one_neq_zero: "- 1 \ 0" by simp lemma zero_neq_neg_one: "0 \ - 1" by simp lemma neg_one_neq_one: "- 1 \ 1" using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemma one_neq_neg_one: "1 \ - 1" using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemmas eq_neg_numeral_simps [simp] = neg_numeral_eq_iff numeral_neq_neg_numeral neg_numeral_neq_numeral one_neq_neg_numeral neg_numeral_neq_one zero_neq_neg_numeral neg_numeral_neq_zero neg_one_neq_numeral numeral_neq_neg_one neg_one_eq_numeral_iff numeral_eq_neg_one_iff neg_one_neq_zero zero_neq_neg_one neg_one_neq_one one_neq_neg_one end subsubsection \Structures with negation and order: class \linordered_idom\\ context linordered_idom begin subclass ring_char_0 .. lemma neg_numeral_le_iff: "- numeral m \ - numeral n \ n \ m" by (simp only: neg_le_iff_le numeral_le_iff) lemma neg_numeral_less_iff: "- numeral m < - numeral n \ n < m" by (simp only: neg_less_iff_less numeral_less_iff) lemma neg_numeral_less_zero: "- numeral n < 0" by (simp only: neg_less_0_iff_less zero_less_numeral) lemma neg_numeral_le_zero: "- numeral n \ 0" by (simp only: neg_le_0_iff_le zero_le_numeral) lemma not_zero_less_neg_numeral: "\ 0 < - numeral n" by (simp only: not_less neg_numeral_le_zero) lemma not_zero_le_neg_numeral: "\ 0 \ - numeral n" by (simp only: not_le neg_numeral_less_zero) lemma neg_numeral_less_numeral: "- numeral m < numeral n" using neg_numeral_less_zero zero_less_numeral by (rule less_trans) lemma neg_numeral_le_numeral: "- numeral m \ numeral n" by (simp only: less_imp_le neg_numeral_less_numeral) lemma not_numeral_less_neg_numeral: "\ numeral m < - numeral n" by (simp only: not_less neg_numeral_le_numeral) lemma not_numeral_le_neg_numeral: "\ numeral m \ - numeral n" by (simp only: not_le neg_numeral_less_numeral) lemma neg_numeral_less_one: "- numeral m < 1" by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) lemma neg_numeral_le_one: "- numeral m \ 1" by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) lemma not_one_less_neg_numeral: "\ 1 < - numeral m" by (simp only: not_less neg_numeral_le_one) lemma not_one_le_neg_numeral: "\ 1 \ - numeral m" by (simp only: not_le neg_numeral_less_one) lemma not_numeral_less_neg_one: "\ numeral m < - 1" using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One) lemma not_numeral_le_neg_one: "\ numeral m \ - 1" using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One) lemma neg_one_less_numeral: "- 1 < numeral m" using neg_numeral_less_numeral [of One m] by (simp add: numeral_One) lemma neg_one_le_numeral: "- 1 \ numeral m" using neg_numeral_le_numeral [of One m] by (simp add: numeral_One) lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \ m \ One" by (cases m) simp_all lemma neg_numeral_le_neg_one: "- numeral m \ - 1" by simp lemma not_neg_one_less_neg_numeral: "\ - 1 < - numeral m" by simp lemma not_neg_one_le_neg_numeral_iff: "\ - 1 \ - numeral m \ m \ One" by (cases m) simp_all lemma sub_non_negative: "sub n m \ 0 \ n \ m" by (simp only: sub_def le_diff_eq) simp lemma sub_positive: "sub n m > 0 \ n > m" by (simp only: sub_def less_diff_eq) simp lemma sub_non_positive: "sub n m \ 0 \ n \ m" by (simp only: sub_def diff_le_eq) simp lemma sub_negative: "sub n m < 0 \ n < m" by (simp only: sub_def diff_less_eq) simp lemmas le_neg_numeral_simps [simp] = neg_numeral_le_iff neg_numeral_le_numeral not_numeral_le_neg_numeral neg_numeral_le_zero not_zero_le_neg_numeral neg_numeral_le_one not_one_le_neg_numeral neg_one_le_numeral not_numeral_le_neg_one neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff lemma le_minus_one_simps [simp]: "- 1 \ 0" "- 1 \ 1" "\ 0 \ - 1" "\ 1 \ - 1" by simp_all lemmas less_neg_numeral_simps [simp] = neg_numeral_less_iff neg_numeral_less_numeral not_numeral_less_neg_numeral neg_numeral_less_zero not_zero_less_neg_numeral neg_numeral_less_one not_one_less_neg_numeral neg_one_less_numeral not_numeral_less_neg_one neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral lemma less_minus_one_simps [simp]: "- 1 < 0" "- 1 < 1" "\ 0 < - 1" "\ 1 < - 1" by (simp_all add: less_le) lemma abs_numeral [simp]: "\numeral n\ = numeral n" by simp lemma abs_neg_numeral [simp]: "\- numeral n\ = numeral n" by (simp only: abs_minus_cancel abs_numeral) lemma abs_neg_one [simp]: "\- 1\ = 1" by simp end subsubsection \Natural numbers\ lemma numeral_num_of_nat: "numeral (num_of_nat n) = n" if "n > 0" using that nat_of_num_numeral num_of_nat_inverse by simp lemma Suc_1 [simp]: "Suc 1 = 2" unfolding Suc_eq_plus1 by (rule one_add_one) lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" unfolding Suc_eq_plus1 by (rule numeral_plus_one) definition pred_numeral :: "num \ nat" where "pred_numeral k = numeral k - 1" declare [[code drop: pred_numeral]] lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" by (simp add: pred_numeral_def) lemma eval_nat_numeral: "numeral One = Suc 0" "numeral (Bit0 n) = Suc (numeral (BitM n))" "numeral (Bit1 n) = Suc (numeral (Bit0 n))" by (simp_all add: numeral.simps BitM_plus_one) lemma pred_numeral_simps [simp]: "pred_numeral One = 0" "pred_numeral (Bit0 k) = numeral (BitM k)" "pred_numeral (Bit1 k) = numeral (Bit0 k)" by (simp_all only: pred_numeral_def eval_nat_numeral diff_Suc_Suc diff_0) lemma pred_numeral_inc [simp]: "pred_numeral (Num.inc k) = numeral k" by (simp only: pred_numeral_def numeral_inc diff_add_inverse2) lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (simp add: eval_nat_numeral) lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" by (simp only: numeral_One One_nat_def) lemma Suc_nat_number_of_add: "Suc (numeral v + n) = numeral (v + One) + n" by simp lemma numerals: "Numeral1 = (1::nat)" "2 = Suc (Suc 0)" by (rule numeral_One) (rule numeral_2_eq_2) lemmas numeral_nat = eval_nat_numeral BitM.simps One_nat_def text \Comparisons involving \<^term>\Suc\.\ lemma eq_numeral_Suc [simp]: "numeral k = Suc n \ pred_numeral k = n" by (simp add: numeral_eq_Suc) lemma Suc_eq_numeral [simp]: "Suc n = numeral k \ n = pred_numeral k" by (simp add: numeral_eq_Suc) lemma less_numeral_Suc [simp]: "numeral k < Suc n \ pred_numeral k < n" by (simp add: numeral_eq_Suc) lemma less_Suc_numeral [simp]: "Suc n < numeral k \ n < pred_numeral k" by (simp add: numeral_eq_Suc) lemma le_numeral_Suc [simp]: "numeral k \ Suc n \ pred_numeral k \ n" by (simp add: numeral_eq_Suc) lemma le_Suc_numeral [simp]: "Suc n \ numeral k \ n \ pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_Suc_numeral [simp]: "Suc n - numeral k = n - pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" by (simp add: numeral_eq_Suc) lemma max_Suc_numeral [simp]: "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma max_numeral_Suc [simp]: "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" by (simp add: numeral_eq_Suc) lemma min_Suc_numeral [simp]: "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma min_numeral_Suc [simp]: "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" by (simp add: numeral_eq_Suc) text \For \<^term>\case_nat\ and \<^term>\rec_nat\.\ lemma case_nat_numeral [simp]: "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" by (simp add: numeral_eq_Suc) lemma case_nat_add_eq_if [simp]: "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" by (simp add: numeral_eq_Suc) lemma rec_nat_numeral [simp]: "rec_nat a f (numeral v) = (let pv = pred_numeral v in f pv (rec_nat a f pv))" by (simp add: numeral_eq_Suc Let_def) lemma rec_nat_add_eq_if [simp]: "rec_nat a f (numeral v + n) = (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" by (simp add: numeral_eq_Suc Let_def) text \Case analysis on \<^term>\n < 2\.\ lemma less_2_cases: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) +lemma less_2_cases_iff: "n < 2 \ n = 0 \ n = Suc 0" + by (auto simp add: numeral_2_eq_2) + text \Removal of Small Numerals: 0, 1 and (in additive positions) 2.\ -text \bh: Are these rules really a good idea?\ +text \bh: Are these rules really a good idea? LCP: well, it already happens for 0 and 1!\ lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" by simp lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" by simp text \Can be used to eliminate long strings of Sucs, but not by default.\ lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *) subsection \Particular lemmas concerning \<^term>\2\\ context linordered_field begin subclass field_char_0 .. lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" by (auto simp add: field_simps) lemma half_gt_zero [simp]: "0 < a \ 0 < a / 2" by (simp add: half_gt_zero_iff) end subsection \Numeral equations as default simplification rules\ declare (in numeral) numeral_One [simp] declare (in numeral) numeral_plus_numeral [simp] declare (in numeral) add_numeral_special [simp] declare (in neg_numeral) add_neg_numeral_simps [simp] declare (in neg_numeral) add_neg_numeral_special [simp] declare (in neg_numeral) diff_numeral_simps [simp] declare (in neg_numeral) diff_numeral_special [simp] declare (in semiring_numeral) numeral_times_numeral [simp] declare (in ring_1) mult_neg_numeral_simps [simp] subsubsection \Special Simplification for Constants\ text \These distributive laws move literals inside sums and differences.\ lemmas distrib_right_numeral [simp] = distrib_right [of _ _ "numeral v"] for v lemmas distrib_left_numeral [simp] = distrib_left [of "numeral v"] for v lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v text \These are actually for fields, like real\ lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w lemmas zero_le_divide_iff_numeral [simp, no_atp] = zero_le_divide_iff [of "numeral w"] for w lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w text \Replaces \inverse #nn\ by \1/#nn\. It looks strange, but then other simprocs simplify the quotient.\ lemmas inverse_eq_divide_numeral [simp] = inverse_eq_divide [of "numeral w"] for w lemmas inverse_eq_divide_neg_numeral [simp] = inverse_eq_divide [of "- numeral w"] for w text \These laws simplify inequalities, moving unary minus from a term into the literal.\ lemmas equation_minus_iff_numeral [no_atp] = equation_minus_iff [of "numeral v"] for v lemmas minus_equation_iff_numeral [no_atp] = minus_equation_iff [of _ "numeral v"] for v lemmas le_minus_iff_numeral [no_atp] = le_minus_iff [of "numeral v"] for v lemmas minus_le_iff_numeral [no_atp] = minus_le_iff [of _ "numeral v"] for v lemmas less_minus_iff_numeral [no_atp] = less_minus_iff [of "numeral v"] for v lemmas minus_less_iff_numeral [no_atp] = minus_less_iff [of _ "numeral v"] for v (* FIXME maybe simproc *) text \Cancellation of constant factors in comparisons (\<\ and \\\)\ lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v lemmas mult_le_cancel_left_numeral [simp, no_atp] = mult_le_cancel_left [of "numeral v"] for v lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v text \Multiplying out constant divisors in comparisons (\<\, \\\ and \=\)\ named_theorems divide_const_simps "simplification rules to simplify comparisons involving constant divisors" lemmas le_divide_eq_numeral1 [simp,divide_const_simps] = pos_le_divide_eq [of "numeral w", OF zero_less_numeral] neg_le_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_le_eq_numeral1 [simp,divide_const_simps] = pos_divide_le_eq [of "numeral w", OF zero_less_numeral] neg_divide_le_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas less_divide_eq_numeral1 [simp,divide_const_simps] = pos_less_divide_eq [of "numeral w", OF zero_less_numeral] neg_less_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_less_eq_numeral1 [simp,divide_const_simps] = pos_divide_less_eq [of "numeral w", OF zero_less_numeral] neg_divide_less_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas eq_divide_eq_numeral1 [simp,divide_const_simps] = eq_divide_eq [of _ _ "numeral w"] eq_divide_eq [of _ _ "- numeral w"] for w lemmas divide_eq_eq_numeral1 [simp,divide_const_simps] = divide_eq_eq [of _ "numeral w"] divide_eq_eq [of _ "- numeral w"] for w subsubsection \Optional Simplification Rules Involving Constants\ text \Simplify quotients that are compared with a literal constant.\ lemmas le_divide_eq_numeral [divide_const_simps] = le_divide_eq [of "numeral w"] le_divide_eq [of "- numeral w"] for w lemmas divide_le_eq_numeral [divide_const_simps] = divide_le_eq [of _ _ "numeral w"] divide_le_eq [of _ _ "- numeral w"] for w lemmas less_divide_eq_numeral [divide_const_simps] = less_divide_eq [of "numeral w"] less_divide_eq [of "- numeral w"] for w lemmas divide_less_eq_numeral [divide_const_simps] = divide_less_eq [of _ _ "numeral w"] divide_less_eq [of _ _ "- numeral w"] for w lemmas eq_divide_eq_numeral [divide_const_simps] = eq_divide_eq [of "numeral w"] eq_divide_eq [of "- numeral w"] for w lemmas divide_eq_eq_numeral [divide_const_simps] = divide_eq_eq [of _ _ "numeral w"] divide_eq_eq [of _ _ "- numeral w"] for w text \Not good as automatic simprules because they cause case splits.\ lemmas [divide_const_simps] = le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 subsection \Setting up simprocs\ lemma mult_numeral_1: "Numeral1 * a = a" for a :: "'a::semiring_numeral" by simp lemma mult_numeral_1_right: "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp lemma divide_numeral_1: "a / Numeral1 = a" for a :: "'a::field" by simp lemma inverse_numeral_1: "inverse Numeral1 = (Numeral1::'a::division_ring)" by simp text \ Theorem lists for the cancellation simprocs. The use of a binary numeral for 1 reduces the number of special cases. \ lemma mult_1s_semiring_numeral: "Numeral1 * a = a" "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp_all lemma mult_1s_ring_1: "- Numeral1 * b = - b" "b * - Numeral1 = - b" for b :: "'a::ring_1" by simp_all lemmas mult_1s = mult_1s_semiring_numeral mult_1s_ring_1 setup \ Reorient_Proc.add (fn Const (\<^const_name>\numeral\, _) $ _ => true | Const (\<^const_name>\uminus\, _) $ (Const (\<^const_name>\numeral\, _) $ _) => true | _ => false) \ simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc subsubsection \Simplification of arithmetic operations on integer constants\ lemmas arith_special = (* already declared simp above *) add_numeral_special add_neg_numeral_special diff_numeral_special lemmas arith_extra_simps = (* rules already in simpset *) numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right minus_zero diff_numeral_simps diff_0 diff_0_right numeral_times_numeral mult_neg_numeral_simps mult_zero_left mult_zero_right abs_numeral abs_neg_numeral text \ For making a minimal simpset, one must include these default simprules. Also include \simp_thms\. \ lemmas arith_simps = add_num_simps mult_num_simps sub_num_simps BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps abs_zero abs_one arith_extra_simps lemmas more_arith_simps = neg_le_iff_le minus_zero left_minus right_minus mult_1_left mult_1_right mult_minus_left mult_minus_right minus_add_distrib minus_minus mult.assoc lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult text \Simplification of relational operations.\ lemmas eq_numeral_extra = zero_neq_one one_neq_zero lemmas rel_simps = le_num_simps less_num_simps eq_num_simps le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. declaration \ let fun number_of ctxt T n = if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\numeral\)) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; in K ( Lin_Arith.set_number_of number_of #> Lin_Arith.add_simps @{thms arith_simps more_arith_simps rel_simps pred_numeral_simps arith_special numeral_One of_nat_simps uminus_numeral_One Suc_numeral Let_numeral Let_neg_numeral Let_0 Let_1 le_Suc_numeral le_numeral_Suc less_Suc_numeral less_numeral_Suc Suc_eq_numeral eq_numeral_Suc mult_Suc mult_Suc_right of_nat_numeral}) end \ subsubsection \Simplification of arithmetic when nested to the right\ lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: "numeral v + (- numeral w + y) = (sub v w + y)" "- numeral v + (numeral w + y) = (sub w v + y)" "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" by (simp_all add: add.assoc [symmetric]) lemma mult_numeral_left_semiring_numeral: "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" by (simp add: mult.assoc [symmetric]) lemma mult_numeral_left_ring_1: "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'a::ring_1)" by (simp_all add: mult.assoc [symmetric]) lemmas mult_numeral_left [simp] = mult_numeral_left_semiring_numeral mult_numeral_left_ring_1 hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec subsection \Code module namespace\ code_identifier code_module Num \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Printing of evaluated natural numbers as numerals\ lemma [code_post]: "Suc 0 = 1" "Suc 1 = 2" "Suc (numeral n) = numeral (Num.inc n)" by (simp_all add: numeral_inc) lemmas [code_post] = Num.inc.simps end