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