src/HOL/Arith_Tools.thy
author nipkow
Fri, 20 Feb 2009 23:46:03 +0100
changeset 30031 bd786c37af84
parent 29012 9140227dc8c5
child 30079 293b896b9c25
permissions -rw-r--r--
Removed redundant lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Arith_Tools.thy
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     4
    Author:     Amine Chaieb, TU Muenchen
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     5
*)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     6
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     7
header {* Setup of arithmetic tools *}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     8
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
     9
theory Arith_Tools
28402
09e4aa3ddc25 clarified dependencies between arith tools
haftmann
parents: 26462
diff changeset
    10
imports NatBin
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    11
uses
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    12
  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    13
  "~~/src/Provers/Arith/extract_common_term.ML"
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28402
diff changeset
    14
  "Tools/int_factor_simprocs.ML"
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28402
diff changeset
    15
  "Tools/nat_simprocs.ML"
28402
09e4aa3ddc25 clarified dependencies between arith tools
haftmann
parents: 26462
diff changeset
    16
  "Tools/Qelim/qelim.ML"
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    17
begin
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    18
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    19
subsection {* Simprocs for the Naturals *}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    20
24075
366d4d234814 arith method setup: proper context;
wenzelm
parents: 23901
diff changeset
    21
declaration {* K nat_simprocs_setup *}
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    22
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    23
subsubsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    24
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    25
text{*Where K above is a literal*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    26
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    27
lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    28
by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    29
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    30
text {*Now just instantiating @{text n} to @{text "number_of v"} does
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    31
  the right simplification, but with some redundant inequality
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    32
  tests.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    33
lemma neg_number_of_pred_iff_0:
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    34
  "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))"
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    35
apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ")
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    36
apply (simp only: less_Suc_eq_le le_0_eq)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    37
apply (subst less_number_of_Suc, simp)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    38
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    39
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    40
text{*No longer required as a simprule because of the @{text inverse_fold}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    41
   simproc*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    42
lemma Suc_diff_number_of:
29012
9140227dc8c5 change lemmas to avoid using neg
huffman
parents: 28952
diff changeset
    43
     "Int.Pls < v ==>
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    44
      Suc m - (number_of v) = m - (number_of (Int.pred v))"
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    45
apply (subst Suc_diff_eq_diff_pred)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    46
apply simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    47
apply (simp del: nat_numeral_1_eq_1)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    48
apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    49
                        neg_number_of_pred_iff_0)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    50
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    51
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    52
lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    53
by (simp add: numerals split add: nat_diff_split)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    54
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    55
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    56
subsubsection{*For @{term nat_case} and @{term nat_rec}*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    57
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    58
lemma nat_case_number_of [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    59
     "nat_case a f (number_of v) =
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    60
        (let pv = number_of (Int.pred v) in
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    61
         if neg pv then a else f (nat pv))"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    62
by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    63
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    64
lemma nat_case_add_eq_if [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    65
     "nat_case a f ((number_of v) + n) =
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    66
       (let pv = number_of (Int.pred v) in
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    67
         if neg pv then nat_case a f n else f (nat pv + n))"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    68
apply (subst add_eq_if)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    69
apply (simp split add: nat.split
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    70
            del: nat_numeral_1_eq_1
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    71
            add: numeral_1_eq_Suc_0 [symmetric] Let_def
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    72
                 neg_imp_number_of_eq_0 neg_number_of_pred_iff_0)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    73
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    74
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    75
lemma nat_rec_number_of [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    76
     "nat_rec a f (number_of v) =
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    77
        (let pv = number_of (Int.pred v) in
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    78
         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    79
apply (case_tac " (number_of v) ::nat")
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    80
apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    81
apply (simp split add: split_if_asm)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    82
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    83
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    84
lemma nat_rec_add_eq_if [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    85
     "nat_rec a f (number_of v + n) =
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
    86
        (let pv = number_of (Int.pred v) in
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    87
         if neg pv then nat_rec a f n
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    88
                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    89
apply (subst add_eq_if)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    90
apply (simp split add: nat.split
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    91
            del: nat_numeral_1_eq_1
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    92
            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    93
                 neg_number_of_pred_iff_0)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    94
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    95
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    96
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    97
subsubsection{*Various Other Lemmas*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    98
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
    99
text {*Evens and Odds, for Mutilated Chess Board*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   100
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   101
text{*Lemmas for specialist use, NOT as default simprules*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   102
lemma nat_mult_2: "2 * z = (z+z::nat)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   103
proof -
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   104
  have "2*z = (1 + 1)*z" by simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   105
  also have "... = z+z" by (simp add: left_distrib)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   106
  finally show ?thesis .
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   107
qed
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   108
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   109
lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   110
by (subst mult_commute, rule nat_mult_2)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   111
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   112
text{*Case analysis on @{term "n<2"}*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   113
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   114
by arith
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   115
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   116
lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   117
by arith
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   118
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   119
lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   120
by (simp add: nat_mult_2 [symmetric])
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   121
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   122
lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   123
apply (subgoal_tac "m mod 2 < 2")
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   124
apply (erule less_2_cases [THEN disjE])
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   125
apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   126
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   127
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   128
lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   129
apply (subgoal_tac "m mod 2 < 2")
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   130
apply (force simp del: mod_less_divisor, simp)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   131
done
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   132
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   133
text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   134
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   135
lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   136
by simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   137
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   138
lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   139
by simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   140
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   141
text{*Can be used to eliminate long strings of Sucs, but not by default*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   142
lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   143
by simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   144
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   145
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   146
text{*These lemmas collapse some needless occurrences of Suc:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   147
    at least three Sucs, since two and fewer are rewritten back to Suc again!
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   148
    We already have some rules to simplify operands smaller than 3.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   149
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   150
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   151
by (simp add: Suc3_eq_add_3)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   152
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   153
lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   154
by (simp add: Suc3_eq_add_3)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   155
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   156
lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   157
by (simp add: Suc3_eq_add_3)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   158
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   159
lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   160
by (simp add: Suc3_eq_add_3)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   161
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   162
lemmas Suc_div_eq_add3_div_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   163
    Suc_div_eq_add3_div [of _ "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   164
declare Suc_div_eq_add3_div_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   165
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   166
lemmas Suc_mod_eq_add3_mod_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   167
    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   168
declare Suc_mod_eq_add3_mod_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   169
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   170
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   171
subsubsection{*Special Simplification for Constants*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   172
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   173
text{*These belong here, late in the development of HOL, to prevent their
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   174
interfering with proofs of abstract properties of instances of the function
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   175
@{term number_of}*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   176
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   177
text{*These distributive laws move literals inside sums and differences.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   178
lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   179
declare left_distrib_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   180
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   181
lemmas right_distrib_number_of = right_distrib [of "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   182
declare right_distrib_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   183
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   184
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   185
lemmas left_diff_distrib_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   186
    left_diff_distrib [of _ _ "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   187
declare left_diff_distrib_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   188
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   189
lemmas right_diff_distrib_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   190
    right_diff_distrib [of "number_of v", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   191
declare right_diff_distrib_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   192
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   193
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   194
text{*These are actually for fields, like real: but where else to put them?*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   195
lemmas zero_less_divide_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   196
    zero_less_divide_iff [of "number_of w", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   197
declare zero_less_divide_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   198
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   199
lemmas divide_less_0_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   200
    divide_less_0_iff [of "number_of w", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   201
declare divide_less_0_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   202
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   203
lemmas zero_le_divide_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   204
    zero_le_divide_iff [of "number_of w", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   205
declare zero_le_divide_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   206
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   207
lemmas divide_le_0_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   208
    divide_le_0_iff [of "number_of w", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   209
declare divide_le_0_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   210
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   211
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   212
(****
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   213
IF times_divide_eq_right and times_divide_eq_left are removed as simprules,
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   214
then these special-case declarations may be useful.
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   215
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   216
text{*These simprules move numerals into numerators and denominators.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   217
lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   218
by (simp add: times_divide_eq)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   219
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   220
lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   221
by (simp add: times_divide_eq)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   222
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   223
lemmas times_divide_eq_right_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   224
    times_divide_eq_right [of "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   225
declare times_divide_eq_right_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   226
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   227
lemmas times_divide_eq_right_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   228
    times_divide_eq_right [of _ _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   229
declare times_divide_eq_right_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   230
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   231
lemmas times_divide_eq_left_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   232
    times_divide_eq_left [of _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   233
declare times_divide_eq_left_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   234
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   235
lemmas times_divide_eq_left_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   236
    times_divide_eq_left [of _ _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   237
declare times_divide_eq_left_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   238
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   239
****)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   240
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   241
text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   242
  strange, but then other simprocs simplify the quotient.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   243
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   244
lemmas inverse_eq_divide_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   245
    inverse_eq_divide [of "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   246
declare inverse_eq_divide_number_of [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   247
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   248
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   249
text {*These laws simplify inequalities, moving unary minus from a term
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   250
into the literal.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   251
lemmas less_minus_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   252
    less_minus_iff [of "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   253
declare less_minus_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   254
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   255
lemmas le_minus_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   256
    le_minus_iff [of "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   257
declare le_minus_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   258
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   259
lemmas equation_minus_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   260
    equation_minus_iff [of "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   261
declare equation_minus_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   262
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   263
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   264
lemmas minus_less_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   265
    minus_less_iff [of _ "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   266
declare minus_less_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   267
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   268
lemmas minus_le_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   269
    minus_le_iff [of _ "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   270
declare minus_le_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   271
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   272
lemmas minus_equation_iff_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   273
    minus_equation_iff [of _ "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   274
declare minus_equation_iff_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   275
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   276
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   277
text{*To Simplify Inequalities Where One Side is the Constant 1*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   278
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   279
lemma less_minus_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   280
  fixes b::"'b::{ordered_idom,number_ring}"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   281
  shows "(1 < - b) = (b < -1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   282
by auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   283
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   284
lemma le_minus_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   285
  fixes b::"'b::{ordered_idom,number_ring}"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   286
  shows "(1 \<le> - b) = (b \<le> -1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   287
by auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   288
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   289
lemma equation_minus_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   290
  fixes b::"'b::number_ring"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   291
  shows "(1 = - b) = (b = -1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   292
by (subst equation_minus_iff, auto)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   293
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   294
lemma minus_less_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   295
  fixes a::"'b::{ordered_idom,number_ring}"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   296
  shows "(- a < 1) = (-1 < a)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   297
by auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   298
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   299
lemma minus_le_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   300
  fixes a::"'b::{ordered_idom,number_ring}"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   301
  shows "(- a \<le> 1) = (-1 \<le> a)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   302
by auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   303
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   304
lemma minus_equation_iff_1 [simp,noatp]:
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   305
  fixes a::"'b::number_ring"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   306
  shows "(- a = 1) = (a = -1)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   307
by (subst minus_equation_iff, auto)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   308
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   309
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   310
text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   311
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   312
lemmas mult_less_cancel_left_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   313
    mult_less_cancel_left [of "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   314
declare mult_less_cancel_left_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   315
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   316
lemmas mult_less_cancel_right_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   317
    mult_less_cancel_right [of _ "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   318
declare mult_less_cancel_right_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   319
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   320
lemmas mult_le_cancel_left_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   321
    mult_le_cancel_left [of "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   322
declare mult_le_cancel_left_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   323
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   324
lemmas mult_le_cancel_right_number_of =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   325
    mult_le_cancel_right [of _ "number_of v", standard]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 24075
diff changeset
   326
declare mult_le_cancel_right_number_of [simp,noatp]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   327
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   328
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   329
text {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   330
26314
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   331
lemmas le_divide_eq_number_of1 [simp] = le_divide_eq [of _ _ "number_of w", standard]
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   332
lemmas divide_le_eq_number_of1 [simp] = divide_le_eq [of _ "number_of w", standard]
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   333
lemmas less_divide_eq_number_of1 [simp] = less_divide_eq [of _ _ "number_of w", standard]
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   334
lemmas divide_less_eq_number_of1 [simp] = divide_less_eq [of _ "number_of w", standard]
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   335
lemmas eq_divide_eq_number_of1 [simp] = eq_divide_eq [of _ _ "number_of w", standard]
9c39fc898fff avoid rebinding of existing facts;
wenzelm
parents: 26154
diff changeset
   336
lemmas divide_eq_eq_number_of1 [simp] = divide_eq_eq [of _ "number_of w", standard]
23462
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   337
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   338
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   339
subsubsection{*Optional Simplification Rules Involving Constants*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   340
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   341
text{*Simplify quotients that are compared with a literal constant.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   342
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   343
lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   344
lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   345
lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   346
lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   347
lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   348
lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   349
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   350
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   351
text{*Not good as automatic simprules because they cause case splits.*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   352
lemmas divide_const_simps =
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   353
  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   354
  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   355
  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   356
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   357
text{*Division By @{text "-1"}*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   358
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   359
lemma divide_minus1 [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   360
     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   361
by simp
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   362
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   363
lemma minus1_divide [simp]:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   364
     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   365
by (simp add: divide_inverse inverse_minus_eq)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   366
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   367
lemma half_gt_zero_iff:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   368
     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   369
by auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   370
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   371
lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   372
declare half_gt_zero [simp]
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   373
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   374
(* The following lemma should appear in Divides.thy, but there the proof
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   375
   doesn't work. *)
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   376
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   377
lemma nat_dvd_not_less:
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   378
  "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)"
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   379
  by (unfold dvd_def) auto
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   380
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   381
ML {*
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   382
val divide_minus1 = @{thm divide_minus1};
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   383
val minus1_divide = @{thm minus1_divide};
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   384
*}
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   385
11728d83794c renamed NatSimprocs.thy to Arith_Tools.thy;
wenzelm
parents:
diff changeset
   386
end