| author | wenzelm | 
| Mon, 23 Jan 2023 22:25:17 +0100 | |
| changeset 77059 | 422c57b75b17 | 
| parent 76224 | 64e8d4afcf10 | 
| child 78099 | 4d9349989d94 | 
| permissions | -rw-r--r-- | 
| 33366 | 1 | (* Author: Various *) | 
| 2 | ||
| 60758 | 3 | section \<open>Combination and Cancellation Simprocs for Numeral Expressions\<close> | 
| 33366 | 4 | |
| 5 | theory Numeral_Simprocs | |
| 76224 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
75880diff
changeset | 6 | imports Parity | 
| 33366 | 7 | begin | 
| 8 | ||
| 69605 | 9 | ML_file \<open>~~/src/Provers/Arith/assoc_fold.ML\<close> | 
| 10 | ML_file \<open>~~/src/Provers/Arith/cancel_numerals.ML\<close> | |
| 11 | ML_file \<open>~~/src/Provers/Arith/combine_numerals.ML\<close> | |
| 12 | ML_file \<open>~~/src/Provers/Arith/cancel_numeral_factor.ML\<close> | |
| 13 | ML_file \<open>~~/src/Provers/Arith/extract_common_term.ML\<close> | |
| 48891 | 14 | |
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47159diff
changeset | 15 | lemmas semiring_norm = | 
| 54249 | 16 | Let_def arith_simps diff_nat_numeral rel_simps | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47159diff
changeset | 17 | if_False if_True | 
| 75879 
24b17460ee60
streamlined simpset building, avoiding duplicated rewrite rules
 haftmann parents: 
75878diff
changeset | 18 | add_Suc add_numeral_left | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47159diff
changeset | 19 | add_neg_numeral_left mult_numeral_left | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 20 | numeral_One [symmetric] uminus_numeral_One [symmetric] Suc_eq_plus1 | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47159diff
changeset | 21 | eq_numeral_iff_iszero not_iszero_Numeral1 | 
| 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
47159diff
changeset | 22 | |
| 61799 | 23 | text \<open>For \<open>combine_numerals\<close>\<close> | 
| 33366 | 24 | |
| 25 | lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" | |
| 26 | by (simp add: add_mult_distrib) | |
| 27 | ||
| 61799 | 28 | text \<open>For \<open>cancel_numerals\<close>\<close> | 
| 33366 | 29 | |
| 30 | lemma nat_diff_add_eq1: | |
| 31 | "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" | |
| 63648 | 32 | by (simp split: nat_diff_split add: add_mult_distrib) | 
| 33366 | 33 | |
| 34 | lemma nat_diff_add_eq2: | |
| 35 | "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" | |
| 63648 | 36 | by (simp split: nat_diff_split add: add_mult_distrib) | 
| 33366 | 37 | |
| 38 | lemma nat_eq_add_iff1: | |
| 39 | "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" | |
| 63648 | 40 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 41 | |
| 42 | lemma nat_eq_add_iff2: | |
| 43 | "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" | |
| 63648 | 44 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 45 | |
| 46 | lemma nat_less_add_iff1: | |
| 47 | "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" | |
| 63648 | 48 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 49 | |
| 50 | lemma nat_less_add_iff2: | |
| 51 | "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" | |
| 63648 | 52 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 53 | |
| 54 | lemma nat_le_add_iff1: | |
| 55 | "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" | |
| 63648 | 56 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 57 | |
| 58 | lemma nat_le_add_iff2: | |
| 59 | "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" | |
| 63648 | 60 | by (auto split: nat_diff_split simp add: add_mult_distrib) | 
| 33366 | 61 | |
| 61799 | 62 | text \<open>For \<open>cancel_numeral_factors\<close>\<close> | 
| 33366 | 63 | |
| 64 | lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" | |
| 65 | by auto | |
| 66 | ||
| 67 | lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m<n)" | |
| 68 | by auto | |
| 69 | ||
| 70 | lemma nat_mult_eq_cancel1: "(0::nat) < k ==> (k*m = k*n) = (m=n)" | |
| 71 | by auto | |
| 72 | ||
| 73 | lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" | |
| 74 | by auto | |
| 75 | ||
| 76 | lemma nat_mult_dvd_cancel_disj[simp]: | |
| 67091 | 77 | "(k*m) dvd (k*n) = (k=0 \<or> m dvd (n::nat))" | 
| 47159 | 78 | by (auto simp: dvd_eq_mod_eq_0 mod_mult_mult1) | 
| 33366 | 79 | |
| 80 | lemma nat_mult_dvd_cancel1: "0 < k \<Longrightarrow> (k*m) dvd (k*n::nat) = (m dvd n)" | |
| 81 | by(auto) | |
| 82 | ||
| 61799 | 83 | text \<open>For \<open>cancel_factor\<close>\<close> | 
| 33366 | 84 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 85 | lemmas nat_mult_le_cancel_disj = mult_le_cancel1 | 
| 33366 | 86 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 87 | lemmas nat_mult_less_cancel_disj = mult_less_cancel1 | 
| 33366 | 88 | |
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 89 | lemma nat_mult_eq_cancel_disj: | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 90 | fixes k m n :: nat | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 91 | shows "k * m = k * n \<longleftrightarrow> k = 0 \<or> m = n" | 
| 75879 
24b17460ee60
streamlined simpset building, avoiding duplicated rewrite rules
 haftmann parents: 
75878diff
changeset | 92 | by (fact mult_cancel_left) | 
| 33366 | 93 | |
| 75879 
24b17460ee60
streamlined simpset building, avoiding duplicated rewrite rules
 haftmann parents: 
75878diff
changeset | 94 | lemma nat_mult_div_cancel_disj: | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 95 | fixes k m n :: nat | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 96 | shows "(k * m) div (k * n) = (if k = 0 then 0 else m div n)" | 
| 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54249diff
changeset | 97 | by (fact div_mult_mult1_if) | 
| 33366 | 98 | |
| 59757 
93d4169e07dc
fixed crash in cancel_numeral_simprocs. NB they still don't work except for type nat!
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 99 | lemma numeral_times_minus_swap: | 
| 
93d4169e07dc
fixed crash in cancel_numeral_simprocs. NB they still don't work except for type nat!
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 100 | fixes x:: "'a::comm_ring_1" shows "numeral w * -x = x * - numeral w" | 
| 75879 
24b17460ee60
streamlined simpset building, avoiding duplicated rewrite rules
 haftmann parents: 
75878diff
changeset | 101 | by (simp add: ac_simps) | 
| 59757 
93d4169e07dc
fixed crash in cancel_numeral_simprocs. NB they still don't work except for type nat!
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 102 | |
| 69605 | 103 | ML_file \<open>Tools/numeral_simprocs.ML\<close> | 
| 33366 | 104 | |
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 105 | simproc_setup semiring_assoc_fold | 
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 106 |   ("(a::'a::comm_semiring_1_cancel) * b") =
 | 
| 60758 | 107 | \<open>fn phi => Numeral_Simprocs.assoc_fold\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 108 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 109 | (* TODO: see whether the type class can be generalized further *) | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 110 | simproc_setup int_combine_numerals | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 111 |   ("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") =
 | 
| 60758 | 112 | \<open>fn phi => Numeral_Simprocs.combine_numerals\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 113 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 114 | simproc_setup field_combine_numerals | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 115 |   ("(i::'a::{field,ring_char_0}) + j"
 | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 116 |   |"(i::'a::{field,ring_char_0}) - j") =
 | 
| 60758 | 117 | \<open>fn phi => Numeral_Simprocs.field_combine_numerals\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 118 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 119 | simproc_setup inteq_cancel_numerals | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 120 |   ("(l::'a::comm_ring_1) + m = n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 121 | |"(l::'a::comm_ring_1) = m + n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 122 | |"(l::'a::comm_ring_1) - m = n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 123 | |"(l::'a::comm_ring_1) = m - n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 124 | |"(l::'a::comm_ring_1) * m = n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 125 | |"(l::'a::comm_ring_1) = m * n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 126 | |"- (l::'a::comm_ring_1) = m" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 127 | |"(l::'a::comm_ring_1) = - m") = | 
| 60758 | 128 | \<open>fn phi => Numeral_Simprocs.eq_cancel_numerals\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 129 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 130 | simproc_setup intless_cancel_numerals | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 131 |   ("(l::'a::linordered_idom) + m < n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 132 | |"(l::'a::linordered_idom) < m + n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 133 | |"(l::'a::linordered_idom) - m < n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 134 | |"(l::'a::linordered_idom) < m - n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 135 | |"(l::'a::linordered_idom) * m < n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 136 | |"(l::'a::linordered_idom) < m * n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 137 | |"- (l::'a::linordered_idom) < m" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 138 | |"(l::'a::linordered_idom) < - m") = | 
| 60758 | 139 | \<open>fn phi => Numeral_Simprocs.less_cancel_numerals\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 140 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 141 | simproc_setup intle_cancel_numerals | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 142 |   ("(l::'a::linordered_idom) + m \<le> n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 143 | |"(l::'a::linordered_idom) \<le> m + n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 144 | |"(l::'a::linordered_idom) - m \<le> n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 145 | |"(l::'a::linordered_idom) \<le> m - n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 146 | |"(l::'a::linordered_idom) * m \<le> n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 147 | |"(l::'a::linordered_idom) \<le> m * n" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 148 | |"- (l::'a::linordered_idom) \<le> m" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 149 | |"(l::'a::linordered_idom) \<le> - m") = | 
| 60758 | 150 | \<open>fn phi => Numeral_Simprocs.le_cancel_numerals\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 151 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 152 | simproc_setup ring_eq_cancel_numeral_factor | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 153 |   ("(l::'a::{idom,ring_char_0}) * m = n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 154 |   |"(l::'a::{idom,ring_char_0}) = m * n") =
 | 
| 60758 | 155 | \<open>fn phi => Numeral_Simprocs.eq_cancel_numeral_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 156 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 157 | simproc_setup ring_less_cancel_numeral_factor | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 158 |   ("(l::'a::linordered_idom) * m < n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 159 | |"(l::'a::linordered_idom) < m * n") = | 
| 60758 | 160 | \<open>fn phi => Numeral_Simprocs.less_cancel_numeral_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 161 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 162 | simproc_setup ring_le_cancel_numeral_factor | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 163 |   ("(l::'a::linordered_idom) * m <= n"
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 164 | |"(l::'a::linordered_idom) <= m * n") = | 
| 60758 | 165 | \<open>fn phi => Numeral_Simprocs.le_cancel_numeral_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 166 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45607diff
changeset | 167 | (* TODO: remove comm_ring_1 constraint if possible *) | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 168 | simproc_setup int_div_cancel_numeral_factors | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 169 |   ("((l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) * m) div n"
 | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 170 |   |"(l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) div (m * n)") =
 | 
| 60758 | 171 | \<open>fn phi => Numeral_Simprocs.div_cancel_numeral_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 172 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 173 | simproc_setup divide_cancel_numeral_factor | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 174 |   ("((l::'a::{field,ring_char_0}) * m) / n"
 | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 175 |   |"(l::'a::{field,ring_char_0}) / (m * n)"
 | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 176 |   |"((numeral v)::'a::{field,ring_char_0}) / (numeral w)") =
 | 
| 60758 | 177 | \<open>fn phi => Numeral_Simprocs.divide_cancel_numeral_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 178 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 179 | simproc_setup ring_eq_cancel_factor | 
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 180 |   ("(l::'a::idom) * m = n" | "(l::'a::idom) = m * n") =
 | 
| 60758 | 181 | \<open>fn phi => Numeral_Simprocs.eq_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 182 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 183 | simproc_setup linordered_ring_le_cancel_factor | 
| 45296 
7a97b2bda137
more accurate class constraints on cancellation simproc patterns
 huffman parents: 
45284diff
changeset | 184 |   ("(l::'a::linordered_idom) * m <= n"
 | 
| 
7a97b2bda137
more accurate class constraints on cancellation simproc patterns
 huffman parents: 
45284diff
changeset | 185 | |"(l::'a::linordered_idom) <= m * n") = | 
| 60758 | 186 | \<open>fn phi => Numeral_Simprocs.le_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 187 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 188 | simproc_setup linordered_ring_less_cancel_factor | 
| 45296 
7a97b2bda137
more accurate class constraints on cancellation simproc patterns
 huffman parents: 
45284diff
changeset | 189 |   ("(l::'a::linordered_idom) * m < n"
 | 
| 
7a97b2bda137
more accurate class constraints on cancellation simproc patterns
 huffman parents: 
45284diff
changeset | 190 | |"(l::'a::linordered_idom) < m * n") = | 
| 60758 | 191 | \<open>fn phi => Numeral_Simprocs.less_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 192 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 193 | simproc_setup int_div_cancel_factor | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 194 |   ("((l::'a::euclidean_semiring_cancel) * m) div n"
 | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 195 | |"(l::'a::euclidean_semiring_cancel) div (m * n)") = | 
| 60758 | 196 | \<open>fn phi => Numeral_Simprocs.div_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 197 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 198 | simproc_setup int_mod_cancel_factor | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 199 |   ("((l::'a::euclidean_semiring_cancel) * m) mod n"
 | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
63648diff
changeset | 200 | |"(l::'a::euclidean_semiring_cancel) mod (m * n)") = | 
| 60758 | 201 | \<open>fn phi => Numeral_Simprocs.mod_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 202 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 203 | simproc_setup dvd_cancel_factor | 
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 204 |   ("((l::'a::idom) * m) dvd n"
 | 
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 205 | |"(l::'a::idom) dvd (m * n)") = | 
| 60758 | 206 | \<open>fn phi => Numeral_Simprocs.dvd_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 207 | |
| 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 208 | simproc_setup divide_cancel_factor | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 209 |   ("((l::'a::field) * m) / n"
 | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59757diff
changeset | 210 | |"(l::'a::field) / (m * n)") = | 
| 60758 | 211 | \<open>fn phi => Numeral_Simprocs.divide_cancel_factor\<close> | 
| 45284 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 huffman parents: 
37886diff
changeset | 212 | |
| 69605 | 213 | ML_file \<open>Tools/nat_numeral_simprocs.ML\<close> | 
| 33366 | 214 | |
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 215 | simproc_setup nat_combine_numerals | 
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 216 |   ("(i::nat) + j" | "Suc (i + j)") =
 | 
| 60758 | 217 | \<open>fn phi => Nat_Numeral_Simprocs.combine_numerals\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 218 | |
| 45436 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 219 | simproc_setup nateq_cancel_numerals | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 220 |   ("(l::nat) + m = n" | "(l::nat) = m + n" |
 | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 221 | "(l::nat) * m = n" | "(l::nat) = m * n" | | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 222 | "Suc m = n" | "m = Suc n") = | 
| 60758 | 223 | \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals\<close> | 
| 45436 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 224 | |
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 225 | simproc_setup natless_cancel_numerals | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 226 |   ("(l::nat) + m < n" | "(l::nat) < m + n" |
 | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 227 | "(l::nat) * m < n" | "(l::nat) < m * n" | | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 228 | "Suc m < n" | "m < Suc n") = | 
| 60758 | 229 | \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numerals\<close> | 
| 45436 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 230 | |
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 231 | simproc_setup natle_cancel_numerals | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 232 |   ("(l::nat) + m \<le> n" | "(l::nat) \<le> m + n" |
 | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 233 | "(l::nat) * m \<le> n" | "(l::nat) \<le> m * n" | | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 234 | "Suc m \<le> n" | "m \<le> Suc n") = | 
| 60758 | 235 | \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numerals\<close> | 
| 45436 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 236 | |
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 237 | simproc_setup natdiff_cancel_numerals | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 238 |   ("((l::nat) + m) - n" | "(l::nat) - (m + n)" |
 | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 239 | "(l::nat) * m - n" | "(l::nat) - m * n" | | 
| 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 240 | "Suc m - n" | "m - Suc n") = | 
| 60758 | 241 | \<open>fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals\<close> | 
| 45436 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 huffman parents: 
45435diff
changeset | 242 | |
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 243 | simproc_setup nat_eq_cancel_numeral_factor | 
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 244 |   ("(l::nat) * m = n" | "(l::nat) = m * n") =
 | 
| 60758 | 245 | \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor\<close> | 
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 246 | |
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 247 | simproc_setup nat_less_cancel_numeral_factor | 
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 248 |   ("(l::nat) * m < n" | "(l::nat) < m * n") =
 | 
| 60758 | 249 | \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor\<close> | 
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 250 | |
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 251 | simproc_setup nat_le_cancel_numeral_factor | 
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 252 |   ("(l::nat) * m <= n" | "(l::nat) <= m * n") =
 | 
| 60758 | 253 | \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor\<close> | 
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 254 | |
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 255 | simproc_setup nat_div_cancel_numeral_factor | 
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 256 |   ("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
 | 
| 60758 | 257 | \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor\<close> | 
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 258 | |
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 259 | simproc_setup nat_dvd_cancel_numeral_factor | 
| 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 260 |   ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
 | 
| 60758 | 261 | \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\<close> | 
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 262 | |
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 263 | simproc_setup nat_eq_cancel_factor | 
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 264 |   ("(l::nat) * m = n" | "(l::nat) = m * n") =
 | 
| 60758 | 265 | \<open>fn phi => Nat_Numeral_Simprocs.eq_cancel_factor\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 266 | |
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 267 | simproc_setup nat_less_cancel_factor | 
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 268 |   ("(l::nat) * m < n" | "(l::nat) < m * n") =
 | 
| 60758 | 269 | \<open>fn phi => Nat_Numeral_Simprocs.less_cancel_factor\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 270 | |
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 271 | simproc_setup nat_le_cancel_factor | 
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 272 |   ("(l::nat) * m <= n" | "(l::nat) <= m * n") =
 | 
| 60758 | 273 | \<open>fn phi => Nat_Numeral_Simprocs.le_cancel_factor\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 274 | |
| 45463 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 huffman parents: 
45462diff
changeset | 275 | simproc_setup nat_div_cancel_factor | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 276 |   ("((l::nat) * m) div n" | "(l::nat) div (m * n)") =
 | 
| 60758 | 277 | \<open>fn phi => Nat_Numeral_Simprocs.div_cancel_factor\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 278 | |
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 279 | simproc_setup nat_dvd_cancel_factor | 
| 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 280 |   ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
 | 
| 60758 | 281 | \<open>fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor\<close> | 
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 282 | |
| 60758 | 283 | declaration \<open> | 
| 54249 | 284 | K (Lin_Arith.add_simprocs | 
| 69593 | 285 | [\<^simproc>\<open>semiring_assoc_fold\<close>, | 
| 286 | \<^simproc>\<open>int_combine_numerals\<close>, | |
| 287 | \<^simproc>\<open>inteq_cancel_numerals\<close>, | |
| 288 | \<^simproc>\<open>intless_cancel_numerals\<close>, | |
| 289 | \<^simproc>\<open>intle_cancel_numerals\<close>, | |
| 70356 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69605diff
changeset | 290 | \<^simproc>\<open>field_combine_numerals\<close>, | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69605diff
changeset | 291 | \<^simproc>\<open>nat_combine_numerals\<close>, | 
| 69593 | 292 | \<^simproc>\<open>nateq_cancel_numerals\<close>, | 
| 293 | \<^simproc>\<open>natless_cancel_numerals\<close>, | |
| 294 | \<^simproc>\<open>natle_cancel_numerals\<close>, | |
| 70356 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69605diff
changeset | 295 | \<^simproc>\<open>natdiff_cancel_numerals\<close>, | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69605diff
changeset | 296 | Numeral_Simprocs.field_divide_cancel_numeral_factor]) | 
| 60758 | 297 | \<close> | 
| 33366 | 298 | |
| 37886 | 299 | end |