| author | wenzelm | 
| Fri, 07 Mar 2014 19:28:34 +0100 | |
| changeset 55982 | b719781c7396 | 
| parent 55509 | bd67ebe275e0 | 
| child 56850 | 13a7bca533a3 | 
| 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: 
36720diff
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: 
36720diff
changeset | 8 | imports Semiring_Normalization | 
| 55509 
bd67ebe275e0
eliminated odd dislocation of keyword declaration and implementation (leftover from 318cd8ac1817)
 haftmann parents: 
55178diff
changeset | 9 | keywords "try0" :: diag | 
| 23252 | 10 | begin | 
| 11 | ||
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 12 | subsection {* Groebner Bases *}
 | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 13 | |
| 54251 | 14 | lemmas bool_simps = simp_thms(1-34) -- {* FIXME move to @{theory HOL} *}
 | 
| 15 | ||
| 16 | lemma nnf_simps: -- {* FIXME shadows fact binding in @{theory HOL} *}
 | |
| 17 | "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" | |
| 18 | "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" | |
| 19 | "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P" | |
| 20 | by blast+ | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 21 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 22 | lemma dnf: | 
| 54251 | 23 | "(P & (Q | R)) = ((P&Q) | (P&R))" | 
| 24 | "((Q | R) & P) = ((Q&P) | (R&P))" | |
| 25 | "(P \<and> Q) = (Q \<and> P)" | |
| 26 | "(P \<or> Q) = (Q \<or> P)" | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 27 | by blast+ | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 28 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 29 | lemmas weak_dnf_simps = dnf bool_simps | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 30 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 31 | lemma PFalse: | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 32 | "P \<equiv> False \<Longrightarrow> \<not> P" | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 33 | "\<not> P \<Longrightarrow> (P \<equiv> False)" | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 34 | by auto | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 35 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 36 | ML {*
 | 
| 45294 | 37 | structure Algebra_Simplification = Named_Thms | 
| 38 | ( | |
| 39 |   val name = @{binding algebra}
 | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 40 | 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: 
36702diff
changeset | 41 | ) | 
| 28402 | 42 | *} | 
| 43 | ||
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 44 | setup Algebra_Simplification.setup | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 45 | |
| 48891 | 46 | ML_file "Tools/groebner.ML" | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36720diff
changeset | 47 | |
| 47432 | 48 | method_setup algebra = {*
 | 
| 49 | let | |
| 50 | fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () | |
| 51 | val addN = "add" | |
| 52 | val delN = "del" | |
| 53 | val any_keyword = keyword addN || keyword delN | |
| 54 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; | |
| 55 | in | |
| 56 | Scan.optional (keyword addN |-- thms) [] -- | |
| 57 | Scan.optional (keyword delN |-- thms) [] >> | |
| 58 | (fn (add_ths, del_ths) => fn ctxt => | |
| 59 | SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt)) | |
| 60 | end | |
| 61 | *} "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: 
36720diff
changeset | 62 | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 63 | declare dvd_def[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 64 | 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: 
36702diff
changeset | 65 | declare mod_div_trivial[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 66 | declare mod_mod_trivial[algebra] | 
| 47142 | 67 | declare div_by_0[algebra] | 
| 68 | declare mod_by_0[algebra] | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 69 | declare zmod_zdiv_equality[symmetric,algebra] | 
| 47165 | 70 | declare div_mod_equality2[symmetric, algebra] | 
| 47159 | 71 | declare div_minus_minus[algebra] | 
| 72 | declare mod_minus_minus[algebra] | |
| 73 | declare div_minus_right[algebra] | |
| 74 | declare mod_minus_right[algebra] | |
| 47142 | 75 | declare div_0[algebra] | 
| 76 | declare mod_0[algebra] | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 77 | declare mod_by_1[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 78 | declare div_by_1[algebra] | 
| 47160 | 79 | declare mod_minus1_right[algebra] | 
| 80 | declare div_minus1_right[algebra] | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 81 | declare mod_mult_self2_is_0[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 82 | declare mod_mult_self1_is_0[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 83 | declare zmod_eq_0_iff[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 84 | declare dvd_0_left_iff[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 85 | declare zdvd1_eq[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 86 | declare zmod_eq_dvd_iff[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 87 | declare nat_mod_eq_iff[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 88 | |
| 55178 | 89 | |
| 90 | subsection {* Try0 *}
 | |
| 91 | ||
| 92 | ML_file "Tools/try0.ML" | |
| 93 | ||
| 28402 | 94 | end |