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