| author | berghofe | 
| Wed, 07 Feb 2007 17:41:11 +0100 | |
| changeset 22270 | 4ccb7e6be929 | 
| parent 22045 | ce5daf09ebfe | 
| child 22803 | 5129e02f4df2 | 
| permissions | -rw-r--r-- | 
| 14128 | 1 | (* Title: HOL/NatSimprocs.thy | 
| 2 | ID: $Id$ | |
| 3 | Copyright 2003 TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | header {*Simprocs for the Naturals*}
 | |
| 7 | ||
| 15131 | 8 | theory NatSimprocs | 
| 15140 | 9 | imports NatBin | 
| 16417 | 10 | uses "int_factor_simprocs.ML" "nat_simprocs.ML" | 
| 15131 | 11 | begin | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
8858diff
changeset | 12 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
8858diff
changeset | 13 | setup nat_simprocs_setup | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: 
8858diff
changeset | 14 | |
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 15 | subsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 16 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 17 | text{*Where K above is a literal*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 18 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 19 | lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 20 | by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 21 | |
| 14577 | 22 | text {*Now just instantiating @{text n} to @{text "number_of v"} does
 | 
| 23 | the right simplification, but with some redundant inequality | |
| 24 | tests.*} | |
| 20485 | 25 | lemma neg_number_of_pred_iff_0: | 
| 20500 | 26 | "neg (number_of (Numeral.pred v)::int) = (number_of v = (0::nat))" | 
| 27 | apply (subgoal_tac "neg (number_of (Numeral.pred v)) = (number_of v < Suc 0) ") | |
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 28 | apply (simp only: less_Suc_eq_le le_0_eq) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 29 | apply (subst less_number_of_Suc, simp) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 30 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 31 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 32 | text{*No longer required as a simprule because of the @{text inverse_fold}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 33 | simproc*} | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 34 | lemma Suc_diff_number_of: | 
| 20485 | 35 | "neg (number_of (uminus v)::int) ==> | 
| 20500 | 36 | Suc m - (number_of v) = m - (number_of (Numeral.pred v))" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 37 | apply (subst Suc_diff_eq_diff_pred) | 
| 18624 | 38 | apply simp | 
| 39 | apply (simp del: nat_numeral_1_eq_1) | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 40 | apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] | 
| 20485 | 41 | neg_number_of_pred_iff_0) | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 42 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 43 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 44 | lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 45 | by (simp add: numerals split add: nat_diff_split) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 46 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 47 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 48 | subsection{*For @{term nat_case} and @{term nat_rec}*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 49 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 50 | lemma nat_case_number_of [simp]: | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 51 | "nat_case a f (number_of v) = | 
| 20500 | 52 | (let pv = number_of (Numeral.pred v) in | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 53 | if neg pv then a else f (nat pv))" | 
| 20485 | 54 | by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0) | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 55 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 56 | lemma nat_case_add_eq_if [simp]: | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 57 | "nat_case a f ((number_of v) + n) = | 
| 20500 | 58 | (let pv = number_of (Numeral.pred v) in | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 59 | if neg pv then nat_case a f n else f (nat pv + n))" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 60 | apply (subst add_eq_if) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 61 | apply (simp split add: nat.split | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 62 | del: nat_numeral_1_eq_1 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 63 | add: numeral_1_eq_Suc_0 [symmetric] Let_def | 
| 20485 | 64 | neg_imp_number_of_eq_0 neg_number_of_pred_iff_0) | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 65 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 66 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 67 | lemma nat_rec_number_of [simp]: | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 68 | "nat_rec a f (number_of v) = | 
| 20500 | 69 | (let pv = number_of (Numeral.pred v) in | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 70 | if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 71 | apply (case_tac " (number_of v) ::nat") | 
| 20485 | 72 | apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0) | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 73 | apply (simp split add: split_if_asm) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 74 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 75 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 76 | lemma nat_rec_add_eq_if [simp]: | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 77 | "nat_rec a f (number_of v + n) = | 
| 20500 | 78 | (let pv = number_of (Numeral.pred v) in | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 79 | if neg pv then nat_rec a f n | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 80 | else f (nat pv + n) (nat_rec a f (nat pv + n)))" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 81 | apply (subst add_eq_if) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 82 | apply (simp split add: nat.split | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 83 | del: nat_numeral_1_eq_1 | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 84 | add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0 | 
| 20485 | 85 | neg_number_of_pred_iff_0) | 
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 86 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 87 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 88 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 89 | subsection{*Various Other Lemmas*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 90 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 91 | subsubsection{*Evens and Odds, for Mutilated Chess Board*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 92 | |
| 14436 | 93 | text{*Lemmas for specialist use, NOT as default simprules*}
 | 
| 94 | lemma nat_mult_2: "2 * z = (z+z::nat)" | |
| 95 | proof - | |
| 96 | have "2*z = (1 + 1)*z" by simp | |
| 97 | also have "... = z+z" by (simp add: left_distrib) | |
| 98 | finally show ?thesis . | |
| 99 | qed | |
| 100 | ||
| 101 | lemma nat_mult_2_right: "z * 2 = (z+z::nat)" | |
| 102 | by (subst mult_commute, rule nat_mult_2) | |
| 103 | ||
| 104 | text{*Case analysis on @{term "n<2"}*}
 | |
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 105 | lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 106 | by arith | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 107 | |
| 14436 | 108 | lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)" | 
| 109 | by arith | |
| 110 | ||
| 111 | lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" | |
| 112 | by (simp add: nat_mult_2 [symmetric]) | |
| 113 | ||
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 114 | lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 115 | apply (subgoal_tac "m mod 2 < 2") | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 116 | apply (erule less_2_cases [THEN disjE]) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 117 | apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 118 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 119 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 120 | lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 121 | apply (subgoal_tac "m mod 2 < 2") | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 122 | apply (force simp del: mod_less_divisor, simp) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 123 | done | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 124 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 125 | subsubsection{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 126 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 127 | lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 128 | by simp | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 129 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 130 | lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 131 | by simp | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 132 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 133 | text{*Can be used to eliminate long strings of Sucs, but not by default*}
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 134 | lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 135 | by simp | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 136 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 137 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 138 | text{*These lemmas collapse some needless occurrences of Suc:
 | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 139 | at least three Sucs, since two and fewer are rewritten back to Suc again! | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 140 | We already have some rules to simplify operands smaller than 3.*} | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 141 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 142 | lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 143 | by (simp add: Suc3_eq_add_3) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 144 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 145 | lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 146 | by (simp add: Suc3_eq_add_3) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 147 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 148 | lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 149 | by (simp add: Suc3_eq_add_3) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 150 | |
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 151 | lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 152 | by (simp add: Suc3_eq_add_3) | 
| 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 153 | |
| 17085 | 154 | lemmas Suc_div_eq_add3_div_number_of = | 
| 155 | Suc_div_eq_add3_div [of _ "number_of v", standard] | |
| 156 | declare Suc_div_eq_add3_div_number_of [simp] | |
| 157 | ||
| 158 | lemmas Suc_mod_eq_add3_mod_number_of = | |
| 159 | Suc_mod_eq_add3_mod [of _ "number_of v", standard] | |
| 160 | declare Suc_mod_eq_add3_mod_number_of [simp] | |
| 161 | ||
| 14273 
e33ffff0123c
further simplifications of the integer development; converting more .ML files
 paulson parents: 
14128diff
changeset | 162 | |
| 14288 | 163 | |
| 164 | subsection{*Special Simplification for Constants*}
 | |
| 165 | ||
| 166 | text{*These belong here, late in the development of HOL, to prevent their
 | |
| 167 | interfering with proofs of abstract properties of instances of the function | |
| 168 | @{term number_of}*}
 | |
| 169 | ||
| 170 | text{*These distributive laws move literals inside sums and differences.*}
 | |
| 17085 | 171 | lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard] | 
| 172 | declare left_distrib_number_of [simp] | |
| 173 | ||
| 174 | lemmas right_distrib_number_of = right_distrib [of "number_of v", standard] | |
| 175 | declare right_distrib_number_of [simp] | |
| 176 | ||
| 14288 | 177 | |
| 17085 | 178 | lemmas left_diff_distrib_number_of = | 
| 179 | left_diff_distrib [of _ _ "number_of v", standard] | |
| 180 | declare left_diff_distrib_number_of [simp] | |
| 181 | ||
| 182 | lemmas right_diff_distrib_number_of = | |
| 183 | right_diff_distrib [of "number_of v", standard] | |
| 184 | declare right_diff_distrib_number_of [simp] | |
| 185 | ||
| 14288 | 186 | |
| 187 | text{*These are actually for fields, like real: but where else to put them?*}
 | |
| 17085 | 188 | lemmas zero_less_divide_iff_number_of = | 
| 189 | zero_less_divide_iff [of "number_of w", standard] | |
| 190 | declare zero_less_divide_iff_number_of [simp] | |
| 191 | ||
| 192 | lemmas divide_less_0_iff_number_of = | |
| 193 | divide_less_0_iff [of "number_of w", standard] | |
| 194 | declare divide_less_0_iff_number_of [simp] | |
| 195 | ||
| 196 | lemmas zero_le_divide_iff_number_of = | |
| 197 | zero_le_divide_iff [of "number_of w", standard] | |
| 198 | declare zero_le_divide_iff_number_of [simp] | |
| 199 | ||
| 200 | lemmas divide_le_0_iff_number_of = | |
| 201 | divide_le_0_iff [of "number_of w", standard] | |
| 202 | declare divide_le_0_iff_number_of [simp] | |
| 203 | ||
| 14288 | 204 | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 205 | (**** | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 206 | IF times_divide_eq_right and times_divide_eq_left are removed as simprules, | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 207 | then these special-case declarations may be useful. | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 208 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 209 | text{*These simprules move numerals into numerators and denominators.*}
 | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 210 | lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 211 | by (simp add: times_divide_eq) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 212 | |
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 213 | lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 214 | by (simp add: times_divide_eq) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 215 | |
| 17085 | 216 | lemmas times_divide_eq_right_number_of = | 
| 217 | times_divide_eq_right [of "number_of w", standard] | |
| 218 | declare times_divide_eq_right_number_of [simp] | |
| 219 | ||
| 220 | lemmas times_divide_eq_right_number_of = | |
| 221 | times_divide_eq_right [of _ _ "number_of w", standard] | |
| 222 | declare times_divide_eq_right_number_of [simp] | |
| 223 | ||
| 224 | lemmas times_divide_eq_left_number_of = | |
| 225 | times_divide_eq_left [of _ "number_of w", standard] | |
| 226 | declare times_divide_eq_left_number_of [simp] | |
| 227 | ||
| 228 | lemmas times_divide_eq_left_number_of = | |
| 229 | times_divide_eq_left [of _ _ "number_of w", standard] | |
| 230 | declare times_divide_eq_left_number_of [simp] | |
| 231 | ||
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 232 | ****) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 233 | |
| 14577 | 234 | text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
 | 
| 235 | strange, but then other simprocs simplify the quotient.*} | |
| 236 | ||
| 17085 | 237 | lemmas inverse_eq_divide_number_of = | 
| 238 | inverse_eq_divide [of "number_of w", standard] | |
| 239 | declare inverse_eq_divide_number_of [simp] | |
| 240 | ||
| 14288 | 241 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 242 | subsubsection{*These laws simplify inequalities, moving unary minus from a term
 | 
| 14288 | 243 | into the literal.*} | 
| 17085 | 244 | lemmas less_minus_iff_number_of = | 
| 245 | less_minus_iff [of "number_of v", standard] | |
| 246 | declare less_minus_iff_number_of [simp] | |
| 247 | ||
| 248 | lemmas le_minus_iff_number_of = | |
| 249 | le_minus_iff [of "number_of v", standard] | |
| 250 | declare le_minus_iff_number_of [simp] | |
| 251 | ||
| 252 | lemmas equation_minus_iff_number_of = | |
| 253 | equation_minus_iff [of "number_of v", standard] | |
| 254 | declare equation_minus_iff_number_of [simp] | |
| 255 | ||
| 14288 | 256 | |
| 17085 | 257 | lemmas minus_less_iff_number_of = | 
| 258 | minus_less_iff [of _ "number_of v", standard] | |
| 259 | declare minus_less_iff_number_of [simp] | |
| 260 | ||
| 261 | lemmas minus_le_iff_number_of = | |
| 262 | minus_le_iff [of _ "number_of v", standard] | |
| 263 | declare minus_le_iff_number_of [simp] | |
| 264 | ||
| 265 | lemmas minus_equation_iff_number_of = | |
| 266 | minus_equation_iff [of _ "number_of v", standard] | |
| 267 | declare minus_equation_iff_number_of [simp] | |
| 268 | ||
| 14288 | 269 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 270 | subsubsection{*To Simplify Inequalities Where One Side is the Constant 1*}
 | 
| 17085 | 271 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 272 | lemma less_minus_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 273 |   fixes b::"'b::{ordered_idom,number_ring}" 
 | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 274 | shows "(1 < - b) = (b < -1)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 275 | by auto | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 276 | |
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 277 | lemma le_minus_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 278 |   fixes b::"'b::{ordered_idom,number_ring}" 
 | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 279 | shows "(1 \<le> - b) = (b \<le> -1)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 280 | by auto | 
| 17085 | 281 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 282 | lemma equation_minus_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 283 | fixes b::"'b::number_ring" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 284 | shows "(1 = - b) = (b = -1)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 285 | by (subst equation_minus_iff, auto) | 
| 14288 | 286 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 287 | lemma minus_less_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 288 |   fixes a::"'b::{ordered_idom,number_ring}" 
 | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 289 | shows "(- a < 1) = (-1 < a)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 290 | by auto | 
| 17085 | 291 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 292 | lemma minus_le_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 293 |   fixes a::"'b::{ordered_idom,number_ring}" 
 | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 294 | shows "(- a \<le> 1) = (-1 \<le> a)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 295 | by auto | 
| 17085 | 296 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 297 | lemma minus_equation_iff_1 [simp]: | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 298 | fixes a::"'b::number_ring" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 299 | shows "(- a = 1) = (a = -1)" | 
| 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 300 | by (subst minus_equation_iff, auto) | 
| 17085 | 301 | |
| 14288 | 302 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 303 | subsubsection {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
 | 
| 14288 | 304 | |
| 17085 | 305 | lemmas mult_less_cancel_left_number_of = | 
| 306 | mult_less_cancel_left [of "number_of v", standard] | |
| 307 | declare mult_less_cancel_left_number_of [simp] | |
| 308 | ||
| 309 | lemmas mult_less_cancel_right_number_of = | |
| 310 | mult_less_cancel_right [of _ "number_of v", standard] | |
| 311 | declare mult_less_cancel_right_number_of [simp] | |
| 312 | ||
| 313 | lemmas mult_le_cancel_left_number_of = | |
| 314 | mult_le_cancel_left [of "number_of v", standard] | |
| 315 | declare mult_le_cancel_left_number_of [simp] | |
| 316 | ||
| 317 | lemmas mult_le_cancel_right_number_of = | |
| 318 | mult_le_cancel_right [of _ "number_of v", standard] | |
| 319 | declare mult_le_cancel_right_number_of [simp] | |
| 320 | ||
| 14288 | 321 | |
| 22045 
ce5daf09ebfe
The "of 1, simplified" stopped working some time ago, leaving these simprules
 paulson parents: 
20500diff
changeset | 322 | subsubsection {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
 | 
| 14288 | 323 | |
| 17085 | 324 | lemmas le_divide_eq_number_of = le_divide_eq [of _ _ "number_of w", standard] | 
| 325 | declare le_divide_eq_number_of [simp] | |
| 326 | ||
| 327 | lemmas divide_le_eq_number_of = divide_le_eq [of _ "number_of w", standard] | |
| 328 | declare divide_le_eq_number_of [simp] | |
| 329 | ||
| 330 | lemmas less_divide_eq_number_of = less_divide_eq [of _ _ "number_of w", standard] | |
| 331 | declare less_divide_eq_number_of [simp] | |
| 332 | ||
| 333 | lemmas divide_less_eq_number_of = divide_less_eq [of _ "number_of w", standard] | |
| 334 | declare divide_less_eq_number_of [simp] | |
| 335 | ||
| 336 | lemmas eq_divide_eq_number_of = eq_divide_eq [of _ _ "number_of w", standard] | |
| 337 | declare eq_divide_eq_number_of [simp] | |
| 338 | ||
| 339 | lemmas divide_eq_eq_number_of = divide_eq_eq [of _ "number_of w", standard] | |
| 340 | declare divide_eq_eq_number_of [simp] | |
| 341 | ||
| 14288 | 342 | |
| 343 | ||
| 15228 | 344 | subsection{*Optional Simplification Rules Involving Constants*}
 | 
| 345 | ||
| 346 | text{*Simplify quotients that are compared with a literal constant.*}
 | |
| 347 | ||
| 348 | lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard] | |
| 349 | lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard] | |
| 350 | lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard] | |
| 351 | lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard] | |
| 352 | lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard] | |
| 353 | lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard] | |
| 354 | ||
| 355 | ||
| 356 | text{*Not good as automatic simprules because they cause case splits.*}
 | |
| 357 | lemmas divide_const_simps = | |
| 358 | le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of | |
| 359 | divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of | |
| 360 | le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 | |
| 361 | ||
| 17472 | 362 | subsubsection{*Division By @{text "-1"}*}
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 363 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 364 | lemma divide_minus1 [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 365 |      "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 366 | by simp | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 367 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 368 | lemma minus1_divide [simp]: | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 369 |      "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
 | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14387diff
changeset | 370 | by (simp add: divide_inverse inverse_minus_eq) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 371 | |
| 14475 | 372 | lemma half_gt_zero_iff: | 
| 373 |      "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
 | |
| 374 | by auto | |
| 375 | ||
| 18648 | 376 | lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard] | 
| 18624 | 377 | declare half_gt_zero [simp] | 
| 14475 | 378 | |
| 17149 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 379 | (* The following lemma should appear in Divides.thy, but there the proof | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 380 | doesn't work. *) | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 381 | |
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 382 | lemma nat_dvd_not_less: | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 383 | "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)" | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 384 | by (unfold dvd_def) auto | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
17085diff
changeset | 385 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 386 | ML | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 387 | {*
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 388 | val divide_minus1 = thm "divide_minus1"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 389 | val minus1_divide = thm "minus1_divide"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 390 | *} | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 391 | |
| 8858 
b739f0ecc1fa
new dummy theory; prevents strange errors when loading NatSimprocs.ML
 paulson parents: diff
changeset | 392 | end |