src/HOL/ex/Simproc_Tests.thy
author huffman
Fri, 28 Oct 2011 11:02:27 +0200
changeset 45284 ae78a4ffa81d
parent 45270 d5b5c9259afd
child 45285 299abd2931d5
permissions -rw-r--r--
use simproc_setup for cancellation simprocs, to get proper name bindings
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
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
     8
imports Rat
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 {*
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    12
  This theory tests the various simprocs defined in
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    13
  @{file "~~/src/HOL/Numeral_Simprocs.thy"}. Many of the tests
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    14
  are derived from commented-out code originally found in
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    15
  @{file "~~/src/HOL/Tools/numeral_simprocs.ML"}.
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    16
*}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    17
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    18
subsection {* ML bindings *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    19
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    20
ML {*
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    21
  val [field_eq_cancel_numeral_factor, field_cancel_numeral_factor]
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    22
    = Numeral_Simprocs.field_cancel_numeral_factors
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
  fun test ps = CHANGED (asm_simp_tac (HOL_basic_ss addsimprocs ps) 1)
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    25
*}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    26
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    27
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    28
subsection {* @{text int_combine_numerals} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    29
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    30
lemma assumes "10 + (2 * l + oo) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    31
  shows "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    32
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    33
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    34
lemma assumes "-3 + (i + (j + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    35
  shows "(i + j + 12 + (k::int)) - 15 = y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    36
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    37
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    38
lemma assumes "7 + (i + (j + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    39
  shows "(i + j + 12 + (k::int)) - 5 = y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    40
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    41
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    42
lemma assumes "-4 * (u * v) + (2 * x + y) = w"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    43
  shows "(2*x - (u*v) + y) - v*3*u = (w::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    44
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    45
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    46
lemma assumes "Numeral0 * (u*v) + (2 * x * u * v + y) = w"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    47
  shows "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    48
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    49
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    50
lemma assumes "3 * (u * v) + (2 * x * u * v + y) = w"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    51
  shows "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    52
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    53
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    54
lemma assumes "-3 * (u * v) + (- (x * u * v) + - y) = w"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    55
  shows "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    56
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    57
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    58
lemma assumes "Numeral0 * b + (a + - c) = d"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    59
  shows "a + -(b+c) + b = (d::int)"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    60
apply (simp only: minus_add_distrib)
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    61
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    62
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    63
lemma assumes "-2 * b + (a + - c) = d"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    64
  shows "a + -(b+c) - b = (d::int)"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    65
apply (simp only: minus_add_distrib)
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    66
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    67
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    68
lemma assumes "-7 + (i + (j + (k + (- u + - y)))) = zz"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    69
  shows "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    70
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    71
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    72
lemma assumes "-27 + (i + (j + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    73
  shows "(i + j + -12 + (k::int)) - 15 = y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    74
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    75
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    76
lemma assumes "27 + (i + (j + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    77
  shows "(i + j + 12 + (k::int)) - -15 = y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    78
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    79
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    80
lemma assumes "3 + (i + (j + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    81
  shows "(i + j + -12 + (k::int)) - -15 = y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    82
by (tactic {* test [@{simproc int_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    83
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    84
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    85
subsection {* @{text inteq_cancel_numerals} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    86
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    87
lemma assumes "u = Numeral0" shows "2*u = (u::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    88
by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    89
(* conclusion matches Rings.ring_1_no_zero_divisors_class.mult_cancel_right2 *)
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    90
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    91
lemma assumes "i + (j + k) = 3 + (u + y)"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    92
  shows "(i + j + 12 + (k::int)) = u + 15 + y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    93
by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    94
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    95
lemma assumes "7 + (j + (i + k)) = y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    96
  shows "(i + j*2 + 12 + (k::int)) = j + 5 + y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
    97
by (tactic {* test [@{simproc inteq_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    98
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
    99
lemma assumes "u + (6*z + (4*y + 6*w)) = 6*z' + (4*y' + (6*w' + vv))"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   100
  shows "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   101
by (tactic {* test [@{simproc int_combine_numerals}, @{simproc inteq_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   102
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   103
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   104
subsection {* @{text intless_cancel_numerals} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   105
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   106
lemma assumes "y < 2 * b" shows "y - b < (b::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   107
by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   108
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   109
lemma assumes "c + y < 4 * b" shows "y - (3*b + c) < (b::int) - 2*c"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   110
by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   111
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   112
lemma assumes "i + (j + k) < 8 + (u + y)"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   113
  shows "(i + j + -3 + (k::int)) < u + 5 + y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   114
by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   115
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   116
lemma assumes "9 + (i + (j + k)) < u + y"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   117
  shows "(i + j + 3 + (k::int)) < u + -6 + y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   118
by (tactic {* test [@{simproc intless_cancel_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   119
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   120
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   121
subsection {* @{text ring_eq_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   122
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   123
lemma assumes "3*x = 4*y" shows "9*x = 12 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   124
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   125
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   126
lemma assumes "-3*x = 4*y" shows "-99*x = 132 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   127
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   128
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   129
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   130
subsection {* @{text int_div_cancel_numeral_factors} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   131
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   132
lemma assumes "(3*x) div (4*y) = z" shows "(9*x) div (12*y) = (z::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   133
by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   134
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   135
lemma assumes "(-3*x) div (4*y) = z" shows "(-99*x) div (132*y) = (z::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   136
by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   137
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   138
lemma assumes "(111*x) div (-44*y) = z" shows "(999*x) div (-396*y) = (z::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   139
by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   140
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   141
lemma assumes "(11*x) div (9*y) = z" shows "(-99*x) div (-81*y) = (z::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   142
by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   143
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   144
lemma assumes "(2*x) div (Numeral1*y) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   145
  shows "(-2 * x) div (-1 * (y::int)) = z"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   146
by (tactic {* test [@{simproc int_div_cancel_numeral_factors}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   147
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   148
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   149
subsection {* @{text ring_less_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   150
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   151
lemma assumes "3*x < 4*y" shows "9*x < 12 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   152
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
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
lemma assumes "-3*x < 4*y" shows "-99*x < 132 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   155
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   156
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   157
lemma assumes "111*x < -44*y" shows "999*x < -396 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   158
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   159
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   160
lemma assumes "9*y < 11*x" shows "-99*x < -81 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   161
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   162
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   163
lemma assumes "Numeral1*y < 2*x" shows "-2 * x < -(y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   164
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   165
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   166
lemma assumes "23*y < Numeral1*x" shows "-x < -23 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   167
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   168
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   169
lemma assumes "3*x < 4*y" shows "9*x < 12 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   170
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   171
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   172
lemma assumes "-3*x < 4*y" shows "-99*x < 132 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   173
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
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
lemma assumes "111*x < -44*y" shows "999*x < -396 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   176
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   177
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   178
lemma assumes "9*y < 11*x" shows "-99*x < -81 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   179
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   180
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   181
lemma assumes "Numeral1*y < 2*x" shows "-2 * x < -(y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   182
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   183
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   184
lemma assumes "23*y < Numeral1*x" shows "-x < -23 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   185
by (tactic {* test [@{simproc ring_less_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   186
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   187
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   188
subsection {* @{text ring_le_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   189
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   190
lemma assumes "3*x \<le> 4*y" shows "9*x \<le> 12 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   191
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   192
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   193
lemma assumes "-3*x \<le> 4*y" shows "-99*x \<le> 132 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   194
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   195
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   196
lemma assumes "111*x \<le> -44*y" shows "999*x \<le> -396 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   197
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   198
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   199
lemma assumes "9*y \<le> 11*x" shows "-99*x \<le> -81 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   200
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   201
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   202
lemma assumes "Numeral1*y \<le> 2*x" shows "-2 * x \<le> -1 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   203
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   204
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   205
lemma assumes "23*y \<le> Numeral1*x" shows "-x \<le> -23 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   206
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   207
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   208
lemma assumes "Numeral1*y \<le> Numeral0" shows "0 \<le> (y::rat) * -2"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   209
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   210
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   211
lemma assumes "3*x \<le> 4*y" shows "9*x \<le> 12 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   212
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   213
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   214
lemma assumes "-3*x \<le> 4*y" shows "-99*x \<le> 132 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   215
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   216
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   217
lemma assumes "111*x \<le> -44*y" shows "999*x \<le> -396 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   218
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   219
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   220
lemma assumes "-1*x \<le> Numeral1*y" shows "- ((2::rat) * x) \<le> 2*y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   221
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
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
lemma assumes "9*y \<le> 11*x" shows "-99*x \<le> -81 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   224
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   225
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   226
lemma assumes "Numeral1*y \<le> 2*x" shows "-2 * x \<le> -1 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   227
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   228
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   229
lemma assumes "23*y \<le> Numeral1*x" shows "-x \<le> -23 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   230
by (tactic {* test [@{simproc ring_le_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   231
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   232
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   233
subsection {* @{text ring_eq_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   234
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   235
lemma assumes "111*x = -44*y" shows "999*x = -396 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   236
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   237
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   238
lemma assumes "11*x = 9*y" shows "-99*x = -81 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   239
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   240
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   241
lemma assumes "2*x = Numeral1*y" shows "-2 * x = -1 * (y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   242
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   243
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   244
lemma assumes "2*x = Numeral1*y" shows "-2 * x = -(y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   245
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   246
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   247
lemma assumes "3*x = 4*y" shows "9*x = 12 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   248
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   249
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   250
lemma assumes "-3*x = 4*y" shows "-99*x = 132 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   251
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   252
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   253
lemma assumes "111*x = -44*y" shows "999*x = -396 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   254
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   255
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   256
lemma assumes "11*x = 9*y" shows "-99*x = -81 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   257
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   258
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   259
lemma assumes "2*x = Numeral1*y" shows "-2 * x = -1 * (y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   260
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   261
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   262
lemma assumes "2*x = Numeral1*y" shows "-2 * x = -(y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   263
by (tactic {* test [@{simproc ring_eq_cancel_numeral_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   264
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   265
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   266
subsection {* @{text field_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   267
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   268
lemma assumes "(3*x) / (4*y) = z" shows "(9*x) / (12 * (y::rat)) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   269
by (tactic {* test [field_cancel_numeral_factor] *}) fact
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   270
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   271
lemma assumes "(-3*x) / (4*y) = z" shows "(-99*x) / (132 * (y::rat)) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   272
by (tactic {* test [field_cancel_numeral_factor] *}) fact
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   273
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   274
lemma assumes "(111*x) / (-44*y) = z" shows "(999*x) / (-396 * (y::rat)) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   275
by (tactic {* test [field_cancel_numeral_factor] *}) fact
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   276
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   277
lemma assumes "(11*x) / (9*y) = z" shows "(-99*x) / (-81 * (y::rat)) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   278
by (tactic {* test [field_cancel_numeral_factor] *}) fact
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   279
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   280
lemma assumes "(2*x) / (Numeral1*y) = z" shows "(-2 * x) / (-1 * (y::rat)) = z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   281
by (tactic {* test [field_cancel_numeral_factor] *}) fact
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   282
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   283
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   284
subsection {* @{text ring_eq_cancel_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   285
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   286
lemma assumes "k = 0 \<or> x = y" shows "x*k = k*(y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   287
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   288
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   289
lemma assumes "k = 0 \<or> 1 = y" shows "k = k*(y::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   290
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   291
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   292
lemma assumes "b = 0 \<or> a*c = 1" shows "a*(b*c) = (b::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   293
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   294
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   295
lemma assumes "a = 0 \<or> b = 0 \<or> c = d*x" shows "a*(b*c) = d*(b::int)*(x*a)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   296
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   297
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   298
lemma assumes "k = 0 \<or> x = y" shows "x*k = k*(y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   299
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
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
lemma assumes "k = 0 \<or> 1 = y" shows "k = k*(y::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   302
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   303
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   304
lemma assumes "b = 0 \<or> a*c = 1" shows "a*(b*c) = (b::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   305
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   306
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   307
lemma assumes "a = 0 \<or> b = 0 \<or> c = d*x" shows "a*(b*c) = d*(b::rat)*(x*a)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   308
by (tactic {* test [@{simproc ring_eq_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   309
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   310
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   311
subsection {* @{text int_div_cancel_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   312
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   313
lemma assumes "(if k = 0 then 0 else x div y) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   314
  shows "(x*k) div (k*(y::int)) = (uu::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   315
by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   316
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   317
lemma assumes "(if k = 0 then 0 else 1 div y) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   318
  shows "(k) div (k*(y::int)) = (uu::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   319
by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   320
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   321
lemma assumes "(if b = 0 then 0 else a * c) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   322
  shows "(a*(b*c)) div ((b::int)) = (uu::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   323
by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   324
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   325
lemma assumes "(if a = 0 then 0 else if b = 0 then 0 else c div (d * x)) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   326
  shows "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   327
by (tactic {* test [@{simproc int_div_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   328
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   329
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   330
subsection {* @{text divide_cancel_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   331
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   332
lemma assumes "(if k = 0 then 0 else x / y) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   333
  shows "(x*k) / (k*(y::rat)) = (uu::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   334
by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   335
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   336
lemma assumes "(if k = 0 then 0 else 1 / y) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   337
  shows "(k) / (k*(y::rat)) = (uu::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   338
by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   339
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   340
lemma assumes "(if b = 0 then 0 else a * c / 1) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   341
  shows "(a*(b*c)) / ((b::rat)) = (uu::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   342
by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   343
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   344
lemma assumes "(if a = 0 then 0 else if b = 0 then 0 else c / (d * x)) = uu"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   345
  shows "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   346
by (tactic {* test [@{simproc divide_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   347
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   348
lemma shows "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   349
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
   350
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   351
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   352
subsection {* @{text linordered_ring_less_cancel_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   353
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   354
lemma assumes "0 < z \<Longrightarrow> x < y" shows "(0::rat) < z \<Longrightarrow> x*z < y*z"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   355
by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   356
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   357
lemma assumes "0 < z \<Longrightarrow> x < y" shows "(0::rat) < z \<Longrightarrow> x*z < z*y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   358
by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   359
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   360
lemma assumes "0 < z \<Longrightarrow> x < y" shows "(0::rat) < z \<Longrightarrow> z*x < y*z"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   361
by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   362
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45224
diff changeset
   363
lemma assumes "0 < z \<Longrightarrow> x < y" shows "(0::rat) < z \<Longrightarrow> z*x < z*y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   364
by (tactic {* test [@{simproc linordered_ring_less_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   365
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   366
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   367
subsection {* @{text linordered_ring_le_cancel_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   368
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   369
lemma assumes "0 < z \<Longrightarrow> x \<le> y" shows "(0::rat) < z \<Longrightarrow> x*z \<le> y*z"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   370
by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   371
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   372
lemma assumes "0 < z \<Longrightarrow> x \<le> y" shows "(0::rat) < z \<Longrightarrow> z*x \<le> z*y"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   373
by (tactic {* test [@{simproc linordered_ring_le_cancel_factor}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   374
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   375
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   376
subsection {* @{text field_combine_numerals} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   377
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   378
lemma assumes "5 / 6 * x = uu" shows "(x::rat) / 2 + x / 3 = uu"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   379
by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   380
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   381
lemma assumes "6 / 9 * x + y = uu" shows "(x::rat) / 3 + y + x / 3 = uu"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   382
by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   383
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   384
lemma assumes "9 / 9 * x = uu" shows "2 * (x::rat) / 3 + x / 3 = uu"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   385
by (tactic {* test [@{simproc field_combine_numerals}] *}) fact
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   386
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   387
lemma "2/3 * (x::rat) + x / 3 = uu"
45284
ae78a4ffa81d use simproc_setup for cancellation simprocs, to get proper name bindings
huffman
parents: 45270
diff changeset
   388
apply (tactic {* test [@{simproc field_combine_numerals}] *})?
45224
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   389
oops -- "FIXME: test fails"
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   390
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   391
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   392
subsection {* @{text field_eq_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   393
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   394
text {* TODO: tests for @{text field_eq_cancel_numeral_factor} simproc *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   395
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   396
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   397
subsection {* @{text field_cancel_numeral_factor} *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   398
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   399
text {* TODO: tests for @{text field_cancel_numeral_factor} simproc *}
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   400
b1d5b3820d82 add HOL/ex/Simproc_Tests.thy: testing for Tools/numeral_simprocs.ML
huffman
parents:
diff changeset
   401
end