author | blanchet |
Thu, 03 Jan 2013 17:10:12 +0100 | |
changeset 50704 | cd1fcda1ea88 |
parent 48891 | c0eafbd55de3 |
child 54251 | adea9f6986b2 |
permissions | -rw-r--r-- |
23252 | 1 |
(* Title: HOL/Groebner_Basis.thy |
2 |
Author: Amine Chaieb, TU Muenchen |
|
3 |
*) |
|
4 |
||
36751
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents:
36720
diff
changeset
|
5 |
header {* Groebner bases *} |
28402 | 6 |
|
23252 | 7 |
theory Groebner_Basis |
36751
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
haftmann
parents:
36720
diff
changeset
|
8 |
imports Semiring_Normalization |
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 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
13 |
lemmas bool_simps = simp_thms(1-34) |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
14 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
15 |
lemma dnf: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
16 |
"(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
17 |
"(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
18 |
by blast+ |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
19 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
20 |
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
|
21 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
22 |
lemma nnf_simps: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
23 |
"(\<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)" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
24 |
"(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
25 |
by blast+ |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
26 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
27 |
lemma PFalse: |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
28 |
"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
|
29 |
"\<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
|
30 |
by auto |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
31 |
|
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
32 |
ML {* |
45294 | 33 |
structure Algebra_Simplification = Named_Thms |
34 |
( |
|
35 |
val name = @{binding algebra} |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
36 |
val description = "pre-simplification rules for algebraic methods" |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
37 |
) |
28402 | 38 |
*} |
39 |
||
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
40 |
setup Algebra_Simplification.setup |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
41 |
|
48891 | 42 |
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
|
43 |
|
47432 | 44 |
method_setup algebra = {* |
45 |
let |
|
46 |
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () |
|
47 |
val addN = "add" |
|
48 |
val delN = "del" |
|
49 |
val any_keyword = keyword addN || keyword delN |
|
50 |
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; |
|
51 |
in |
|
52 |
Scan.optional (keyword addN |-- thms) [] -- |
|
53 |
Scan.optional (keyword delN |-- thms) [] >> |
|
54 |
(fn (add_ths, del_ths) => fn ctxt => |
|
55 |
SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt)) |
|
56 |
end |
|
57 |
*} "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
|
58 |
|
36712
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
59 |
declare dvd_def[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
60 |
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
|
61 |
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
|
62 |
declare mod_mod_trivial[algebra] |
47142 | 63 |
declare div_by_0[algebra] |
64 |
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
|
65 |
declare zmod_zdiv_equality[symmetric,algebra] |
47165 | 66 |
declare div_mod_equality2[symmetric, algebra] |
47159 | 67 |
declare div_minus_minus[algebra] |
68 |
declare mod_minus_minus[algebra] |
|
69 |
declare div_minus_right[algebra] |
|
70 |
declare mod_minus_right[algebra] |
|
47142 | 71 |
declare div_0[algebra] |
72 |
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
|
73 |
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
|
74 |
declare div_by_1[algebra] |
47160 | 75 |
declare mod_minus1_right[algebra] |
76 |
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
|
77 |
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
|
78 |
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
|
79 |
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
|
80 |
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
|
81 |
declare zdvd1_eq[algebra] |
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
haftmann
parents:
36702
diff
changeset
|
82 |
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
|
83 |
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
|
84 |
|
28402 | 85 |
end |