src/HOL/NatSimprocs.thy
author wenzelm
Wed, 13 Jun 2007 00:01:57 +0200
changeset 23354 a189707c1d76
parent 23329 0dbb30302259
child 23404 8659acd81f9d
permissions -rw-r--r--
Context positions.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/NatSimprocs.thy
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
    Copyright   2003 TU Muenchen
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     6
header {*Simprocs for the Naturals*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     7
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     8
theory NatSimprocs
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents: 23164
diff changeset
     9
imports Groebner_Basis
23263
0c227412b285 tuned boostrap
haftmann
parents: 23252
diff changeset
    10
uses
0c227412b285 tuned boostrap
haftmann
parents: 23252
diff changeset
    11
  "~~/src/Provers/Arith/cancel_numeral_factor.ML"
0c227412b285 tuned boostrap
haftmann
parents: 23252
diff changeset
    12
  "~~/src/Provers/Arith/extract_common_term.ML"
0c227412b285 tuned boostrap
haftmann
parents: 23252
diff changeset
    13
  "int_factor_simprocs.ML"
0c227412b285 tuned boostrap
haftmann
parents: 23252
diff changeset
    14
  "nat_simprocs.ML"
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
begin
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    16
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    17
setup nat_simprocs_setup
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    18
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    19
subsection{*For simplifying @{term "Suc m - K"} and  @{term "K - Suc m"}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    20
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    21
text{*Where K above is a literal*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    22
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    23
lemma Suc_diff_eq_diff_pred: "Numeral0 < n ==> Suc m - n = m - (n - Numeral1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    24
by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    25
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    26
text {*Now just instantiating @{text n} to @{text "number_of v"} does
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    27
  the right simplification, but with some redundant inequality
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    28
  tests.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    29
lemma neg_number_of_pred_iff_0:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    30
  "neg (number_of (Numeral.pred v)::int) = (number_of v = (0::nat))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    31
apply (subgoal_tac "neg (number_of (Numeral.pred v)) = (number_of v < Suc 0) ")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    32
apply (simp only: less_Suc_eq_le le_0_eq)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    33
apply (subst less_number_of_Suc, simp)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    34
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    35
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    36
text{*No longer required as a simprule because of the @{text inverse_fold}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    37
   simproc*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    38
lemma Suc_diff_number_of:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    39
     "neg (number_of (uminus v)::int) ==>  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    40
      Suc m - (number_of v) = m - (number_of (Numeral.pred v))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    41
apply (subst Suc_diff_eq_diff_pred)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    42
apply simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    43
apply (simp del: nat_numeral_1_eq_1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    44
apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    45
                        neg_number_of_pred_iff_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    46
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    47
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    48
lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    49
by (simp add: numerals split add: nat_diff_split)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    50
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    51
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    52
subsection{*For @{term nat_case} and @{term nat_rec}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    53
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    54
lemma nat_case_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    55
     "nat_case a f (number_of v) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    56
        (let pv = number_of (Numeral.pred v) in  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    57
         if neg pv then a else f (nat pv))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    58
by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    59
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    60
lemma nat_case_add_eq_if [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    61
     "nat_case a f ((number_of v) + n) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    62
       (let pv = number_of (Numeral.pred v) in  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    63
         if neg pv then nat_case a f n else f (nat pv + n))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    64
apply (subst add_eq_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    65
apply (simp split add: nat.split
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    66
            del: nat_numeral_1_eq_1
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    67
	    add: numeral_1_eq_Suc_0 [symmetric] Let_def 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    68
                 neg_imp_number_of_eq_0 neg_number_of_pred_iff_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    69
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    70
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    71
lemma nat_rec_number_of [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    72
     "nat_rec a f (number_of v) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    73
        (let pv = number_of (Numeral.pred v) in  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    74
         if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    75
apply (case_tac " (number_of v) ::nat")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    76
apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    77
apply (simp split add: split_if_asm)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    78
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    79
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    80
lemma nat_rec_add_eq_if [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    81
     "nat_rec a f (number_of v + n) =  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    82
        (let pv = number_of (Numeral.pred v) in  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    83
         if neg pv then nat_rec a f n  
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    84
                   else f (nat pv + n) (nat_rec a f (nat pv + n)))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    85
apply (subst add_eq_if)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
apply (simp split add: nat.split
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    87
            del: nat_numeral_1_eq_1
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    88
            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    89
                 neg_number_of_pred_iff_0)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    90
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    91
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    93
subsection{*Various Other Lemmas*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    94
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    95
subsubsection{*Evens and Odds, for Mutilated Chess Board*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    96
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
text{*Lemmas for specialist use, NOT as default simprules*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    98
lemma nat_mult_2: "2 * z = (z+z::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    99
proof -
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
  have "2*z = (1 + 1)*z" by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   101
  also have "... = z+z" by (simp add: left_distrib)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
  finally show ?thesis .
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   103
qed
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   104
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   105
lemma nat_mult_2_right: "z * 2 = (z+z::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   106
by (subst mult_commute, rule nat_mult_2)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   107
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
text{*Case analysis on @{term "n<2"}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
by arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   111
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   113
by arith
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   114
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   115
lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   116
by (simp add: nat_mult_2 [symmetric])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   117
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   118
lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   119
apply (subgoal_tac "m mod 2 < 2")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   120
apply (erule less_2_cases [THEN disjE])
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   121
apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   122
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   123
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   124
lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   125
apply (subgoal_tac "m mod 2 < 2")
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   126
apply (force simp del: mod_less_divisor, simp) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   127
done
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   128
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   129
subsubsection{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   130
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   131
lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   132
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   133
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   134
lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   135
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   136
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   137
text{*Can be used to eliminate long strings of Sucs, but not by default*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   138
lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   139
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   140
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   141
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   142
text{*These lemmas collapse some needless occurrences of Suc:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   143
    at least three Sucs, since two and fewer are rewritten back to Suc again!
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   144
    We already have some rules to simplify operands smaller than 3.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   145
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   146
lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   147
by (simp add: Suc3_eq_add_3)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   148
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   149
lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   150
by (simp add: Suc3_eq_add_3)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   151
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   152
lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   153
by (simp add: Suc3_eq_add_3)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   154
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   155
lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   156
by (simp add: Suc3_eq_add_3)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   157
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   158
lemmas Suc_div_eq_add3_div_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   159
    Suc_div_eq_add3_div [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   160
declare Suc_div_eq_add3_div_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   161
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   162
lemmas Suc_mod_eq_add3_mod_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   163
    Suc_mod_eq_add3_mod [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   164
declare Suc_mod_eq_add3_mod_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   165
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   166
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   167
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   168
subsection{*Special Simplification for Constants*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   169
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   170
text{*These belong here, late in the development of HOL, to prevent their
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   171
interfering with proofs of abstract properties of instances of the function
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   172
@{term number_of}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   173
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   174
text{*These distributive laws move literals inside sums and differences.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   175
lemmas left_distrib_number_of = left_distrib [of _ _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   176
declare left_distrib_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   177
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   178
lemmas right_distrib_number_of = right_distrib [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   179
declare right_distrib_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   180
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   181
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   182
lemmas left_diff_distrib_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   183
    left_diff_distrib [of _ _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   184
declare left_diff_distrib_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   185
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   186
lemmas right_diff_distrib_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   187
    right_diff_distrib [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   188
declare right_diff_distrib_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   189
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   190
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   191
text{*These are actually for fields, like real: but where else to put them?*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   192
lemmas zero_less_divide_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   193
    zero_less_divide_iff [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   194
declare zero_less_divide_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   195
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   196
lemmas divide_less_0_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   197
    divide_less_0_iff [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   198
declare divide_less_0_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   199
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   200
lemmas zero_le_divide_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   201
    zero_le_divide_iff [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   202
declare zero_le_divide_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   203
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   204
lemmas divide_le_0_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   205
    divide_le_0_iff [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   206
declare divide_le_0_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   207
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   208
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   209
(****
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   210
IF times_divide_eq_right and times_divide_eq_left are removed as simprules,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   211
then these special-case declarations may be useful.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   212
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   213
text{*These simprules move numerals into numerators and denominators.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   214
lemma times_recip_eq_right [simp]: "a * (1/c) = a / (c::'a::field)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   215
by (simp add: times_divide_eq)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   216
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   217
lemma times_recip_eq_left [simp]: "(1/c) * a = a / (c::'a::field)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   218
by (simp add: times_divide_eq)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   219
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   220
lemmas times_divide_eq_right_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   221
    times_divide_eq_right [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   222
declare times_divide_eq_right_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   223
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   224
lemmas times_divide_eq_right_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   225
    times_divide_eq_right [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   226
declare times_divide_eq_right_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   227
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   228
lemmas times_divide_eq_left_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   229
    times_divide_eq_left [of _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   230
declare times_divide_eq_left_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   231
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   232
lemmas times_divide_eq_left_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   233
    times_divide_eq_left [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   234
declare times_divide_eq_left_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   235
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   236
****)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   237
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   238
text {*Replaces @{text "inverse #nn"} by @{text "1/#nn"}.  It looks
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   239
  strange, but then other simprocs simplify the quotient.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   240
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   241
lemmas inverse_eq_divide_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   242
    inverse_eq_divide [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   243
declare inverse_eq_divide_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   244
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   245
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   246
subsubsection{*These laws simplify inequalities, moving unary minus from a term
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   247
into the literal.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   248
lemmas less_minus_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   249
    less_minus_iff [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   250
declare less_minus_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   251
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   252
lemmas le_minus_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   253
    le_minus_iff [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   254
declare le_minus_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   255
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   256
lemmas equation_minus_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   257
    equation_minus_iff [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   258
declare equation_minus_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   259
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   260
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   261
lemmas minus_less_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   262
    minus_less_iff [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   263
declare minus_less_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   264
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   265
lemmas minus_le_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   266
    minus_le_iff [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   267
declare minus_le_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   268
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   269
lemmas minus_equation_iff_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   270
    minus_equation_iff [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   271
declare minus_equation_iff_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   272
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   273
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   274
subsubsection{*To Simplify Inequalities Where One Side is the Constant 1*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   275
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   276
lemma less_minus_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   277
  fixes b::"'b::{ordered_idom,number_ring}" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   278
  shows "(1 < - b) = (b < -1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   279
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   280
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   281
lemma le_minus_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   282
  fixes b::"'b::{ordered_idom,number_ring}" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   283
  shows "(1 \<le> - b) = (b \<le> -1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   284
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   285
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   286
lemma equation_minus_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   287
  fixes b::"'b::number_ring" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   288
  shows "(1 = - b) = (b = -1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   289
by (subst equation_minus_iff, auto) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   290
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   291
lemma minus_less_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   292
  fixes a::"'b::{ordered_idom,number_ring}" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   293
  shows "(- a < 1) = (-1 < a)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   294
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   295
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   296
lemma minus_le_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   297
  fixes a::"'b::{ordered_idom,number_ring}" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   298
  shows "(- a \<le> 1) = (-1 \<le> a)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   299
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   300
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   301
lemma minus_equation_iff_1 [simp]: 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   302
  fixes a::"'b::number_ring" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   303
  shows "(- a = 1) = (a = -1)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   304
by (subst minus_equation_iff, auto) 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   305
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   306
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   307
subsubsection {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   308
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   309
lemmas mult_less_cancel_left_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   310
    mult_less_cancel_left [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   311
declare mult_less_cancel_left_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   312
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   313
lemmas mult_less_cancel_right_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   314
    mult_less_cancel_right [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   315
declare mult_less_cancel_right_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   316
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   317
lemmas mult_le_cancel_left_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   318
    mult_le_cancel_left [of "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   319
declare mult_le_cancel_left_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   320
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   321
lemmas mult_le_cancel_right_number_of =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   322
    mult_le_cancel_right [of _ "number_of v", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   323
declare mult_le_cancel_right_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   324
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   325
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   326
subsubsection {*Multiplying out constant divisors in comparisons (@{text "<"}, @{text "\<le>"} and @{text "="}) *}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   327
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   328
lemmas le_divide_eq_number_of = le_divide_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   329
declare le_divide_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   330
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   331
lemmas divide_le_eq_number_of = divide_le_eq [of _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   332
declare divide_le_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   333
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   334
lemmas less_divide_eq_number_of = less_divide_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   335
declare less_divide_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   336
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   337
lemmas divide_less_eq_number_of = divide_less_eq [of _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   338
declare divide_less_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   339
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   340
lemmas eq_divide_eq_number_of = eq_divide_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   341
declare eq_divide_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   342
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   343
lemmas divide_eq_eq_number_of = divide_eq_eq [of _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   344
declare divide_eq_eq_number_of [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   345
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   346
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   347
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   348
subsection{*Optional Simplification Rules Involving Constants*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   349
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   350
text{*Simplify quotients that are compared with a literal constant.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   351
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   352
lemmas le_divide_eq_number_of = le_divide_eq [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   353
lemmas divide_le_eq_number_of = divide_le_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   354
lemmas less_divide_eq_number_of = less_divide_eq [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   355
lemmas divide_less_eq_number_of = divide_less_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   356
lemmas eq_divide_eq_number_of = eq_divide_eq [of "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   357
lemmas divide_eq_eq_number_of = divide_eq_eq [of _ _ "number_of w", standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   358
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   359
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   360
text{*Not good as automatic simprules because they cause case splits.*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   361
lemmas divide_const_simps =
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   362
  le_divide_eq_number_of divide_le_eq_number_of less_divide_eq_number_of
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   363
  divide_less_eq_number_of eq_divide_eq_number_of divide_eq_eq_number_of
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   364
  le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   365
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   366
subsubsection{*Division By @{text "-1"}*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   367
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   368
lemma divide_minus1 [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   369
     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   370
by simp
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   371
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   372
lemma minus1_divide [simp]:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   373
     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   374
by (simp add: divide_inverse inverse_minus_eq)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   375
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   376
lemma half_gt_zero_iff:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   377
     "(0 < r/2) = (0 < (r::'a::{ordered_field,division_by_zero,number_ring}))"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   378
by auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   379
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   380
lemmas half_gt_zero = half_gt_zero_iff [THEN iffD2, standard]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   381
declare half_gt_zero [simp]
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   382
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   383
(* The following lemma should appear in Divides.thy, but there the proof
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   384
   doesn't work. *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   385
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   386
lemma nat_dvd_not_less:
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   387
  "[| 0 < m; m < n |] ==> \<not> n dvd (m::nat)"
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   388
  by (unfold dvd_def) auto
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   389
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   390
ML {*
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   391
val divide_minus1 = @{thm divide_minus1};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   392
val minus1_divide = @{thm minus1_divide};
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   393
*}
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   394
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   395
section{* Installing Groebner Bases for Fields *}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   396
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   397
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   398
interpretation class_fieldgb: 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   399
  fieldgb["op +" "op *" "op ^" "0::'a::{field,recpower,number_ring}" "1" "op -" "uminus" "op /" "inverse"] apply (unfold_locales) by (simp_all add: divide_inverse)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   400
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   401
lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   402
lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   403
  by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   404
lemma mult_frac_frac: "((x::'a::{field,division_by_zero}) / y) * (z / w) = (x*z) / (y*w)" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   405
  by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   406
lemma mult_frac_num: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   407
  by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   408
lemma mult_num_frac: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   409
  by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   410
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   411
lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   412
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   413
lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_by_zero}) / y + z = (x + z*y) / y" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   414
  by (simp add: add_divide_distrib)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   415
lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_by_zero}) / y = (x + z*y) / y" 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   416
  by (simp add: add_divide_distrib)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   417
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   418
declaration{*
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   419
let
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   420
 val zr = @{cpat "0"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   421
 val zT = ctyp_of_term zr
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   422
 val geq = @{cpat "op ="}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   423
 val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   424
 val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"} 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   425
 val add_frac_num = mk_meta_eq @{thm "add_frac_num"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   426
 val add_num_frac = mk_meta_eq @{thm "add_num_frac"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   427
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   428
 fun prove_nz ctxt =
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   429
  let val ss = local_simpset_of ctxt 
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   430
  in fn T => fn t => 
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   431
    let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   432
      val z = instantiate_cterm ([(zT,T)],[]) zr 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   433
      val eq = instantiate_cterm ([(eqT,T)],[]) geq
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   434
      val th = Simplifier.rewrite (ss addsimps simp_thms) 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   435
           (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   436
                  (Thm.capply (Thm.capply eq t) z)))
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   437
    in equal_elim (symmetric th) TrueI
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   438
    end
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   439
  end
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   440
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   441
 fun proc ctxt phi ss ct = 
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   442
  let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   443
    val ((x,y),(w,z)) = 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   444
         (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   445
    val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   446
    val T = ctyp_of_term x
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   447
    val [y_nz, z_nz] = map (prove_nz ctxt T) [y, z]
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   448
    val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   449
  in SOME (implies_elim (implies_elim th y_nz) z_nz)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   450
  end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   451
  handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   452
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   453
 fun proc2 ctxt phi ss ct = 
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   454
  let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   455
    val (l,r) = Thm.dest_binop ct
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   456
    val T = ctyp_of_term l
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   457
  in (case (term_of l, term_of r) of
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   458
      (Const(@{const_name "HOL.divide"},_)$_$_, _) => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   459
        let val (x,y) = Thm.dest_binop l val z = r
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   460
            val _ = map (HOLogic.dest_number o term_of) [x,y,z]
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   461
            val ynz = prove_nz ctxt T y
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   462
        in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   463
        end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   464
     | (_, Const (@{const_name "HOL.divide"},_)$_$_) => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   465
        let val (x,y) = Thm.dest_binop r val z = l
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   466
            val _ = map (HOLogic.dest_number o term_of) [x,y,z]
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   467
            val ynz = prove_nz ctxt T y
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   468
        in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   469
        end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   470
     | _ => NONE)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   471
  end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   472
  handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   473
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   474
 fun is_number (Const(@{const_name "HOL.divide"},_)$a$b) = is_number a andalso is_number b
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   475
   | is_number t = can HOLogic.dest_number t
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   476
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   477
 val is_number = is_number o term_of
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   478
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   479
 fun proc3 phi ss ct =
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   480
  (case term_of ct of
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   481
    Const(@{const_name "Orderings.less"},_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   482
      let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   483
        val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   484
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   485
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   486
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   487
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   488
  | Const(@{const_name "Orderings.less_eq"},_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   489
      let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   490
        val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   491
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   492
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   493
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   494
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   495
  | Const("op =",_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   496
      let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   497
        val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   498
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   499
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   500
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   501
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   502
  | Const(@{const_name "Orderings.less"},_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   503
    let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   504
      val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   505
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   506
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   507
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   508
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   509
  | Const(@{const_name "Orderings.less_eq"},_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   510
    let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   511
      val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   512
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   513
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   514
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   515
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   516
  | Const("op =",_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   517
    let 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   518
      val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   519
        val _ = map is_number [a,b,c]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   520
        val T = ctyp_of_term c
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   521
        val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   522
      in SOME (mk_meta_eq th) end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   523
  | _ => NONE)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   524
  handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   525
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   526
fun add_frac_frac_simproc ctxt = 
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   527
       make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}], 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   528
                     name = "add_frac_frac_simproc",
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   529
                     proc = proc ctxt, identifier = []}
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   530
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   531
fun add_frac_num_simproc ctxt = 
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   532
       make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}], 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   533
                     name = "add_frac_num_simproc",
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   534
                     proc = proc2 ctxt, identifier = []}
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   535
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   536
val ord_frac_simproc = 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   537
  make_simproc 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   538
    {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   539
             @{cpat "(?a::(?'a::{field, ord}))/?b \<le> ?c"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   540
             @{cpat "?c < (?a::(?'a::{field, ord}))/?b"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   541
             @{cpat "?c \<le> (?a::(?'a::{field, ord}))/?b"},
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   542
             @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"},
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   543
             @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}],
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   544
             name = "ord_frac_simproc", proc = proc3, identifier = []}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   545
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   546
val nat_arith = map thm ["add_nat_number_of", "diff_nat_number_of", 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   547
               "mult_nat_number_of", "eq_nat_number_of", "less_nat_number_of"]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   548
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   549
val comp_arith = (map thm ["Let_def", "if_False", "if_True", "add_0", 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   550
                 "add_Suc", "add_number_of_left", "mult_number_of_left", 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   551
                 "Suc_eq_add_numeral_1"])@
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   552
                 (map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"])
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   553
                 @ arith_simps@ nat_arith @ rel_simps 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   554
val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   555
           @{thm "divide_Numeral1"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   556
           @{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"},
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   557
           @{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"},
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   558
           @{thm "mult_num_frac"}, @{thm "mult_frac_num"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   559
           @{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   560
           @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   561
           @{thm "diff_def"}, @{thm "minus_divide_left"}, 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   562
           @{thm "Numeral1_eq1_nat"}, @{thm "add_divide_distrib"} RS sym]
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   563
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   564
23329
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   565
fun comp_conv ctxt = Simplifier.rewrite  
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   566
(HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"} 
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   567
              addsimps ths addsimps comp_arith addsimps simp_thms
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   568
              addsimprocs field_cancel_numeral_factors
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   569
               addsimprocs [add_frac_frac_simproc ctxt, add_frac_num_simproc ctxt,
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   570
                            ord_frac_simproc]
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   571
                addcongs [@{thm "if_weak_cong"}])
0dbb30302259 tuned setup for the fields instantiation for Groebner Bases;
chaieb
parents: 23327
diff changeset
   572
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   573
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   574
fun numeral_is_const ct = 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   575
  case term_of ct of 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   576
   Const (@{const_name "HOL.divide"},_) $ a $ b => 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   577
     can HOLogic.dest_number a andalso can HOLogic.dest_number b
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   578
 | t => can HOLogic.dest_number t
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   579
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   580
fun dest_const ct = case term_of ct of
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   581
   Const (@{const_name "HOL.divide"},_) $ a $ b=>
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   582
    Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b))
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   583
 | t => Rat.rat_of_int (snd (HOLogic.dest_number t))
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   584
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   585
fun mk_const phi cT x = 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   586
 let val (a, b) = Rat.quotient_of_rat x
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   587
 in if b = 1 then Normalizer.mk_cnumber cT a
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   588
    else Thm.capply 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   589
         (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"}) 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   590
                     (Normalizer.mk_cnumber cT a))
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   591
         (Normalizer.mk_cnumber cT b)
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   592
  end
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   593
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   594
in 
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   595
 NormalizerData.funs @{thm class_fieldgb.axioms}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   596
   {is_const = K numeral_is_const,
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   597
    dest_const = K dest_const,
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   598
    mk_const = mk_const,
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   599
    conv = K comp_conv}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   600
end
23327
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   601
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   602
*}
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   603
1654013ec97c Added instantiation of algebra method to fields
chaieb
parents: 23263
diff changeset
   604
end