| author | wenzelm | 
| Wed, 25 Mar 2015 13:31:47 +0100 | |
| changeset 59810 | e749a0f2f401 | 
| parent 58925 | 1b655309617c | 
| child 60758 | d8d85a8172b5 | 
| permissions | -rw-r--r-- | 
| 23252 | 1  | 
(* Title: HOL/Groebner_Basis.thy  | 
2  | 
Author: Amine Chaieb, TU Muenchen  | 
|
3  | 
*)  | 
|
4  | 
||
| 58889 | 5  | 
section {* Groebner bases *}
 | 
| 28402 | 6  | 
|
| 23252 | 7  | 
theory Groebner_Basis  | 
| 
58777
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
8  | 
imports Semiring_Normalization Parity  | 
| 23252 | 9  | 
begin  | 
10  | 
||
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
11  | 
subsection {* Groebner Bases *}
 | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
12  | 
|
| 54251 | 13  | 
lemmas bool_simps = simp_thms(1-34) -- {* FIXME move to @{theory HOL} *}
 | 
14  | 
||
15  | 
lemma nnf_simps: -- {* FIXME shadows fact binding in @{theory HOL} *}
 | 
|
16  | 
"(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)"  | 
|
17  | 
"(P \<longrightarrow> Q) = (\<not>P \<or> Q)"  | 
|
18  | 
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"  | 
|
19  | 
by blast+  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
20  | 
|
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
21  | 
lemma dnf:  | 
| 54251 | 22  | 
"(P & (Q | R)) = ((P&Q) | (P&R))"  | 
23  | 
"((Q | R) & P) = ((Q&P) | (R&P))"  | 
|
24  | 
"(P \<and> Q) = (Q \<and> P)"  | 
|
25  | 
"(P \<or> Q) = (Q \<or> P)"  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
26  | 
by blast+  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
27  | 
|
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
28  | 
lemmas weak_dnf_simps = dnf bool_simps  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
29  | 
|
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
30  | 
lemma PFalse:  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
31  | 
"P \<equiv> False \<Longrightarrow> \<not> P"  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
32  | 
"\<not> P \<Longrightarrow> (P \<equiv> False)"  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
33  | 
by auto  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
34  | 
|
| 57951 | 35  | 
named_theorems algebra "pre-simplification rules for algebraic methods"  | 
| 48891 | 36  | 
ML_file "Tools/groebner.ML"  | 
| 
36751
 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 
haftmann 
parents: 
36720 
diff
changeset
 | 
37  | 
|
| 47432 | 38  | 
method_setup algebra = {*
 | 
39  | 
let  | 
|
40  | 
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()  | 
|
41  | 
val addN = "add"  | 
|
42  | 
val delN = "del"  | 
|
43  | 
val any_keyword = keyword addN || keyword delN  | 
|
44  | 
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;  | 
|
45  | 
in  | 
|
46  | 
Scan.optional (keyword addN |-- thms) [] --  | 
|
47  | 
Scan.optional (keyword delN |-- thms) [] >>  | 
|
48  | 
(fn (add_ths, del_ths) => fn ctxt =>  | 
|
49  | 
SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))  | 
|
50  | 
end  | 
|
51  | 
*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"  | 
|
| 
36751
 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 
haftmann 
parents: 
36720 
diff
changeset
 | 
52  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
53  | 
declare dvd_def[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
54  | 
declare dvd_eq_mod_eq_0[symmetric, algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
55  | 
declare mod_div_trivial[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
56  | 
declare mod_mod_trivial[algebra]  | 
| 47142 | 57  | 
declare div_by_0[algebra]  | 
58  | 
declare mod_by_0[algebra]  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
59  | 
declare zmod_zdiv_equality[symmetric,algebra]  | 
| 47165 | 60  | 
declare div_mod_equality2[symmetric, algebra]  | 
| 47159 | 61  | 
declare div_minus_minus[algebra]  | 
62  | 
declare mod_minus_minus[algebra]  | 
|
63  | 
declare div_minus_right[algebra]  | 
|
64  | 
declare mod_minus_right[algebra]  | 
|
| 47142 | 65  | 
declare div_0[algebra]  | 
66  | 
declare mod_0[algebra]  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
67  | 
declare mod_by_1[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
68  | 
declare div_by_1[algebra]  | 
| 47160 | 69  | 
declare mod_minus1_right[algebra]  | 
70  | 
declare div_minus1_right[algebra]  | 
|
| 
36712
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
71  | 
declare mod_mult_self2_is_0[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
72  | 
declare mod_mult_self1_is_0[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
73  | 
declare zmod_eq_0_iff[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
74  | 
declare dvd_0_left_iff[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
75  | 
declare zdvd1_eq[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
76  | 
declare zmod_eq_dvd_iff[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
77  | 
declare nat_mod_eq_iff[algebra]  | 
| 
 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 
haftmann 
parents: 
36702 
diff
changeset
 | 
78  | 
|
| 
58777
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
79  | 
context semiring_parity  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
80  | 
begin  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
81  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
82  | 
declare even_times_iff [algebra]  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
83  | 
declare even_power [algebra]  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
84  | 
|
| 28402 | 85  | 
end  | 
| 
58777
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
86  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
87  | 
context ring_parity  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
88  | 
begin  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
89  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
90  | 
declare even_minus [algebra]  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
91  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
92  | 
end  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
93  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
94  | 
declare even_Suc [algebra]  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
95  | 
declare even_diff_nat [algebra]  | 
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
96  | 
|
| 
 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 
haftmann 
parents: 
57951 
diff
changeset
 | 
97  | 
end  |