| author | wenzelm | 
| Wed, 05 Mar 2008 21:24:03 +0100 | |
| changeset 26199 | 04817a8802f2 | 
| parent 26086 | 3c243098b64a | 
| child 26314 | 9c39fc898fff | 
| permissions | -rw-r--r-- | 
| 23252 | 1  | 
(* Title: HOL/Groebner_Basis.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Amine Chaieb, TU Muenchen  | 
|
4  | 
*)  | 
|
5  | 
||
6  | 
header {* Semiring normalization and Groebner Bases *}
 | 
|
7  | 
theory Groebner_Basis  | 
|
8  | 
imports NatBin  | 
|
9  | 
uses  | 
|
10  | 
"Tools/Groebner_Basis/misc.ML"  | 
|
11  | 
"Tools/Groebner_Basis/normalizer_data.ML"  | 
|
12  | 
  ("Tools/Groebner_Basis/normalizer.ML")
 | 
|
| 23312 | 13  | 
  ("Tools/Groebner_Basis/groebner.ML")
 | 
| 23252 | 14  | 
begin  | 
15  | 
||
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
16  | 
|
| 23252 | 17  | 
subsection {* Semiring normalization *}
 | 
18  | 
||
19  | 
setup NormalizerData.setup  | 
|
20  | 
||
21  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
22  | 
locale gb_semiring =  | 
| 23252 | 23  | 
fixes add mul pwr r0 r1  | 
24  | 
assumes add_a:"(add x (add y z) = add (add x y) z)"  | 
|
25  | 
and add_c: "add x y = add y x" and add_0:"add r0 x = x"  | 
|
26  | 
and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x"  | 
|
27  | 
and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0"  | 
|
28  | 
and mul_d:"mul x (add y z) = add (mul x y) (mul x z)"  | 
|
29  | 
and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)"  | 
|
30  | 
begin  | 
|
31  | 
||
32  | 
lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)"  | 
|
33  | 
proof (induct p)  | 
|
34  | 
case 0  | 
|
35  | 
then show ?case by (auto simp add: pwr_0 mul_1)  | 
|
36  | 
next  | 
|
37  | 
case Suc  | 
|
38  | 
from this [symmetric] show ?case  | 
|
39  | 
by (auto simp add: pwr_Suc mul_1 mul_a)  | 
|
40  | 
qed  | 
|
41  | 
||
42  | 
lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
43  | 
proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1)  | 
|
44  | 
fix q x y  | 
|
45  | 
assume "\<And>x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
46  | 
have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))"  | 
|
47  | 
by (simp add: mul_a)  | 
|
48  | 
also have "\<dots> = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c)  | 
|
49  | 
also have "\<dots> = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a)  | 
|
50  | 
finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) =  | 
|
51  | 
mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c)  | 
|
52  | 
qed  | 
|
53  | 
||
54  | 
lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)"  | 
|
55  | 
proof (induct p arbitrary: q)  | 
|
56  | 
case 0  | 
|
57  | 
show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto  | 
|
58  | 
next  | 
|
59  | 
case Suc  | 
|
60  | 
thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc)  | 
|
61  | 
qed  | 
|
62  | 
||
63  | 
||
64  | 
subsubsection {* Declaring the abstract theory *}
 | 
|
65  | 
||
66  | 
lemma semiring_ops:  | 
|
67  | 
includes meta_term_syntax  | 
|
68  | 
shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)"  | 
|
69  | 
and "TERM r0" and "TERM r1"  | 
|
70  | 
by rule+  | 
|
71  | 
||
72  | 
lemma semiring_rules:  | 
|
73  | 
"add (mul a m) (mul b m) = mul (add a b) m"  | 
|
74  | 
"add (mul a m) m = mul (add a r1) m"  | 
|
75  | 
"add m (mul a m) = mul (add a r1) m"  | 
|
76  | 
"add m m = mul (add r1 r1) m"  | 
|
77  | 
"add r0 a = a"  | 
|
78  | 
"add a r0 = a"  | 
|
79  | 
"mul a b = mul b a"  | 
|
80  | 
"mul (add a b) c = add (mul a c) (mul b c)"  | 
|
81  | 
"mul r0 a = r0"  | 
|
82  | 
"mul a r0 = r0"  | 
|
83  | 
"mul r1 a = a"  | 
|
84  | 
"mul a r1 = a"  | 
|
85  | 
"mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)"  | 
|
86  | 
"mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))"  | 
|
87  | 
"mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)"  | 
|
88  | 
"mul (mul lx ly) rx = mul (mul lx rx) ly"  | 
|
89  | 
"mul (mul lx ly) rx = mul lx (mul ly rx)"  | 
|
90  | 
"mul lx (mul rx ry) = mul (mul lx rx) ry"  | 
|
91  | 
"mul lx (mul rx ry) = mul rx (mul lx ry)"  | 
|
92  | 
"add (add a b) (add c d) = add (add a c) (add b d)"  | 
|
93  | 
"add (add a b) c = add a (add b c)"  | 
|
94  | 
"add a (add c d) = add c (add a d)"  | 
|
95  | 
"add (add a b) c = add (add a c) b"  | 
|
96  | 
"add a c = add c a"  | 
|
97  | 
"add a (add c d) = add (add a c) d"  | 
|
98  | 
"mul (pwr x p) (pwr x q) = pwr x (p + q)"  | 
|
99  | 
"mul x (pwr x q) = pwr x (Suc q)"  | 
|
100  | 
"mul (pwr x q) x = pwr x (Suc q)"  | 
|
101  | 
"mul x x = pwr x 2"  | 
|
102  | 
"pwr (mul x y) q = mul (pwr x q) (pwr y q)"  | 
|
103  | 
"pwr (pwr x p) q = pwr x (p * q)"  | 
|
104  | 
"pwr x 0 = r1"  | 
|
105  | 
"pwr x 1 = x"  | 
|
106  | 
"mul x (add y z) = add (mul x y) (mul x z)"  | 
|
107  | 
"pwr x (Suc q) = mul x (pwr x q)"  | 
|
108  | 
"pwr x (2*n) = mul (pwr x n) (pwr x n)"  | 
|
109  | 
"pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))"  | 
|
110  | 
proof -  | 
|
111  | 
show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp  | 
|
112  | 
next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp  | 
|
113  | 
next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp  | 
|
114  | 
next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp  | 
|
115  | 
next show "add r0 a = a" using add_0 by simp  | 
|
116  | 
next show "add a r0 = a" using add_0 add_c by simp  | 
|
117  | 
next show "mul a b = mul b a" using mul_c by simp  | 
|
118  | 
next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp  | 
|
119  | 
next show "mul r0 a = r0" using mul_0 by simp  | 
|
120  | 
next show "mul a r0 = r0" using mul_0 mul_c by simp  | 
|
121  | 
next show "mul r1 a = a" using mul_1 by simp  | 
|
122  | 
next show "mul a r1 = a" using mul_1 mul_c by simp  | 
|
123  | 
next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)"  | 
|
124  | 
using mul_c mul_a by simp  | 
|
125  | 
next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))"  | 
|
126  | 
using mul_a by simp  | 
|
127  | 
next  | 
|
128  | 
have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c)  | 
|
129  | 
also have "\<dots> = mul rx (mul ry (mul lx ly))" using mul_a by simp  | 
|
130  | 
finally  | 
|
131  | 
show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)"  | 
|
132  | 
using mul_c by simp  | 
|
133  | 
next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp  | 
|
134  | 
next  | 
|
135  | 
show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a)  | 
|
136  | 
next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a )  | 
|
137  | 
next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c)  | 
|
138  | 
next show "add (add a b) (add c d) = add (add a c) (add b d)"  | 
|
139  | 
using add_c add_a by simp  | 
|
140  | 
next show "add (add a b) c = add a (add b c)" using add_a by simp  | 
|
141  | 
next show "add a (add c d) = add c (add a d)"  | 
|
142  | 
apply (simp add: add_a) by (simp only: add_c)  | 
|
143  | 
next show "add (add a b) c = add (add a c) b" using add_a add_c by simp  | 
|
144  | 
next show "add a c = add c a" by (rule add_c)  | 
|
145  | 
next show "add a (add c d) = add (add a c) d" using add_a by simp  | 
|
146  | 
next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr)  | 
|
147  | 
next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp  | 
|
148  | 
next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp  | 
|
149  | 
next show "mul x x = pwr x 2" by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c)  | 
|
150  | 
next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul)  | 
|
151  | 
next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr)  | 
|
152  | 
next show "pwr x 0 = r1" using pwr_0 .  | 
|
153  | 
next show "pwr x 1 = x" by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c)  | 
|
154  | 
next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp  | 
|
155  | 
next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp  | 
|
156  | 
next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number mul_pwr)  | 
|
157  | 
next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))"  | 
|
158  | 
by (simp add: nat_number pwr_Suc mul_pwr)  | 
|
159  | 
qed  | 
|
160  | 
||
161  | 
||
162  | 
lemma "axioms" [normalizer  | 
|
163  | 
semiring ops: semiring_ops  | 
|
164  | 
semiring rules: semiring_rules]:  | 
|
| 26199 | 165  | 
"gb_semiring add mul pwr r0 r1" by (rule gb_semiring_axioms)  | 
| 23252 | 166  | 
|
167  | 
end  | 
|
168  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
169  | 
interpretation class_semiring: gb_semiring  | 
| 23252 | 170  | 
    ["op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1"]
 | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
171  | 
by unfold_locales (auto simp add: ring_simps power_Suc)  | 
| 23252 | 172  | 
|
173  | 
lemmas nat_arith =  | 
|
174  | 
add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of  | 
|
175  | 
||
176  | 
lemma not_iszero_Numeral1: "\<not> iszero (Numeral1::'a::number_ring)"  | 
|
177  | 
by (simp add: numeral_1_eq_1)  | 
|
178  | 
lemmas comp_arith = Let_def arith_simps nat_arith rel_simps if_False  | 
|
179  | 
if_True add_0 add_Suc add_number_of_left mult_number_of_left  | 
|
180  | 
numeral_1_eq_1[symmetric] Suc_eq_add_numeral_1  | 
|
181  | 
numeral_0_eq_0[symmetric] numerals[symmetric] not_iszero_1  | 
|
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
25250 
diff
changeset
 | 
182  | 
iszero_number_of_Bit1 iszero_number_of_Bit0 nonzero_number_of_Min  | 
| 23252 | 183  | 
iszero_number_of_Pls iszero_0 not_iszero_Numeral1  | 
184  | 
||
185  | 
lemmas semiring_norm = comp_arith  | 
|
186  | 
||
187  | 
ML {*
 | 
|
| 23573 | 188  | 
local  | 
| 23252 | 189  | 
|
| 23573 | 190  | 
open Conv;  | 
191  | 
||
192  | 
fun numeral_is_const ct =  | 
|
193  | 
can HOLogic.dest_number (Thm.term_of ct);  | 
|
| 23252 | 194  | 
|
| 23573 | 195  | 
fun int_of_rat x =  | 
196  | 
(case Rat.quotient_of_rat x of (i, 1) => i  | 
|
197  | 
| _ => error "int_of_rat: bad int");  | 
|
| 23252 | 198  | 
|
| 23573 | 199  | 
val numeral_conv =  | 
200  | 
  Simplifier.rewrite (HOL_basic_ss addsimps @{thms semiring_norm}) then_conv
 | 
|
201  | 
Simplifier.rewrite (HOL_basic_ss addsimps  | 
|
202  | 
    (@{thms numeral_1_eq_1} @ @{thms numeral_0_eq_0} @ @{thms numerals(1-2)}));
 | 
|
203  | 
||
204  | 
in  | 
|
205  | 
||
206  | 
fun normalizer_funs key =  | 
|
207  | 
NormalizerData.funs key  | 
|
| 23252 | 208  | 
   {is_const = fn phi => numeral_is_const,
 | 
209  | 
dest_const = fn phi => fn ct =>  | 
|
210  | 
Rat.rat_of_int (snd  | 
|
211  | 
(HOLogic.dest_number (Thm.term_of ct)  | 
|
212  | 
handle TERM _ => error "ring_dest_const")),  | 
|
| 23573 | 213  | 
mk_const = fn phi => fn cT => fn x => Numeral.mk_cnumber cT (int_of_rat x),  | 
| 
23330
 
01c09922ce59
Conversion for computation on constants now depends on the context
 
chaieb 
parents: 
23327 
diff
changeset
 | 
214  | 
conv = fn phi => K numeral_conv}  | 
| 23573 | 215  | 
|
216  | 
end  | 
|
| 23252 | 217  | 
*}  | 
218  | 
||
| 23573 | 219  | 
declaration {* normalizer_funs @{thm class_semiring.axioms} *}
 | 
220  | 
||
| 23252 | 221  | 
|
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
222  | 
locale gb_ring = gb_semiring +  | 
| 23252 | 223  | 
fixes sub :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
224  | 
and neg :: "'a \<Rightarrow> 'a"  | 
|
225  | 
assumes neg_mul: "neg x = mul (neg r1) x"  | 
|
226  | 
and sub_add: "sub x y = add x (neg y)"  | 
|
227  | 
begin  | 
|
228  | 
||
229  | 
lemma ring_ops:  | 
|
230  | 
includes meta_term_syntax  | 
|
231  | 
shows "TERM (sub x y)" and "TERM (neg x)" .  | 
|
232  | 
||
233  | 
lemmas ring_rules = neg_mul sub_add  | 
|
234  | 
||
235  | 
lemma "axioms" [normalizer  | 
|
236  | 
semiring ops: semiring_ops  | 
|
237  | 
semiring rules: semiring_rules  | 
|
238  | 
ring ops: ring_ops  | 
|
239  | 
ring rules: ring_rules]:  | 
|
| 26199 | 240  | 
"gb_ring add mul pwr r0 r1 sub neg" by (rule gb_ring_axioms)  | 
| 23252 | 241  | 
|
242  | 
end  | 
|
243  | 
||
244  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
245  | 
interpretation class_ring: gb_ring ["op +" "op *" "op ^"  | 
| 23252 | 246  | 
    "0::'a::{comm_semiring_1,recpower,number_ring}" 1 "op -" "uminus"]
 | 
247  | 
by unfold_locales simp_all  | 
|
248  | 
||
249  | 
||
| 23573 | 250  | 
declaration {* normalizer_funs @{thm class_ring.axioms} *}
 | 
| 23252 | 251  | 
|
252  | 
use "Tools/Groebner_Basis/normalizer.ML"  | 
|
253  | 
||
254  | 
method_setup sring_norm = {*
 | 
|
255  | 
Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' (Normalizer.semiring_normalize_tac ctxt))  | 
|
| 23458 | 256  | 
*} "semiring normalizer"  | 
| 23252 | 257  | 
|
258  | 
||
| 23327 | 259  | 
locale gb_field = gb_ring +  | 
260  | 
fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
|
261  | 
and inverse:: "'a \<Rightarrow> 'a"  | 
|
262  | 
assumes divide: "divide x y = mul x (inverse y)"  | 
|
263  | 
and inverse: "inverse x = divide r1 x"  | 
|
264  | 
begin  | 
|
265  | 
||
266  | 
lemma "axioms" [normalizer  | 
|
267  | 
semiring ops: semiring_ops  | 
|
268  | 
semiring rules: semiring_rules  | 
|
269  | 
ring ops: ring_ops  | 
|
270  | 
ring rules: ring_rules]:  | 
|
| 26199 | 271  | 
"gb_field add mul pwr r0 r1 sub neg divide inverse" by (rule gb_field_axioms)  | 
| 23327 | 272  | 
|
273  | 
end  | 
|
274  | 
||
| 23458 | 275  | 
|
| 23266 | 276  | 
subsection {* Groebner Bases *}
 | 
| 23252 | 277  | 
|
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
278  | 
locale semiringb = gb_semiring +  | 
| 23252 | 279  | 
assumes add_cancel: "add (x::'a) y = add x z \<longleftrightarrow> y = z"  | 
280  | 
and add_mul_solve: "add (mul w y) (mul x z) =  | 
|
281  | 
add (mul w z) (mul x y) \<longleftrightarrow> w = x \<or> y = z"  | 
|
282  | 
begin  | 
|
283  | 
||
284  | 
lemma noteq_reduce: "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
285  | 
proof-  | 
|
286  | 
have "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> \<not> (a = b \<or> c = d)" by simp  | 
|
287  | 
also have "\<dots> \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
288  | 
using add_mul_solve by blast  | 
|
289  | 
finally show "a \<noteq> b \<and> c \<noteq> d \<longleftrightarrow> add (mul a c) (mul b d) \<noteq> add (mul a d) (mul b c)"  | 
|
290  | 
by simp  | 
|
291  | 
qed  | 
|
292  | 
||
293  | 
lemma add_scale_eq_noteq: "\<lbrakk>r \<noteq> r0 ; (a = b) \<and> ~(c = d)\<rbrakk>  | 
|
294  | 
\<Longrightarrow> add a (mul r c) \<noteq> add b (mul r d)"  | 
|
295  | 
proof(clarify)  | 
|
296  | 
assume nz: "r\<noteq> r0" and cnd: "c\<noteq>d"  | 
|
297  | 
and eq: "add b (mul r c) = add b (mul r d)"  | 
|
298  | 
hence "mul r c = mul r d" using cnd add_cancel by simp  | 
|
299  | 
hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)"  | 
|
300  | 
using mul_0 add_cancel by simp  | 
|
301  | 
thus "False" using add_mul_solve nz cnd by simp  | 
|
302  | 
qed  | 
|
303  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
304  | 
lemma add_r0_iff: " x = add x a \<longleftrightarrow> a = r0"  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
305  | 
proof-  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
306  | 
have "a = r0 \<longleftrightarrow> add x a = add x r0" by (simp add: add_cancel)  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
307  | 
thus "x = add x a \<longleftrightarrow> a = r0" by (auto simp add: add_c add_0)  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
308  | 
qed  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
309  | 
|
| 23252 | 310  | 
declare "axioms" [normalizer del]  | 
311  | 
||
312  | 
lemma "axioms" [normalizer  | 
|
313  | 
semiring ops: semiring_ops  | 
|
314  | 
semiring rules: semiring_rules  | 
|
315  | 
idom rules: noteq_reduce add_scale_eq_noteq]:  | 
|
| 26199 | 316  | 
"semiringb add mul pwr r0 r1" by (rule semiringb_axioms)  | 
| 23252 | 317  | 
|
318  | 
end  | 
|
319  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
320  | 
locale ringb = semiringb + gb_ring +  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
321  | 
assumes subr0_iff: "sub x y = r0 \<longleftrightarrow> x = y"  | 
| 23252 | 322  | 
begin  | 
323  | 
||
324  | 
declare "axioms" [normalizer del]  | 
|
325  | 
||
326  | 
lemma "axioms" [normalizer  | 
|
327  | 
semiring ops: semiring_ops  | 
|
328  | 
semiring rules: semiring_rules  | 
|
329  | 
ring ops: ring_ops  | 
|
330  | 
ring rules: ring_rules  | 
|
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
331  | 
idom rules: noteq_reduce add_scale_eq_noteq  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
332  | 
ideal rules: subr0_iff add_r0_iff]:  | 
| 26199 | 333  | 
"ringb add mul pwr r0 r1 sub neg" by (rule ringb_axioms)  | 
| 23252 | 334  | 
|
335  | 
end  | 
|
336  | 
||
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
337  | 
|
| 23252 | 338  | 
lemma no_zero_divirors_neq0:  | 
339  | 
assumes az: "(a::'a::no_zero_divisors) \<noteq> 0"  | 
|
340  | 
and ab: "a*b = 0" shows "b = 0"  | 
|
341  | 
proof -  | 
|
342  | 
  { assume bz: "b \<noteq> 0"
 | 
|
343  | 
from no_zero_divisors [OF az bz] ab have False by blast }  | 
|
344  | 
thus "b = 0" by blast  | 
|
345  | 
qed  | 
|
346  | 
||
347  | 
interpretation class_ringb: ringb  | 
|
348  | 
  ["op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus"]
 | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
349  | 
proof(unfold_locales, simp add: ring_simps power_Suc, auto)  | 
| 23252 | 350  | 
  fix w x y z ::"'a::{idom,recpower,number_ring}"
 | 
351  | 
assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"  | 
|
352  | 
hence ynz': "y - z \<noteq> 0" by simp  | 
|
353  | 
from p have "w * y + x* z - w*z - x*y = 0" by simp  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
354  | 
hence "w* (y - z) - x * (y - z) = 0" by (simp add: ring_simps)  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
355  | 
hence "(y - z) * (w - x) = 0" by (simp add: ring_simps)  | 
| 23252 | 356  | 
with no_zero_divirors_neq0 [OF ynz']  | 
357  | 
have "w - x = 0" by blast  | 
|
358  | 
thus "w = x" by simp  | 
|
359  | 
qed  | 
|
360  | 
||
| 23573 | 361  | 
declaration {* normalizer_funs @{thm class_ringb.axioms} *}
 | 
| 23252 | 362  | 
|
363  | 
interpretation natgb: semiringb  | 
|
364  | 
["op +" "op *" "op ^" "0::nat" "1"]  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
365  | 
proof (unfold_locales, simp add: ring_simps power_Suc)  | 
| 23252 | 366  | 
fix w x y z ::"nat"  | 
367  | 
  { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \<noteq> z"
 | 
|
368  | 
hence "y < z \<or> y > z" by arith  | 
|
369  | 
    moreover {
 | 
|
370  | 
assume lt:"y <z" hence "\<exists>k. z = y + k \<and> k > 0" by (rule_tac x="z - y" in exI, auto)  | 
|
371  | 
then obtain k where kp: "k>0" and yz:"z = y + k" by blast  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
372  | 
from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz ring_simps)  | 
| 23252 | 373  | 
hence "x*k = w*k" by simp  | 
374  | 
hence "w = x" using kp by (simp add: mult_cancel2) }  | 
|
375  | 
    moreover {
 | 
|
376  | 
assume lt: "y >z" hence "\<exists>k. y = z + k \<and> k>0" by (rule_tac x="y - z" in exI, auto)  | 
|
377  | 
then obtain k where kp: "k>0" and yz:"y = z + k" by blast  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23458 
diff
changeset
 | 
378  | 
from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz ring_simps)  | 
| 23252 | 379  | 
hence "w*k = x*k" by simp  | 
380  | 
hence "w = x" using kp by (simp add: mult_cancel2)}  | 
|
381  | 
ultimately have "w=x" by blast }  | 
|
382  | 
thus "(w * y + x * z = w * z + x * y) = (w = x \<or> y = z)" by auto  | 
|
383  | 
qed  | 
|
384  | 
||
| 23573 | 385  | 
declaration {* normalizer_funs @{thm natgb.axioms} *}
 | 
| 23252 | 386  | 
|
| 23327 | 387  | 
locale fieldgb = ringb + gb_field  | 
388  | 
begin  | 
|
389  | 
||
390  | 
declare "axioms" [normalizer del]  | 
|
391  | 
||
392  | 
lemma "axioms" [normalizer  | 
|
393  | 
semiring ops: semiring_ops  | 
|
394  | 
semiring rules: semiring_rules  | 
|
395  | 
ring ops: ring_ops  | 
|
396  | 
ring rules: ring_rules  | 
|
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
397  | 
idom rules: noteq_reduce add_scale_eq_noteq  | 
| 
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
398  | 
ideal rules: subr0_iff add_r0_iff]:  | 
| 23327 | 399  | 
"fieldgb add mul pwr r0 r1 sub neg divide inverse" by unfold_locales  | 
400  | 
end  | 
|
401  | 
||
402  | 
||
| 
23258
 
9062e98fdab1
renamed locale ring/semiring to gb_ring/gb_semiring to avoid clash with Ring_and_Field versions;
 
wenzelm 
parents: 
23252 
diff
changeset
 | 
403  | 
lemmas bool_simps = simp_thms(1-34)  | 
| 23252 | 404  | 
lemma dnf:  | 
405  | 
"(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"  | 
|
406  | 
"(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"  | 
|
407  | 
by blast+  | 
|
408  | 
||
409  | 
lemmas weak_dnf_simps = dnf bool_simps  | 
|
410  | 
||
411  | 
lemma nnf_simps:  | 
|
412  | 
"(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"  | 
|
413  | 
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"  | 
|
414  | 
by blast+  | 
|
415  | 
||
416  | 
lemma PFalse:  | 
|
417  | 
"P \<equiv> False \<Longrightarrow> \<not> P"  | 
|
418  | 
"\<not> P \<Longrightarrow> (P \<equiv> False)"  | 
|
419  | 
by auto  | 
|
420  | 
||
421  | 
use "Tools/Groebner_Basis/groebner.ML"  | 
|
422  | 
||
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
423  | 
method_setup algebra =  | 
| 23458 | 424  | 
{*
 | 
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
425  | 
let  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
426  | 
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
427  | 
val addN = "add"  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
428  | 
val delN = "del"  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
429  | 
val any_keyword = keyword addN || keyword delN  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
430  | 
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
431  | 
in  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
432  | 
fn src => Method.syntax  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
433  | 
((Scan.optional (keyword addN |-- thms) []) --  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
434  | 
(Scan.optional (keyword delN |-- thms) [])) src  | 
| 
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
435  | 
#> (fn ((add_ths, del_ths), ctxt) =>  | 
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
436  | 
Method.SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))  | 
| 
23332
 
b91295432e6d
algebra_tac moved to file Tools/Groebner_Basis/groebner.ML; Method now takes theorems to be added or deleted from a simpset for simplificatio *before* the core method starts
 
chaieb 
parents: 
23330 
diff
changeset
 | 
437  | 
end  | 
| 
25250
 
b3a485b98963
(1) added axiom to ringb and theorems to enable algebra to prove the ideal membership problem; (2) Method algebra now calls algebra_tac which first tries to solve a universal formula, then in case of failure trie to solve the ideal membership problem (see HOL/Tools/Groebner_Basis/groebner.ML)
 
chaieb 
parents: 
23573 
diff
changeset
 | 
438  | 
*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"  | 
| 23252 | 439  | 
|
440  | 
end  |