| author | wenzelm | 
| Sat, 18 Jan 2025 11:03:18 +0100 | |
| changeset 81904 | aa28d82d6b66 | 
| parent 76224 | 64e8d4afcf10 | 
| permissions | -rw-r--r-- | 
| 23252 | 1 | (* Title: HOL/Groebner_Basis.thy | 
| 2 | Author: Amine Chaieb, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 60758 | 5 | section \<open>Groebner bases\<close> | 
| 28402 | 6 | |
| 23252 | 7 | theory Groebner_Basis | 
| 58777 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 8 | imports Semiring_Normalization Parity | 
| 23252 | 9 | begin | 
| 10 | ||
| 60758 | 11 | subsection \<open>Groebner Bases\<close> | 
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 12 | |
| 69593 | 13 | lemmas bool_simps = simp_thms(1-34) \<comment> \<open>FIXME move to \<^theory>\<open>HOL.HOL\<close>\<close> | 
| 54251 | 14 | |
| 69593 | 15 | lemma nnf_simps: \<comment> \<open>FIXME shadows fact binding in \<^theory>\<open>HOL.HOL\<close>\<close> | 
| 54251 | 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: 
36702diff
changeset | 20 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 21 | lemma dnf: | 
| 67091 | 22 | "(P \<and> (Q \<or> R)) = ((P\<and>Q) \<or> (P\<and>R))" | 
| 23 | "((Q \<or> R) \<and> P) = ((Q\<and>P) \<or> (R\<and>P))" | |
| 54251 | 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: 
36702diff
changeset | 26 | by blast+ | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 27 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
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: 
36702diff
changeset | 29 | |
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 30 | lemma PFalse: | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
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: 
36702diff
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: 
36702diff
changeset | 33 | by auto | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 34 | |
| 57951 | 35 | named_theorems algebra "pre-simplification rules for algebraic methods" | 
| 69605 | 36 | ML_file \<open>Tools/groebner.ML\<close> | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36720diff
changeset | 37 | |
| 60758 | 38 | method_setup algebra = \<open> | 
| 47432 | 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 | |
| 61476 | 44 | val thms = Scan.repeats (Scan.unless any_keyword Attrib.multi_thm); | 
| 47432 | 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 | |
| 60758 | 51 | \<close> "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 | 52 | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 53 | declare dvd_def[algebra] | 
| 64248 | 54 | declare mod_eq_0_iff_dvd[algebra] | 
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 55 | declare mod_div_trivial[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 56 | declare mod_mod_trivial[algebra] | 
| 47142 | 57 | declare div_by_0[algebra] | 
| 58 | declare mod_by_0[algebra] | |
| 64246 | 59 | declare mult_div_mod_eq[algebra] | 
| 47159 | 60 | declare div_minus_minus[algebra] | 
| 61 | declare mod_minus_minus[algebra] | |
| 62 | declare div_minus_right[algebra] | |
| 63 | declare mod_minus_right[algebra] | |
| 47142 | 64 | declare div_0[algebra] | 
| 65 | declare mod_0[algebra] | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 66 | declare mod_by_1[algebra] | 
| 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 67 | declare div_by_1[algebra] | 
| 47160 | 68 | declare mod_minus1_right[algebra] | 
| 69 | 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 | 70 | 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 | 71 | declare mod_mult_self1_is_0[algebra] | 
| 76224 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
70341diff
changeset | 72 | |
| 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
70341diff
changeset | 73 | lemma zmod_eq_0_iff [algebra]: | 
| 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
70341diff
changeset | 74 | \<open>m mod d = 0 \<longleftrightarrow> (\<exists>q. m = d * q)\<close> for m d :: int | 
| 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
70341diff
changeset | 75 | by (auto simp add: mod_eq_0_iff_dvd) | 
| 
64e8d4afcf10
moved relevant theorems from theory Divides to theory Euclidean_Division
 haftmann parents: 
70341diff
changeset | 76 | |
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 77 | 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 | 78 | declare zdvd1_eq[algebra] | 
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64249diff
changeset | 79 | declare mod_eq_dvd_iff[algebra] | 
| 36712 
2f4c318861b3
avoid references to groebner bases in names which have no references to groebner bases
 haftmann parents: 
36702diff
changeset | 80 | 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 | 81 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 82 | context semiring_parity | 
| 58777 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 83 | begin | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 84 | |
| 68100 | 85 | declare even_mult_iff [algebra] | 
| 58777 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 86 | declare even_power [algebra] | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 87 | |
| 28402 | 88 | end | 
| 58777 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 89 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 90 | context ring_parity | 
| 58777 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 91 | begin | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 92 | |
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 93 | declare even_minus [algebra] | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 94 | |
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 95 | end | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 96 | |
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 97 | declare even_Suc [algebra] | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 98 | declare even_diff_nat [algebra] | 
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 99 | |
| 
6ba2f1fa243b
further downshift of theory Parity in the hierarchy
 haftmann parents: 
57951diff
changeset | 100 | end |