src/HOL/Integ/NatSimprocs.thy
author paulson
Thu, 04 Mar 2004 12:06:07 +0100
changeset 14430 5cb24165a2e1
parent 14387 e96d5c42c4b0
child 14436 77017c49c004
permissions -rw-r--r--
new material from Avigad, and simplified treatment of division by 0
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14128
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     1
(*  Title:      HOL/NatSimprocs.thy
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     2
    ID:         $Id$
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     3
    Copyright   2003 TU Muenchen
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     4
*)
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     5
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     6
header {*Simprocs for the Naturals*}
fd6d20c2371c header comment
paulson
parents: 10536
diff changeset
     7
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 8858
diff changeset
     8
theory NatSimprocs = NatBin
10536
8f34ecae1446 invoking CancelNumeralFactorFun
paulson
parents: 9436
diff changeset
     9
files "int_factor_simprocs.ML" "nat_simprocs.ML":
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 8858
diff changeset
    10
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 8858
diff changeset
    11
setup nat_simprocs_setup
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 8858
diff changeset
    12
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    13
subsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    14
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    15
text{*Where K above is a literal*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    16
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    17
lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    18
by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    19
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    20
(*Now just instantiating n to (number_of v) does the right simplification,
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    21
  but with some redundant inequality tests.*)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    22
lemma neg_number_of_bin_pred_iff_0:
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14353
diff changeset
    23
     "neg (number_of (bin_pred v)::int) = (number_of v = (0::nat))"
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    24
apply (subgoal_tac "neg (number_of (bin_pred v)) = (number_of v < Suc 0) ")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    25
apply (simp only: less_Suc_eq_le le_0_eq)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    26
apply (subst less_number_of_Suc, simp)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    27
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    28
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    29
text{*No longer required as a simprule because of the @{text inverse_fold}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    30
   simproc*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    31
lemma Suc_diff_number_of:
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14353
diff changeset
    32
     "neg (number_of (bin_minus v)::int) ==>  
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    33
      Suc m - (number_of v) = m - (number_of (bin_pred v))"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    34
apply (subst Suc_diff_eq_diff_pred)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    35
apply (simp add: ); 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    36
apply (simp del: nat_numeral_1_eq_1); 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    37
apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] 
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    38
                        neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    39
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    40
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    41
lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    42
by (simp add: numerals split add: nat_diff_split)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    43
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    44
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    45
subsection{*For @{term nat_case} and @{term nat_rec}*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    46
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    47
lemma nat_case_number_of [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    48
     "nat_case a f (number_of v) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    49
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    50
         if neg pv then a else f (nat pv))"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    51
by (simp split add: nat.split add: Let_def neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    52
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    53
lemma nat_case_add_eq_if [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    54
     "nat_case a f ((number_of v) + n) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    55
       (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    56
         if neg pv then nat_case a f n else f (nat pv + n))"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    57
apply (subst add_eq_if)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    58
apply (simp split add: nat.split
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    59
            del: nat_numeral_1_eq_1
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    60
	    add: numeral_1_eq_Suc_0 [symmetric] Let_def 
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    61
                 neg_imp_number_of_eq_0 neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    62
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    63
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    64
lemma nat_rec_number_of [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    65
     "nat_rec a f (number_of v) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    66
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    67
         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    68
apply (case_tac " (number_of v) ::nat")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    69
apply (simp_all (no_asm_simp) add: Let_def neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    70
apply (simp split add: split_if_asm)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    71
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    72
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    73
lemma nat_rec_add_eq_if [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    74
     "nat_rec a f (number_of v + n) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    75
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    76
         if neg pv then nat_rec a f n  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    77
                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    78
apply (subst add_eq_if)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    79
apply (simp split add: nat.split
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    80
            del: nat_numeral_1_eq_1
14273
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    81
            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    82
                 neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    83
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    84
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    85
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    86
subsection{*Various Other Lemmas*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    87
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    88
subsubsection{*Evens and Odds, for Mutilated Chess Board*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    89
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    90
(*Case analysis on b<2*)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    91
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    92
by arith
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    93
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    94
lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    95
apply (subgoal_tac "m mod 2 < 2")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    96
apply (erule less_2_cases [THEN disjE])
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    97
apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    98
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    99
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   100
lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   101
apply (subgoal_tac "m mod 2 < 2")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   102
apply (force simp del: mod_less_divisor, simp) 
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   103
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   104
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   105
subsubsection{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   106
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   107
lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   108
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   109
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   110
lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   111
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   112
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   113
text{*Can be used to eliminate long strings of Sucs, but not by default*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   114
lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   115
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   116
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   117
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   118
text{*These lemmas collapse some needless occurrences of Suc:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   119
    at least three Sucs, since two and fewer are rewritten back to Suc again!
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   120
    We already have some rules to simplify operands smaller than 3.*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   121
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   122
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   123
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   124
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   125
lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   126
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   127
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   128
lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   129
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   130
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   131
lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   132
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   133
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   134
declare Suc_div_eq_add3_div [of _ "number_of v", standard, simp]
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   135
declare Suc_mod_eq_add3_mod [of _ "number_of v", standard, simp]
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   136
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   137
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   138
subsection{*Special Simplification for Constants*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   139
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   140
text{*These belong here, late in the development of HOL, to prevent their
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   141
interfering with proofs of abstract properties of instances of the function
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   142
@{term number_of}*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   143
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   144
text{*These distributive laws move literals inside sums and differences.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   145
declare left_distrib [of _ _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   146
declare right_distrib [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   147
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   148
declare left_diff_distrib [of _ _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   149
declare right_diff_distrib [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   150
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   151
text{*These are actually for fields, like real: but where else to put them?*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   152
declare zero_less_divide_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   153
declare divide_less_0_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   154
declare zero_le_divide_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   155
declare divide_le_0_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   156
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14288
diff changeset
   157
(*Replaces "inverse #nn" by 1/#nn.  It looks strange, but then other simprocs
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14288
diff changeset
   158
simplify the quotient.*)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   159
declare inverse_eq_divide [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   160
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   161
text{*These laws simplify inequalities, moving unary minus from a term
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   162
into the literal.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   163
declare less_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   164
declare le_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   165
declare equation_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   166
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   167
declare minus_less_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   168
declare minus_le_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   169
declare minus_equation_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   170
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   171
text{*These simplify inequalities where one side is the constant 1.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   172
declare less_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   173
declare le_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   174
declare equation_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   175
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   176
declare minus_less_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   177
declare minus_le_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   178
declare minus_equation_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   179
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   180
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   181
(*Cancellation of constant factors in comparisons (< and \<le>) *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   182
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   183
declare mult_less_cancel_left [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   184
declare mult_less_cancel_right [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   185
declare mult_le_cancel_left [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   186
declare mult_le_cancel_right [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   187
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   188
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   189
(*Multiplying out constant divisors in comparisons (< \<le> and =) *)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   190
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   191
declare le_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   192
declare divide_le_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   193
declare less_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   194
declare divide_less_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   195
declare eq_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   196
declare divide_eq_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   197
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   198
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   199
subsubsection{*Division By @{term "-1"}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   200
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   201
lemma divide_minus1 [simp]:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   202
     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   203
by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   204
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   205
lemma minus1_divide [simp]:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   206
     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14387
diff changeset
   207
by (simp add: divide_inverse inverse_minus_eq)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   208
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   209
ML
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   210
{*
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   211
val divide_minus1 = thm "divide_minus1";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   212
val minus1_divide = thm "minus1_divide";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   213
*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   214
8858
b739f0ecc1fa new dummy theory; prevents strange errors when loading NatSimprocs.ML
paulson
parents:
diff changeset
   215
end