src/HOL/Integ/NatSimprocs.thy
author paulson
Tue, 10 Feb 2004 12:02:11 +0100
changeset 14378 69c4d5997669
parent 14353 79f9fbef9106
child 14387 e96d5c42c4b0
permissions -rw-r--r--
generic of_nat and of_int functions, and generalization of iszero and neg
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))"
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    34
apply (subst Suc_diff_eq_diff_pred, simp, simp)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    35
apply (force simp only: diff_nat_number_of less_0_number_of [symmetric] 
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    36
                        neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    37
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    38
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    39
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
    40
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
    41
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    42
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    43
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
    44
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    45
lemma nat_case_number_of [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    46
     "nat_case a f (number_of v) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    47
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    48
         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
    49
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
    50
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    51
lemma nat_case_add_eq_if [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    52
     "nat_case a f ((number_of v) + n) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    53
       (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    54
         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
    55
apply (subst add_eq_if)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    56
apply (simp split add: nat.split
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    57
            add: numeral_1_eq_Suc_0 [symmetric] Let_def 
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    58
                 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
    59
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    60
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    61
lemma nat_rec_number_of [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    62
     "nat_rec a f (number_of v) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    63
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    64
         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
    65
apply (case_tac " (number_of v) ::nat")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    66
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
    67
apply (simp split add: split_if_asm)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    68
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    69
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    70
lemma nat_rec_add_eq_if [simp]:
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    71
     "nat_rec a f (number_of v + n) =  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    72
        (let pv = number_of (bin_pred v) in  
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    73
         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
    74
                   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
    75
apply (subst add_eq_if)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    76
apply (simp split add: nat.split
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    77
            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
    78
                 neg_number_of_bin_pred_iff_0)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    79
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    80
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    81
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    82
subsection{*Various Other Lemmas*}
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    83
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    84
subsubsection{*Evens and Odds, for Mutilated Chess Board*}
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
(*Case analysis on b<2*)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    87
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
    88
by arith
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
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
    91
apply (subgoal_tac "m mod 2 < 2")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    92
apply (erule less_2_cases [THEN disjE])
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    93
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
    94
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    95
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    96
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
    97
apply (subgoal_tac "m mod 2 < 2")
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    98
apply (force simp del: mod_less_divisor, simp) 
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
    99
done
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   100
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   101
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
   102
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   103
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
   104
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   105
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   106
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
   107
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   108
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   109
declare numeral_0_eq_0 [simp] numeral_1_eq_1 [simp] 
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   110
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   111
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
   112
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
   113
by simp
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   114
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   115
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   116
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
   117
    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
   118
    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
   119
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   120
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
   121
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   122
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   123
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
   124
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   125
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   126
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
   127
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   128
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   129
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
   130
by (simp add: Suc3_eq_add_3)
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   131
e33ffff0123c further simplifications of the integer development; converting more .ML files
paulson
parents: 14128
diff changeset
   132
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
   133
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
   134
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   135
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   136
subsection{*Special Simplification for Constants*}
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
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
   139
interfering with proofs of abstract properties of instances of the function
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   140
@{term number_of}*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   141
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   142
text{*These distributive laws move literals inside sums and differences.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   143
declare left_distrib [of _ _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   144
declare right_distrib [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   145
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   146
declare left_diff_distrib [of _ _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   147
declare right_diff_distrib [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   148
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   149
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
   150
declare zero_less_divide_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   151
declare divide_less_0_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   152
declare zero_le_divide_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   153
declare divide_le_0_iff [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   154
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14288
diff changeset
   155
(*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
   156
simplify the quotient.*)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   157
declare inverse_eq_divide [of "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   158
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   159
text{*These laws simplify inequalities, moving unary minus from a term
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   160
into the literal.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   161
declare less_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   162
declare le_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   163
declare equation_minus_iff [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   164
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   165
declare minus_less_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   166
declare minus_le_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   167
declare minus_equation_iff [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   168
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   169
text{*These simplify inequalities where one side is the constant 1.*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   170
declare less_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   171
declare le_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   172
declare equation_minus_iff [of 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   173
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   174
declare minus_less_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   175
declare minus_le_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   176
declare minus_equation_iff [of _ 1, simplified, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   177
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   178
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   179
(*Cancellation of constant factors in comparisons (< and \<le>) *)
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
declare mult_less_cancel_left [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   182
declare mult_less_cancel_right [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   183
declare mult_le_cancel_left [of "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   184
declare mult_le_cancel_right [of _ "number_of v", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   185
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   186
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   187
(*Multiplying out constant divisors in comparisons (< \<le> and =) *)
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
declare le_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   190
declare divide_le_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   191
declare less_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   192
declare divide_less_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   193
declare eq_divide_eq [of _ _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   194
declare divide_eq_eq [of _ "number_of w", standard, simp]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   195
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14273
diff changeset
   196
8858
b739f0ecc1fa new dummy theory; prevents strange errors when loading NatSimprocs.ML
paulson
parents:
diff changeset
   197
end