| author | haftmann | 
| Sat, 15 Sep 2012 20:14:29 +0200 | |
| changeset 49388 | 1ffd5a055acf | 
| parent 48559 | 686cc7c47589 | 
| child 51717 | 9e7d1c139569 | 
| permissions | -rw-r--r-- | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/ex/Simproc_Tests.thy  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
2  | 
Author: Brian Huffman  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
5  | 
header {* Testing of arithmetic simprocs *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
7  | 
theory Simproc_Tests  | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
8  | 
imports Main  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
11  | 
text {*
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
12  | 
  This theory tests the various simprocs defined in @{file
 | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
13  | 
  "~~/src/HOL/Nat.thy"} and @{file "~~/src/HOL/Numeral_Simprocs.thy"}.
 | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
14  | 
Many of the tests are derived from commented-out code originally  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
15  | 
  found in @{file "~~/src/HOL/Tools/numeral_simprocs.ML"}.
 | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
16  | 
*}  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
18  | 
subsection {* ML bindings *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
20  | 
ML {*
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
21  | 
fun test ps = CHANGED (asm_simp_tac (HOL_basic_ss addsimprocs ps) 1)  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
22  | 
*}  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
23  | 
|
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
24  | 
subsection {* Cancellation simprocs from @{text Nat.thy} *}
 | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
25  | 
|
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
26  | 
notepad begin  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
27  | 
fix a b c d :: nat  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
28  | 
  {
 | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
29  | 
assume "b = Suc c" have "a + b = Suc (c + a)"  | 
| 
48559
 
686cc7c47589
give Nat_Arith simprocs proper name bindings by using simproc_setup
 
huffman 
parents: 
48556 
diff
changeset
 | 
30  | 
      by (tactic {* test [@{simproc nateq_cancel_sums}] *}) fact
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
31  | 
next  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
32  | 
assume "b < Suc c" have "a + b < Suc (c + a)"  | 
| 
48559
 
686cc7c47589
give Nat_Arith simprocs proper name bindings by using simproc_setup
 
huffman 
parents: 
48556 
diff
changeset
 | 
33  | 
      by (tactic {* test [@{simproc natless_cancel_sums}] *}) fact
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
34  | 
next  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
35  | 
assume "b \<le> Suc c" have "a + b \<le> Suc (c + a)"  | 
| 
48559
 
686cc7c47589
give Nat_Arith simprocs proper name bindings by using simproc_setup
 
huffman 
parents: 
48556 
diff
changeset
 | 
36  | 
      by (tactic {* test [@{simproc natle_cancel_sums}] *}) fact
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
37  | 
next  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
38  | 
assume "b - Suc c = d" have "a + b - Suc (c + a) = d"  | 
| 
48559
 
686cc7c47589
give Nat_Arith simprocs proper name bindings by using simproc_setup
 
huffman 
parents: 
48556 
diff
changeset
 | 
39  | 
      by (tactic {* test [@{simproc natdiff_cancel_sums}] *}) fact
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
40  | 
}  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
41  | 
end  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
42  | 
|
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
43  | 
schematic_lemma "\<And>(y::?'b::size). size (?x::?'a::size) \<le> size y + size ?x"  | 
| 
48559
 
686cc7c47589
give Nat_Arith simprocs proper name bindings by using simproc_setup
 
huffman 
parents: 
48556 
diff
changeset
 | 
44  | 
  by (tactic {* test [@{simproc natle_cancel_sums}] *}) (rule le0)
 | 
| 
48372
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
45  | 
(* TODO: test more simprocs with schematic variables *)  | 
| 
 
868dc809c8a2
make nat_cancel_sums simprocs robust in the presence of schematic variables; add regression tests
 
huffman 
parents: 
47108 
diff
changeset
 | 
46  | 
|
| 
45464
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
47  | 
subsection {* Abelian group cancellation simprocs *}
 | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
48  | 
|
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
49  | 
notepad begin  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
50  | 
fix a b c u :: "'a::ab_group_add"  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
51  | 
  {
 | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
52  | 
assume "(a + 0) - (b + 0) = u" have "(a + c) - (b + c) = u"  | 
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
48372 
diff
changeset
 | 
53  | 
      by (tactic {* test [@{simproc group_cancel_diff}] *}) fact
 | 
| 
45464
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
54  | 
next  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
55  | 
assume "a + 0 = b + 0" have "a + c = b + c"  | 
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
48372 
diff
changeset
 | 
56  | 
      by (tactic {* test [@{simproc group_cancel_eq}] *}) fact
 | 
| 
45464
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
57  | 
}  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
58  | 
end  | 
| 
48556
 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 
huffman 
parents: 
48372 
diff
changeset
 | 
59  | 
(* TODO: more tests for Groups.group_cancel_{add,diff,eq,less,le} *)
 | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
61  | 
subsection {* @{text int_combine_numerals} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
62  | 
|
| 
45464
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
63  | 
(* FIXME: int_combine_numerals often unnecessarily regroups addition  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
64  | 
and rewrites subtraction to negation. Ideally it should behave more  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
65  | 
like Groups.abel_cancel_sum, preserving the shape of terms as much as  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
66  | 
possible. *)  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
67  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
68  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
69  | 
fix a b c d oo uu i j k l u v w x y z :: "'a::comm_ring_1"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
70  | 
  {
 | 
| 
45464
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
71  | 
assume "a + - b = u" have "(a + c) - (b + c) = u"  | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
72  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 
huffman 
parents: 
45463 
diff
changeset
 | 
73  | 
next  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
74  | 
assume "10 + (2 * l + oo) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
75  | 
have "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
76  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
77  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
78  | 
assume "-3 + (i + (j + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
79  | 
have "(i + j + 12 + k) - 15 = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
80  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
81  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
82  | 
assume "7 + (i + (j + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
83  | 
have "(i + j + 12 + k) - 5 = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
84  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
85  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
86  | 
assume "-4 * (u * v) + (2 * x + y) = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
87  | 
have "(2*x - (u*v) + y) - v*3*u = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
88  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
89  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
90  | 
assume "2 * x * u * v + y = w"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
91  | 
have "(2*x*u*v + (u*v)*4 + y) - v*u*4 = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
92  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
93  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
94  | 
assume "3 * (u * v) + (2 * x * u * v + y) = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
95  | 
have "(2*x*u*v + (u*v)*4 + y) - v*u = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
96  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
97  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
98  | 
assume "-3 * (u * v) + (- (x * u * v) + - y) = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
99  | 
have "u*v - (x*u*v + (u*v)*4 + y) = w"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
100  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
101  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
102  | 
assume "a + - c = d"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
103  | 
have "a + -(b+c) + b = d"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
104  | 
apply (simp only: minus_add_distrib)  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
105  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
106  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
107  | 
assume "-2 * b + (a + - c) = d"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
108  | 
have "a + -(b+c) - b = d"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
109  | 
apply (simp only: minus_add_distrib)  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
110  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
111  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
112  | 
assume "-7 + (i + (j + (k + (- u + - y)))) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
113  | 
have "(i + j + -2 + k) - (u + 5 + y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
114  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
115  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
116  | 
assume "-27 + (i + (j + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
117  | 
have "(i + j + -12 + k) - 15 = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
118  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
119  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
120  | 
assume "27 + (i + (j + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
121  | 
have "(i + j + 12 + k) - -15 = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
122  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
123  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
124  | 
assume "3 + (i + (j + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
125  | 
have "(i + j + -12 + k) - -15 = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
126  | 
      by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
127  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
128  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
129  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
130  | 
subsection {* @{text inteq_cancel_numerals} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
131  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
132  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
133  | 
fix i j k u vv w y z w' y' z' :: "'a::comm_ring_1"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
134  | 
  {
 | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
135  | 
assume "u = 0" have "2*u = u"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
136  | 
      by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
 | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
137  | 
(* conclusion matches Rings.ring_1_no_zero_divisors_class.mult_cancel_right2 *)  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
138  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
139  | 
assume "i + (j + k) = 3 + (u + y)"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
140  | 
have "(i + j + 12 + k) = u + 15 + y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
141  | 
      by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
142  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
143  | 
assume "7 + (j + (i + k)) = y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
144  | 
have "(i + j*2 + 12 + k) = j + 5 + y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
145  | 
      by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
146  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
147  | 
assume "u + (6*z + (4*y + 6*w)) = 6*z' + (4*y' + (6*w' + vv))"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
148  | 
have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
149  | 
      by (tactic {* test [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
150  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
151  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
153  | 
subsection {* @{text intless_cancel_numerals} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
154  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
155  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
156  | 
fix b c i j k u y :: "'a::linordered_idom"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
157  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
158  | 
assume "y < 2 * b" have "y - b < b"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
159  | 
      by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
160  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
161  | 
assume "c + y < 4 * b" have "y - (3*b + c) < b - 2*c"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
162  | 
      by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
163  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
164  | 
assume "i + (j + k) < 8 + (u + y)"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
165  | 
have "(i + j + -3 + k) < u + 5 + y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
166  | 
      by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
167  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
168  | 
assume "9 + (i + (j + k)) < u + y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
169  | 
have "(i + j + 3 + k) < u + -6 + y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
170  | 
      by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
171  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
172  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
174  | 
subsection {* @{text ring_eq_cancel_numeral_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
175  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
176  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
177  | 
  fix x y :: "'a::{idom,ring_char_0}"
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
178  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
179  | 
assume "3*x = 4*y" have "9*x = 12 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
180  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
181  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
182  | 
assume "-3*x = 4*y" have "-99*x = 132 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
183  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
184  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
185  | 
assume "111*x = -44*y" have "999*x = -396 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
186  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
187  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
188  | 
assume "11*x = 9*y" have "-99*x = -81 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
189  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
190  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
191  | 
assume "2*x = y" have "-2 * x = -1 * y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
192  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
193  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
194  | 
assume "2*x = y" have "-2 * x = -y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
195  | 
      by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
196  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
197  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
199  | 
subsection {* @{text int_div_cancel_numeral_factors} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
200  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
201  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
202  | 
  fix x y z :: "'a::{semiring_div,comm_ring_1,ring_char_0}"
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
203  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
204  | 
assume "(3*x) div (4*y) = z" have "(9*x) div (12*y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
205  | 
      by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
206  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
207  | 
assume "(-3*x) div (4*y) = z" have "(-99*x) div (132*y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
208  | 
      by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
209  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
210  | 
assume "(111*x) div (-44*y) = z" have "(999*x) div (-396*y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
211  | 
      by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
212  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
213  | 
assume "(11*x) div (9*y) = z" have "(-99*x) div (-81*y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
214  | 
      by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
215  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
216  | 
assume "(2*x) div y = z"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
217  | 
have "(-2 * x) div (-1 * y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
218  | 
      by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
219  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
220  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
222  | 
subsection {* @{text ring_less_cancel_numeral_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
223  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
224  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
225  | 
fix x y :: "'a::linordered_idom"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
226  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
227  | 
assume "3*x < 4*y" have "9*x < 12 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
228  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
229  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
230  | 
assume "-3*x < 4*y" have "-99*x < 132 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
231  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
232  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
233  | 
assume "111*x < -44*y" have "999*x < -396 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
234  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
235  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
236  | 
assume "9*y < 11*x" have "-99*x < -81 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
237  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
238  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
239  | 
assume "y < 2*x" have "-2 * x < -y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
240  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
241  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
242  | 
assume "23*y < x" have "-x < -23 * y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
243  | 
      by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
244  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
245  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
246  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
247  | 
subsection {* @{text ring_le_cancel_numeral_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
248  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
249  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
250  | 
fix x y :: "'a::linordered_idom"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
251  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
252  | 
assume "3*x \<le> 4*y" have "9*x \<le> 12 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
253  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
254  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
255  | 
assume "-3*x \<le> 4*y" have "-99*x \<le> 132 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
256  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
257  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
258  | 
assume "111*x \<le> -44*y" have "999*x \<le> -396 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
259  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
260  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
261  | 
assume "9*y \<le> 11*x" have "-99*x \<le> -81 * y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
262  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
263  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
264  | 
assume "y \<le> 2*x" have "-2 * x \<le> -1 * y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
265  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
266  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
267  | 
assume "23*y \<le> x" have "-x \<le> -23 * y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
268  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
269  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
270  | 
assume "y \<le> 0" have "0 \<le> y * -2"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
271  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
272  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
273  | 
assume "- x \<le> y" have "- (2 * x) \<le> 2*y"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
274  | 
      by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
275  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
276  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
277  | 
|
| 45285 | 278  | 
subsection {* @{text divide_cancel_numeral_factor} *}
 | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
279  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
280  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
281  | 
  fix x y z :: "'a::{field_inverse_zero,ring_char_0}"
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
282  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
283  | 
assume "(3*x) / (4*y) = z" have "(9*x) / (12 * y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
284  | 
      by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
285  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
286  | 
assume "(-3*x) / (4*y) = z" have "(-99*x) / (132 * y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
287  | 
      by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
288  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
289  | 
assume "(111*x) / (-44*y) = z" have "(999*x) / (-396 * y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
290  | 
      by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
291  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
292  | 
assume "(11*x) / (9*y) = z" have "(-99*x) / (-81 * y) = z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
293  | 
      by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
294  | 
next  | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
295  | 
assume "(2*x) / y = z" have "(-2 * x) / (-1 * y) = z"  | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
296  | 
      by (tactic {* test [@{simproc divide_cancel_numeral_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
297  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
298  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
299  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
300  | 
subsection {* @{text ring_eq_cancel_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
301  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
302  | 
notepad begin  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
303  | 
fix a b c d k x y :: "'a::idom"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
304  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
305  | 
assume "k = 0 \<or> x = y" have "x*k = k*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
306  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
307  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
308  | 
assume "k = 0 \<or> 1 = y" have "k = k*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
309  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
310  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
311  | 
assume "b = 0 \<or> a*c = 1" have "a*(b*c) = b"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
312  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
313  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
314  | 
assume "a = 0 \<or> b = 0 \<or> c = d*x" have "a*(b*c) = d*b*(x*a)"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
315  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
316  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
317  | 
assume "k = 0 \<or> x = y" have "x*k = k*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
318  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
319  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
320  | 
assume "k = 0 \<or> 1 = y" have "k = k*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
321  | 
      by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
322  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
323  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
324  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
325  | 
subsection {* @{text int_div_cancel_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
326  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
327  | 
notepad begin  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
328  | 
fix a b c d k uu x y :: "'a::semiring_div"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
329  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
330  | 
assume "(if k = 0 then 0 else x div y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
331  | 
have "(x*k) div (k*y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
332  | 
      by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
333  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
334  | 
assume "(if k = 0 then 0 else 1 div y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
335  | 
have "(k) div (k*y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
336  | 
      by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
337  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
338  | 
assume "(if b = 0 then 0 else a * c) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
339  | 
have "(a*(b*c)) div b = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
340  | 
      by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
341  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
342  | 
assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
343  | 
have "(a*(b*c)) div (d*b*(x*a)) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
344  | 
      by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
345  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
346  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
347  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
348  | 
lemma shows "a*(b*c)/(y*z) = d*(b::'a::linordered_field_inverse_zero)*(x*a)/z"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
349  | 
oops -- "FIXME: need simproc to cover this case"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
350  | 
|
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
351  | 
subsection {* @{text divide_cancel_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
352  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
353  | 
notepad begin  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
354  | 
fix a b c d k uu x y :: "'a::field_inverse_zero"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
355  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
356  | 
assume "(if k = 0 then 0 else x / y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
357  | 
have "(x*k) / (k*y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
358  | 
      by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
359  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
360  | 
assume "(if k = 0 then 0 else 1 / y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
361  | 
have "(k) / (k*y) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
362  | 
      by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
363  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
364  | 
assume "(if b = 0 then 0 else a * c / 1) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
365  | 
have "(a*(b*c)) / b = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
366  | 
      by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
367  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
368  | 
assume "(if a = 0 then 0 else if b = 0 then 0 else c / (d * x)) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
369  | 
have "(a*(b*c)) / (d*b*(x*a)) = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
370  | 
      by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
371  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
372  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
373  | 
|
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
374  | 
lemma  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
375  | 
fixes a b c d x y z :: "'a::linordered_field_inverse_zero"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
376  | 
shows "a*(b*c)/(y*z) = d*(b)*(x*a)/z"  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
377  | 
oops -- "FIXME: need simproc to cover this case"  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
378  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
379  | 
subsection {* @{text linordered_ring_less_cancel_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
380  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
381  | 
notepad begin  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
382  | 
fix x y z :: "'a::linordered_idom"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
383  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
384  | 
assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> x*z < y*z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
385  | 
      by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
386  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
387  | 
assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> x*z < z*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
388  | 
      by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
389  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
390  | 
assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> z*x < y*z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
391  | 
      by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
392  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
393  | 
assume "0 < z \<Longrightarrow> x < y" have "0 < z \<Longrightarrow> z*x < z*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
394  | 
      by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
 | 
| 
46240
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
395  | 
next  | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
396  | 
txt "This simproc now uses the simplifier to prove that terms to  | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
397  | 
be canceled are positive/negative."  | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
398  | 
assume z_pos: "0 < z"  | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
399  | 
assume "x < y" have "z*x < z*y"  | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
400  | 
      by (tactic {* CHANGED (asm_simp_tac (HOL_basic_ss
 | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
401  | 
        addsimprocs [@{simproc linordered_ring_less_cancel_factor}]
 | 
| 
 
933f35c4e126
factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
 
huffman 
parents: 
45530 
diff
changeset
 | 
402  | 
        addsimps [@{thm z_pos}]) 1) *}) fact
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
403  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
404  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
405  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
406  | 
subsection {* @{text linordered_ring_le_cancel_factor} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
407  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
408  | 
notepad begin  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
409  | 
fix x y z :: "'a::linordered_idom"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
410  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
411  | 
assume "0 < z \<Longrightarrow> x \<le> y" have "0 < z \<Longrightarrow> x*z \<le> y*z"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
412  | 
      by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
413  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
414  | 
assume "0 < z \<Longrightarrow> x \<le> y" have "0 < z \<Longrightarrow> z*x \<le> z*y"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
415  | 
      by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
416  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
417  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
418  | 
|
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
419  | 
subsection {* @{text field_combine_numerals} *}
 | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
420  | 
|
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
421  | 
notepad begin  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
422  | 
  fix x y z uu :: "'a::{field_inverse_zero,ring_char_0}"
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
423  | 
  {
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
424  | 
assume "5 / 6 * x = uu" have "x / 2 + x / 3 = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
425  | 
      by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
426  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
427  | 
assume "6 / 9 * x + y = uu" have "x / 3 + y + x / 3 = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
428  | 
      by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
 | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
429  | 
next  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
430  | 
assume "9 / 9 * x = uu" have "2 * x / 3 + x / 3 = uu"  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
431  | 
      by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
 | 
| 
45437
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
432  | 
next  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
433  | 
assume "y + z = uu"  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
434  | 
have "x / 2 + y - 3 * x / 6 + z = uu"  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
435  | 
      by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
 | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
436  | 
next  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
437  | 
assume "1 / 15 * x + y = uu"  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
438  | 
have "7 * x / 5 + y - 4 * x / 3 = uu"  | 
| 
 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 
huffman 
parents: 
45436 
diff
changeset
 | 
439  | 
      by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
 | 
| 
45435
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
440  | 
}  | 
| 
 
d660c4b9daa6
add ring_char_0 class constraints to several simprocs (internal proofs of #n ~= 0 fail for type
 
huffman 
parents: 
45285 
diff
changeset
 | 
441  | 
end  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
442  | 
|
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
443  | 
lemma  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
444  | 
  fixes x :: "'a::{linordered_field_inverse_zero}"
 | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
445  | 
shows "2/3 * x + x / 3 = uu"  | 
| 
45284
 
ae78a4ffa81d
use simproc_setup for cancellation simprocs, to get proper name bindings
 
huffman 
parents: 
45270 
diff
changeset
 | 
446  | 
apply (tactic {* test [@{simproc field_combine_numerals}] *})?
 | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
447  | 
oops -- "FIXME: test fails"  | 
| 
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
448  | 
|
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
449  | 
subsection {* @{text nat_combine_numerals} *}
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
450  | 
|
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
451  | 
notepad begin  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
452  | 
fix i j k m n u :: nat  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
453  | 
  {
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
454  | 
assume "4*k = u" have "k + 3*k = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
455  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
456  | 
next  | 
| 
45530
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
457  | 
(* FIXME "Suc (i + 3) \<equiv> i + 4" *)  | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
458  | 
assume "4 * Suc 0 + i = u" have "Suc (i + 3) = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
459  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
460  | 
next  | 
| 
45530
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
461  | 
(* FIXME "Suc (i + j + 3 + k) \<equiv> i + j + 4 + k" *)  | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
462  | 
assume "4 * Suc 0 + (i + (j + k)) = u" have "Suc (i + j + 3 + k) = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
463  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
464  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
465  | 
assume "2 * j + 4 * k = u" have "k + j + 3*k + j = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
466  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
467  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
468  | 
assume "6 * Suc 0 + (5 * (i * j) + (4 * k + i)) = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
469  | 
have "Suc (j*i + i + k + 5 + 3*k + i*j*4) = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
470  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
471  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
472  | 
assume "5 * (m * n) = u" have "(2*n*m) + (3*(m*n)) = u"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
473  | 
      by (tactic {* test [@{simproc nat_combine_numerals}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
474  | 
}  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
475  | 
end  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
476  | 
|
| 
45436
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
477  | 
subsection {* @{text nateq_cancel_numerals} *}
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
478  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
479  | 
notepad begin  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
480  | 
fix i j k l oo u uu vv w y z w' y' z' :: "nat"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
481  | 
  {
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
482  | 
assume "Suc 0 * u = 0" have "2*u = (u::nat)"  | 
| 
45436
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
483  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
484  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
485  | 
assume "Suc 0 * u = Suc 0" have "2*u = Suc (u)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
486  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
487  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
488  | 
assume "i + (j + k) = 3 * Suc 0 + (u + y)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
489  | 
have "(i + j + 12 + k) = u + 15 + y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
490  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
491  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
492  | 
assume "7 * Suc 0 + (i + (j + k)) = u + y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
493  | 
have "(i + j + 12 + k) = u + 5 + y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
494  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
495  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
496  | 
assume "11 * Suc 0 + (i + (j + k)) = u + y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
497  | 
have "(i + j + 12 + k) = Suc (u + y)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
498  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
499  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
500  | 
assume "i + (j + k) = 2 * Suc 0 + (u + y)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
501  | 
have "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
502  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
503  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
504  | 
assume "Suc 0 * u + (2 * y + 3 * z) = Suc 0"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
505  | 
have "2*y + 3*z + 2*u = Suc (u)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
506  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
507  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
508  | 
assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) = Suc 0"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
509  | 
have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
510  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
511  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
512  | 
assume "Suc 0 * u + (2 * y + (3 * z + (6 * w + (2 * y + 3 * z)))) =  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
513  | 
2 * y' + (3 * z' + (6 * w' + (2 * y' + (3 * z' + vv))))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
514  | 
have "2*y + 3*z + 6*w + 2*y + 3*z + 2*u =  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
515  | 
2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + vv"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
516  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
517  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
518  | 
assume "2 * u + (2 * z + (5 * Suc 0 + 2 * y)) = vv"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
519  | 
have "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
520  | 
      by (tactic {* test [@{simproc nateq_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
521  | 
}  | 
| 
45224
 
b1d5b3820d82
add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
 
huffman 
parents:  
diff
changeset
 | 
522  | 
end  | 
| 
45436
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
523  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
524  | 
subsection {* @{text natless_cancel_numerals} *}
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
525  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
526  | 
notepad begin  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
527  | 
fix length :: "'a \<Rightarrow> nat" and l1 l2 xs :: "'a" and f :: "nat \<Rightarrow> 'a"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46240 
diff
changeset
 | 
528  | 
fix c i j k l m oo u uu vv w y z w' y' z' :: "nat"  | 
| 
45436
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
529  | 
  {
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
530  | 
assume "0 < j" have "(2*length xs < 2*length xs + j)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
531  | 
      by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
532  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
533  | 
assume "0 < j" have "(2*length xs < length xs * 2 + j)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
534  | 
      by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
535  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
536  | 
assume "i + (j + k) < u + y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
537  | 
have "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
538  | 
      by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
539  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
540  | 
assume "0 < Suc 0 * (m * n) + u" have "(2*n*m) < (3*(m*n)) + u"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
541  | 
      by (tactic {* test [@{simproc natless_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
542  | 
}  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
543  | 
end  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
544  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
545  | 
subsection {* @{text natle_cancel_numerals} *}
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
546  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
547  | 
notepad begin  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
548  | 
fix length :: "'a \<Rightarrow> nat" and l2 l3 :: "'a" and f :: "nat \<Rightarrow> 'a"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
549  | 
fix c e i j k l oo u uu vv w y z w' y' z' :: "nat"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
550  | 
  {
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
551  | 
assume "u + y \<le> 36 * Suc 0 + (i + (j + k))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
552  | 
have "Suc (Suc (Suc (Suc (Suc (u + y))))) \<le> ((i + j) + 41 + k)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
553  | 
      by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
554  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
555  | 
assume "5 * Suc 0 + (case length (f c) of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> k) = 0"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
556  | 
have "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) \<le> Suc 0)"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
557  | 
      by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
558  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
559  | 
assume "6 + length l2 = 0" have "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) \<le> length l1"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
560  | 
      by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
561  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
562  | 
assume "5 + length l3 = 0"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
563  | 
have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) \<le> length (compT P E A ST mxr e))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
564  | 
      by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
565  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
566  | 
assume "5 + length (compT P E (A \<union> A' e) ST mxr c) = 0"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
567  | 
have "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un A' e) ST mxr c))))))) \<le> length (compT P E A ST mxr e))"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
568  | 
      by (tactic {* test [@{simproc natle_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
569  | 
}  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
570  | 
end  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
571  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
572  | 
subsection {* @{text natdiff_cancel_numerals} *}
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
573  | 
|
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
574  | 
notepad begin  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
575  | 
fix length :: "'a \<Rightarrow> nat" and l2 l3 :: "'a" and f :: "nat \<Rightarrow> 'a"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
576  | 
fix c e i j k l oo u uu vv v w x y z zz w' y' z' :: "nat"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
577  | 
  {
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
578  | 
assume "i + (j + k) - 3 * Suc 0 = y" have "(i + j + 12 + k) - 15 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
579  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
580  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
581  | 
assume "7 * Suc 0 + (i + (j + k)) - 0 = y" have "(i + j + 12 + k) - 5 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
582  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
583  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
584  | 
assume "u - Suc 0 * Suc 0 = y" have "Suc u - 2 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
585  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
586  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
587  | 
assume "Suc 0 * Suc 0 + u - 0 = y" have "Suc (Suc (Suc u)) - 2 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
588  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
589  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
590  | 
assume "Suc 0 * Suc 0 + (i + (j + k)) - 0 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
591  | 
have "(i + j + 2 + k) - 1 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
592  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
593  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
594  | 
assume "i + (j + k) - Suc 0 * Suc 0 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
595  | 
have "(i + j + 1 + k) - 2 = y"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
596  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
597  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
598  | 
assume "2 * x + y - 2 * (u * v) = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
599  | 
have "(2*x + (u*v) + y) - v*3*u = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
600  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
601  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
602  | 
assume "2 * x * u * v + (5 + y) - 0 = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
603  | 
have "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
604  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
605  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
606  | 
assume "3 * (u * v) + (2 * x * u * v + y) - 0 = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
607  | 
have "(2*x*u*v + (u*v)*4 + y) - v*u = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
608  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
609  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
610  | 
assume "3 * u + (2 + (2 * x * u * v + y)) - 0 = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
611  | 
have "Suc (Suc (2*x*u*v + u*4 + y)) - u = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
612  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
613  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
614  | 
assume "Suc (Suc 0 * (u * v)) - 0 = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
615  | 
have "Suc ((u*v)*4) - v*3*u = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
616  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
617  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
618  | 
assume "2 - 0 = w" have "Suc (Suc ((u*v)*3)) - v*3*u = w"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
619  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
620  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
621  | 
assume "17 * Suc 0 + (i + (j + k)) - (u + y) = zz"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
622  | 
have "(i + j + 32 + k) - (u + 15 + y) = zz"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
623  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
624  | 
next  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
625  | 
assume "u + y - 0 = v" have "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v"  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
626  | 
      by (tactic {* test [@{simproc natdiff_cancel_numerals}] *}) fact
 | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
627  | 
}  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
628  | 
end  | 
| 
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
629  | 
|
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
630  | 
subsection {* Factor-cancellation simprocs for type @{typ nat} *}
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
631  | 
|
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
632  | 
text {* @{text nat_eq_cancel_factor}, @{text nat_less_cancel_factor},
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
633  | 
@{text nat_le_cancel_factor}, @{text nat_divide_cancel_factor}, and
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
634  | 
@{text nat_dvd_cancel_factor}. *}
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
635  | 
|
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
636  | 
notepad begin  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
637  | 
fix a b c d k x y uu :: nat  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
638  | 
  {
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
639  | 
assume "k = 0 \<or> x = y" have "x*k = k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
640  | 
      by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
641  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
642  | 
assume "k = 0 \<or> Suc 0 = y" have "k = k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
643  | 
      by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
644  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
645  | 
assume "b = 0 \<or> a * c = Suc 0" have "a*(b*c) = b"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
646  | 
      by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
647  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
648  | 
assume "a = 0 \<or> b = 0 \<or> c = d * x" have "a*(b*c) = d*b*(x*a)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
649  | 
      by (tactic {* test [@{simproc nat_eq_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
650  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
651  | 
assume "0 < k \<and> x < y" have "x*k < k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
652  | 
      by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
653  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
654  | 
assume "0 < k \<and> Suc 0 < y" have "k < k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
655  | 
      by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
656  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
657  | 
assume "0 < b \<and> a * c < Suc 0" have "a*(b*c) < b"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
658  | 
      by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
659  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
660  | 
assume "0 < a \<and> 0 < b \<and> c < d * x" have "a*(b*c) < d*b*(x*a)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
661  | 
      by (tactic {* test [@{simproc nat_less_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
662  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
663  | 
assume "0 < k \<longrightarrow> x \<le> y" have "x*k \<le> k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
664  | 
      by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
665  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
666  | 
assume "0 < k \<longrightarrow> Suc 0 \<le> y" have "k \<le> k*y"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
667  | 
      by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
668  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
669  | 
assume "0 < b \<longrightarrow> a * c \<le> Suc 0" have "a*(b*c) \<le> b"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
670  | 
      by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
671  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
672  | 
assume "0 < a \<longrightarrow> 0 < b \<longrightarrow> c \<le> d * x" have "a*(b*c) \<le> d*b*(x*a)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
673  | 
      by (tactic {* test [@{simproc nat_le_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
674  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
675  | 
assume "(if k = 0 then 0 else x div y) = uu" have "(x*k) div (k*y) = uu"  | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
676  | 
      by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact
 | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
677  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
678  | 
assume "(if k = 0 then 0 else Suc 0 div y) = uu" have "k div (k*y) = uu"  | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
679  | 
      by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact
 | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
680  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
681  | 
assume "(if b = 0 then 0 else a * c) = uu" have "(a*(b*c)) div (b) = uu"  | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
682  | 
      by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact
 | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
683  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
684  | 
assume "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
685  | 
have "(a*(b*c)) div (d*b*(x*a)) = uu"  | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
686  | 
      by (tactic {* test [@{simproc nat_div_cancel_factor}] *}) fact
 | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
687  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
688  | 
assume "k = 0 \<or> x dvd y" have "(x*k) dvd (k*y)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
689  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
690  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
691  | 
assume "k = 0 \<or> Suc 0 dvd y" have "k dvd (k*y)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
692  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
693  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
694  | 
assume "b = 0 \<or> a * c dvd Suc 0" have "(a*(b*c)) dvd (b)"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
695  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
696  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
697  | 
assume "b = 0 \<or> Suc 0 dvd a * c" have "b dvd (a*(b*c))"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
698  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
699  | 
next  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
700  | 
assume "a = 0 \<or> b = 0 \<or> c dvd d * x" have "(a*(b*c)) dvd (d*b*(x*a))"  | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
701  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_factor}] *}) fact
 | 
| 
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
702  | 
}  | 
| 
45436
 
62bc9474d04b
use simproc_setup for some nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45435 
diff
changeset
 | 
703  | 
end  | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
704  | 
|
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
705  | 
subsection {* Numeral-cancellation simprocs for type @{typ nat} *}
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
706  | 
|
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
707  | 
notepad begin  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
708  | 
fix x y z :: nat  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
709  | 
  {
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
710  | 
assume "3 * x = 4 * y" have "9*x = 12 * y"  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
711  | 
      by (tactic {* test [@{simproc nat_eq_cancel_numeral_factor}] *}) fact
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
712  | 
next  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
713  | 
assume "3 * x < 4 * y" have "9*x < 12 * y"  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
714  | 
      by (tactic {* test [@{simproc nat_less_cancel_numeral_factor}] *}) fact
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
715  | 
next  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
716  | 
assume "3 * x \<le> 4 * y" have "9*x \<le> 12 * y"  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
717  | 
      by (tactic {* test [@{simproc nat_le_cancel_numeral_factor}] *}) fact
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
718  | 
next  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
719  | 
assume "(3 * x) div (4 * y) = z" have "(9*x) div (12 * y) = z"  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
720  | 
      by (tactic {* test [@{simproc nat_div_cancel_numeral_factor}] *}) fact
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
721  | 
next  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
722  | 
assume "(3 * x) dvd (4 * y)" have "(9*x) dvd (12 * y)"  | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
723  | 
      by (tactic {* test [@{simproc nat_dvd_cancel_numeral_factor}] *}) fact
 | 
| 
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
724  | 
}  | 
| 
45462
 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 
huffman 
parents: 
45437 
diff
changeset
 | 
725  | 
end  | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
726  | 
|
| 
45530
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
727  | 
subsection {* Integer numeral div/mod simprocs *}
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
728  | 
|
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
729  | 
notepad begin  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
730  | 
have "(10::int) div 3 = 3"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
731  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
732  | 
have "(10::int) mod 3 = 1"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
733  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
734  | 
have "(10::int) div -3 = -4"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
735  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
736  | 
have "(10::int) mod -3 = -2"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
737  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
738  | 
have "(-10::int) div 3 = -4"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
739  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
740  | 
have "(-10::int) mod 3 = 2"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
741  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
742  | 
have "(-10::int) div -3 = 3"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
743  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
744  | 
have "(-10::int) mod -3 = -1"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
745  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
746  | 
have "(8452::int) mod 3 = 1"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
747  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
748  | 
have "(59485::int) div 434 = 137"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
749  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
750  | 
have "(1000006::int) mod 10 = 6"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
751  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
752  | 
have "10000000 div 2 = (5000000::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
753  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
754  | 
have "10000001 mod 2 = (1::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
755  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
756  | 
have "10000055 div 32 = (312501::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
757  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
758  | 
have "10000055 mod 32 = (23::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
759  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
760  | 
have "100094 div 144 = (695::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
761  | 
    by (tactic {* test [@{simproc binary_int_div}] *})
 | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
762  | 
have "100094 mod 144 = (14::int)"  | 
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
763  | 
    by (tactic {* test [@{simproc binary_int_mod}] *})
 | 
| 
45463
 
9a588a835c1e
use simproc_setup for the remaining nat_numeral simprocs
 
huffman 
parents: 
45462 
diff
changeset
 | 
764  | 
end  | 
| 
45530
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
765  | 
|
| 
 
0c4853bb77bf
rewrite integer numeral div/mod simprocs to not return conditional rewrites; add regression tests
 
huffman 
parents: 
45464 
diff
changeset
 | 
766  | 
end  |