merged
authorhaftmann
Thu, 06 May 2010 19:27:51 +0200
changeset 36717 2a72455be88b
parent 36694 978e6469b504 (current diff)
parent 36716 b09f3ad3208f (diff)
child 36718 30cdc863a4f8
merged
src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML
src/HOL/Library/normarith.ML
src/HOL/Tools/Groebner_Basis/misc.ML
src/HOL/Tools/Groebner_Basis/normalizer_data.ML
src/HOL/Tools/Qelim/cooper.ML
src/HOL/Tools/Qelim/presburger.ML
--- a/NEWS	Thu May 06 10:55:09 2010 +0200
+++ b/NEWS	Thu May 06 19:27:51 2010 +0200
@@ -140,6 +140,8 @@
 
 *** HOL ***
 
+* Dropped theorem duplicate comp_arith; use semiring_norm instead.  INCOMPATIBILITY.
+
 * Theory 'Finite_Set': various folding_* locales facilitate the application
 of the various fold combinators on finite sets.
 
--- a/src/HOL/Groebner_Basis.thy	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Groebner_Basis.thy	Thu May 06 19:27:51 2010 +0200
@@ -5,20 +5,17 @@
 header {* Semiring normalization and Groebner Bases *}
 
 theory Groebner_Basis
-imports Numeral_Simprocs
+imports Numeral_Simprocs Nat_Transfer
 uses
-  "Tools/Groebner_Basis/misc.ML"
-  "Tools/Groebner_Basis/normalizer_data.ML"
-  ("Tools/Groebner_Basis/normalizer.ML")
+  "Tools/Groebner_Basis/normalizer.ML"
   ("Tools/Groebner_Basis/groebner.ML")
 begin
 
 subsection {* Semiring normalization *}
 
-setup NormalizerData.setup
+setup Normalizer.setup
 
-
-locale gb_semiring =
+locale normalizing_semiring =
   fixes add mul pwr r0 r1
   assumes add_a:"(add x (add y z) = add (add x y) z)"
     and add_c: "add x y = add y x" and add_0:"add r0 x = x"
@@ -59,9 +56,6 @@
   thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc)
 qed
 
-
-subsubsection {* Declaring the abstract theory *}
-
 lemma semiring_ops:
   shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)"
     and "TERM r0" and "TERM r1" .
@@ -156,41 +150,21 @@
 qed
 
 
-lemmas gb_semiring_axioms' =
-  gb_semiring_axioms [normalizer
+lemmas normalizing_semiring_axioms' =
+  normalizing_semiring_axioms [normalizer
     semiring ops: semiring_ops
     semiring rules: semiring_rules]
 
 end
 
-interpretation class_semiring: gb_semiring
-    "op +" "op *" "op ^" "0::'a::{comm_semiring_1}" "1"
-  proof qed (auto simp add: algebra_simps)
-
-lemmas nat_arith =
-  add_nat_number_of
-  diff_nat_number_of
-  mult_nat_number_of
-  eq_nat_number_of
-  less_nat_number_of
-
-lemma not_iszero_Numeral1: "\<not> iszero (Numeral1::'a::number_ring)"
-  by simp
-
-lemmas comp_arith =
-  Let_def arith_simps nat_arith rel_simps neg_simps if_False
-  if_True add_0 add_Suc add_number_of_left mult_number_of_left
-  numeral_1_eq_1[symmetric] Suc_eq_plus1
-  numeral_0_eq_0[symmetric] numerals[symmetric]
-  iszero_simps not_iszero_Numeral1
-
-lemmas semiring_norm = comp_arith
+sublocale comm_semiring_1
+  < normalizing!: normalizing_semiring plus times power zero one
+proof
+qed (simp_all add: algebra_simps)
 
 ML {*
 local
 
-open Conv;
-
 fun numeral_is_const ct = can HOLogic.dest_number (Thm.term_of ct);
 
 fun int_of_rat x =
@@ -204,8 +178,8 @@
 
 in
 
-fun normalizer_funs key =
-  NormalizerData.funs key
+fun normalizer_funs' key =
+  Normalizer.funs key
    {is_const = fn phi => numeral_is_const,
     dest_const = fn phi => fn ct =>
       Rat.rat_of_int (snd
@@ -217,10 +191,9 @@
 end
 *}
 
-declaration {* normalizer_funs @{thm class_semiring.gb_semiring_axioms'} *}
+declaration {* normalizer_funs' @{thm normalizing.normalizing_semiring_axioms'} *}
 
-
-locale gb_ring = gb_semiring +
+locale normalizing_ring = normalizing_semiring +
   fixes sub :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
     and neg :: "'a \<Rightarrow> 'a"
   assumes neg_mul: "neg x = mul (neg r1) x"
@@ -231,8 +204,8 @@
 
 lemmas ring_rules = neg_mul sub_add
 
-lemmas gb_ring_axioms' =
-  gb_ring_axioms [normalizer
+lemmas normalizing_ring_axioms' =
+  normalizing_ring_axioms [normalizer
     semiring ops: semiring_ops
     semiring rules: semiring_rules
     ring ops: ring_ops
@@ -240,23 +213,14 @@
 
 end
 
-
-interpretation class_ring: gb_ring "op +" "op *" "op ^"
-    "0::'a::{comm_semiring_1,number_ring}" 1 "op -" "uminus"
-  proof qed simp_all
-
-
-declaration {* normalizer_funs @{thm class_ring.gb_ring_axioms'} *}
+(*FIXME add class*)
+interpretation normalizing!: normalizing_ring plus times power
+  "0::'a::{comm_semiring_1,number_ring}" 1 minus uminus proof
+qed simp_all
 
-use "Tools/Groebner_Basis/normalizer.ML"
-
+declaration {* normalizer_funs' @{thm normalizing.normalizing_ring_axioms'} *}
 
-method_setup sring_norm = {*
-  Scan.succeed (SIMPLE_METHOD' o Normalizer.semiring_normalize_tac)
-*} "semiring normalizer"
-
-
-locale gb_field = gb_ring +
+locale normalizing_field = normalizing_ring +
   fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
     and inverse:: "'a \<Rightarrow> 'a"
   assumes divide_inverse: "divide x y = mul x (inverse y)"
@@ -267,8 +231,8 @@
 
 lemmas field_rules = divide_inverse inverse_divide
 
-lemmas gb_field_axioms' =
-  gb_field_axioms [normalizer
+lemmas normalizing_field_axioms' =
+  normalizing_field_axioms [normalizer
     semiring ops: semiring_ops
     semiring rules: semiring_rules
     ring ops: ring_ops
@@ -278,10 +242,7 @@
 
 end
 
-
-subsection {* Groebner Bases *}
-
-locale semiringb = gb_semiring +
+locale normalizing_semiring_cancel = normalizing_semiring +
   assumes add_cancel: "add (x::'a) y = add x z \<longleftrightarrow> y = z"
   and add_mul_solve: "add (mul w y) (mul x z) =
     add (mul w z) (mul x y) \<longleftrightarrow> w = x \<or> y = z"
@@ -313,22 +274,23 @@
   thus "x = add x a \<longleftrightarrow> a = r0" by (auto simp add: add_c add_0)
 qed
 
-declare gb_semiring_axioms' [normalizer del]
+declare normalizing_semiring_axioms' [normalizer del]
 
-lemmas semiringb_axioms' = semiringb_axioms [normalizer
-  semiring ops: semiring_ops
-  semiring rules: semiring_rules
-  idom rules: noteq_reduce add_scale_eq_noteq]
+lemmas normalizing_semiring_cancel_axioms' =
+  normalizing_semiring_cancel_axioms [normalizer
+    semiring ops: semiring_ops
+    semiring rules: semiring_rules
+    idom rules: noteq_reduce add_scale_eq_noteq]
 
 end
 
-locale ringb = semiringb + gb_ring + 
+locale normalizing_ring_cancel = normalizing_semiring_cancel + normalizing_ring + 
   assumes subr0_iff: "sub x y = r0 \<longleftrightarrow> x = y"
 begin
 
-declare gb_ring_axioms' [normalizer del]
+declare normalizing_ring_axioms' [normalizer del]
 
-lemmas ringb_axioms' = ringb_axioms [normalizer
+lemmas normalizing_ring_cancel_axioms' = normalizing_ring_cancel_axioms [normalizer
   semiring ops: semiring_ops
   semiring rules: semiring_rules
   ring ops: ring_ops
@@ -338,17 +300,16 @@
 
 end
 
-
-lemma no_zero_divirors_neq0:
-  assumes az: "(a::'a::no_zero_divisors) \<noteq> 0"
-    and ab: "a*b = 0" shows "b = 0"
-proof -
-  { assume bz: "b \<noteq> 0"
-    from no_zero_divisors [OF az bz] ab have False by blast }
-  thus "b = 0" by blast
+lemma (in no_zero_divisors) prod_eq_zero_eq_zero:
+  assumes "a * b = 0" and "a \<noteq> 0"
+  shows "b = 0"
+proof (rule classical)
+  assume "b \<noteq> 0" with `a \<noteq> 0` no_zero_divisors have "a * b \<noteq> 0" by blast
+  with `a * b = 0` show ?thesis by simp
 qed
 
-interpretation class_ringb: ringb
+(*FIXME introduce class*)
+interpretation normalizing!: normalizing_ring_cancel
   "op +" "op *" "op ^" "0::'a::{idom,number_ring}" "1" "op -" "uminus"
 proof(unfold_locales, simp add: algebra_simps, auto)
   fix w x y z ::"'a::{idom,number_ring}"
@@ -357,14 +318,14 @@
   from p have "w * y + x* z - w*z - x*y = 0" by simp
   hence "w* (y - z) - x * (y - z) = 0" by (simp add: algebra_simps)
   hence "(y - z) * (w - x) = 0" by (simp add: algebra_simps)
-  with  no_zero_divirors_neq0 [OF ynz']
+  with  prod_eq_zero_eq_zero [OF _ ynz']
   have "w - x = 0" by blast
   thus "w = x"  by simp
 qed
 
-declaration {* normalizer_funs @{thm class_ringb.ringb_axioms'} *}
+declaration {* normalizer_funs' @{thm normalizing.normalizing_ring_cancel_axioms'} *}
 
-interpretation natgb: semiringb
+interpretation normalizing_nat!: normalizing_semiring_cancel
   "op +" "op *" "op ^" "0::nat" "1"
 proof (unfold_locales, simp add: algebra_simps)
   fix w x y z ::"nat"
@@ -386,14 +347,14 @@
   thus "(w * y + x * z = w * z + x * y) = (w = x \<or> y = z)" by auto
 qed
 
-declaration {* normalizer_funs @{thm natgb.semiringb_axioms'} *}
+declaration {* normalizer_funs' @{thm normalizing_nat.normalizing_semiring_cancel_axioms'} *}
 
-locale fieldgb = ringb + gb_field
+locale normalizing_field_cancel = normalizing_ring_cancel + normalizing_field
 begin
 
-declare gb_field_axioms' [normalizer del]
+declare normalizing_field_axioms' [normalizer del]
 
-lemmas fieldgb_axioms' = fieldgb_axioms [normalizer
+lemmas normalizing_field_cancel_axioms' = normalizing_field_cancel_axioms [normalizer
   semiring ops: semiring_ops
   semiring rules: semiring_rules
   ring ops: ring_ops
@@ -405,73 +366,10 @@
 
 end
 
-
-lemmas bool_simps = simp_thms(1-34)
-lemma dnf:
-    "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"
-    "(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"
-  by blast+
-
-lemmas weak_dnf_simps = dnf bool_simps
-
-lemma nnf_simps:
-    "(\<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)"
-    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
-  by blast+
-
-lemma PFalse:
-    "P \<equiv> False \<Longrightarrow> \<not> P"
-    "\<not> P \<Longrightarrow> (P \<equiv> False)"
-  by auto
-use "Tools/Groebner_Basis/groebner.ML"
-
-method_setup algebra =
-{*
-let
- fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
- val addN = "add"
- val delN = "del"
- val any_keyword = keyword addN || keyword delN
- val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-in
-  ((Scan.optional (keyword addN |-- thms) []) -- 
-   (Scan.optional (keyword delN |-- thms) [])) >>
-  (fn (add_ths, del_ths) => fn ctxt =>
-       SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))
-end
-*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
-declare dvd_def[algebra]
-declare dvd_eq_mod_eq_0[symmetric, algebra]
-declare mod_div_trivial[algebra]
-declare mod_mod_trivial[algebra]
-declare conjunct1[OF DIVISION_BY_ZERO, algebra]
-declare conjunct2[OF DIVISION_BY_ZERO, algebra]
-declare zmod_zdiv_equality[symmetric,algebra]
-declare zdiv_zmod_equality[symmetric, algebra]
-declare zdiv_zminus_zminus[algebra]
-declare zmod_zminus_zminus[algebra]
-declare zdiv_zminus2[algebra]
-declare zmod_zminus2[algebra]
-declare zdiv_zero[algebra]
-declare zmod_zero[algebra]
-declare mod_by_1[algebra]
-declare div_by_1[algebra]
-declare zmod_minus1_right[algebra]
-declare zdiv_minus1_right[algebra]
-declare mod_div_trivial[algebra]
-declare mod_mod_trivial[algebra]
-declare mod_mult_self2_is_0[algebra]
-declare mod_mult_self1_is_0[algebra]
-declare zmod_eq_0_iff[algebra]
-declare dvd_0_left_iff[algebra]
-declare zdvd1_eq[algebra]
-declare zmod_eq_dvd_iff[algebra]
-declare nat_mod_eq_iff[algebra]
-
-subsection{* Groebner Bases for fields *}
-
-interpretation class_fieldgb:
-  fieldgb "op +" "op *" "op ^" "0::'a::{field,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse)
+(*FIXME introduce class*)
+interpretation normalizing!: normalizing_field_cancel "op +" "op *" "op ^"
+  "0::'a::{field,number_ring}" "1" "op -" "uminus" "op /" "inverse"
+apply (unfold_locales) by (simp_all add: divide_inverse)
 
 lemma divide_Numeral1: "(x::'a::{field, number_ring}) / Numeral1 = x" by simp
 lemma divide_Numeral0: "(x::'a::{field_inverse_zero, number_ring}) / Numeral0 = 0"
@@ -479,9 +377,9 @@
 lemma mult_frac_frac: "((x::'a::field_inverse_zero) / y) * (z / w) = (x*z) / (y*w)"
   by simp
 lemma mult_frac_num: "((x::'a::field_inverse_zero) / y) * z  = (x*z) / y"
-  by simp
+  by (fact times_divide_eq_left)
 lemma mult_num_frac: "((x::'a::field_inverse_zero) / y) * z  = (x*z) / y"
-  by simp
+  by (fact times_divide_eq_left)
 
 lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp
 
@@ -490,13 +388,7 @@
 lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::field_inverse_zero) / y = (x + z*y) / y"
   by (simp add: add_divide_distrib)
 
-ML {*
-let open Conv
-in fconv_rule (arg_conv (arg1_conv (rewr_conv (mk_meta_eq @{thm mult_commute})))) (@{thm field_divide_inverse} RS sym)
-end
-*}
-
-ML{* 
+ML {* 
 local
  val zr = @{cpat "0"}
  val zT = ctyp_of_term zr
@@ -621,10 +513,6 @@
              @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}],
              name = "ord_frac_simproc", proc = proc3, identifier = []}
 
-local
-open Conv
-in
-
 val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
            @{thm "divide_Numeral1"},
            @{thm "divide_zero"}, @{thm "divide_Numeral0"},
@@ -635,11 +523,13 @@
            @{thm "diff_def"}, @{thm "minus_divide_left"},
            @{thm "Numeral1_eq1_nat"}, @{thm "add_divide_distrib"} RS sym,
            @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, 
-           fconv_rule (arg_conv (arg1_conv (rewr_conv (mk_meta_eq @{thm mult_commute}))))   
+           Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute}))))   
            (@{thm field_divide_inverse} RS sym)]
 
-val comp_conv = (Simplifier.rewrite
-(HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"}
+in
+
+val field_comp_conv = (Simplifier.rewrite
+(HOL_basic_ss addsimps @{thms "semiring_norm"}
               addsimps ths addsimps @{thms simp_thms}
               addsimprocs Numeral_Simprocs.field_cancel_numeral_factors
                addsimprocs [add_frac_frac_simproc, add_frac_num_simproc,
@@ -647,7 +537,12 @@
                 addcongs [@{thm "if_weak_cong"}]))
 then_conv (Simplifier.rewrite (HOL_basic_ss addsimps
   [@{thm numeral_1_eq_1},@{thm numeral_0_eq_0}] @ @{thms numerals(1-2)}))
+
 end
+*}
+
+declaration {*
+let
 
 fun numeral_is_const ct =
   case term_of ct of
@@ -674,16 +569,93 @@
   end
 
 in
- val field_comp_conv = comp_conv;
- val fieldgb_declaration = 
-  NormalizerData.funs @{thm class_fieldgb.fieldgb_axioms'}
+ 
+  Normalizer.funs @{thm normalizing.normalizing_field_cancel_axioms'}
    {is_const = K numeral_is_const,
     dest_const = K dest_const,
     mk_const = mk_const,
-    conv = K (K comp_conv)}
-end;
+    conv = K (K field_comp_conv)}
+
+end
+*}
+
+lemmas comp_arith = semiring_norm (*FIXME*)
+
+
+subsection {* Groebner Bases *}
+
+lemmas bool_simps = simp_thms(1-34)
+
+lemma dnf:
+    "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"
+    "(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"
+  by blast+
+
+lemmas weak_dnf_simps = dnf bool_simps
+
+lemma nnf_simps:
+    "(\<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)"
+    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
+  by blast+
+
+lemma PFalse:
+    "P \<equiv> False \<Longrightarrow> \<not> P"
+    "\<not> P \<Longrightarrow> (P \<equiv> False)"
+  by auto
+
+ML {*
+structure Algebra_Simplification = Named_Thms(
+  val name = "algebra"
+  val description = "pre-simplification rules for algebraic methods"
+)
 *}
 
-declaration fieldgb_declaration
+setup Algebra_Simplification.setup
+
+declare dvd_def[algebra]
+declare dvd_eq_mod_eq_0[symmetric, algebra]
+declare mod_div_trivial[algebra]
+declare mod_mod_trivial[algebra]
+declare conjunct1[OF DIVISION_BY_ZERO, algebra]
+declare conjunct2[OF DIVISION_BY_ZERO, algebra]
+declare zmod_zdiv_equality[symmetric,algebra]
+declare zdiv_zmod_equality[symmetric, algebra]
+declare zdiv_zminus_zminus[algebra]
+declare zmod_zminus_zminus[algebra]
+declare zdiv_zminus2[algebra]
+declare zmod_zminus2[algebra]
+declare zdiv_zero[algebra]
+declare zmod_zero[algebra]
+declare mod_by_1[algebra]
+declare div_by_1[algebra]
+declare zmod_minus1_right[algebra]
+declare zdiv_minus1_right[algebra]
+declare mod_div_trivial[algebra]
+declare mod_mod_trivial[algebra]
+declare mod_mult_self2_is_0[algebra]
+declare mod_mult_self1_is_0[algebra]
+declare zmod_eq_0_iff[algebra]
+declare dvd_0_left_iff[algebra]
+declare zdvd1_eq[algebra]
+declare zmod_eq_dvd_iff[algebra]
+declare nat_mod_eq_iff[algebra]
+
+use "Tools/Groebner_Basis/groebner.ML"
+
+method_setup algebra =
+{*
+let
+ fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+ val addN = "add"
+ val delN = "del"
+ val any_keyword = keyword addN || keyword delN
+ val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
+in
+  ((Scan.optional (keyword addN |-- thms) []) -- 
+   (Scan.optional (keyword delN |-- thms) [])) >>
+  (fn (add_ths, del_ths) => fn ctxt =>
+       SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))
+end
+*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
 
 end
--- a/src/HOL/Int.thy	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Int.thy	Thu May 06 19:27:51 2010 +0200
@@ -1063,20 +1063,24 @@
 
 text {* First version by Norbert Voelker *}
 
-definition (*for simplifying equalities*)
-  iszero :: "'a\<Colon>semiring_1 \<Rightarrow> bool"
-where
+definition (*for simplifying equalities*) iszero :: "'a\<Colon>semiring_1 \<Rightarrow> bool" where
   "iszero z \<longleftrightarrow> z = 0"
 
 lemma iszero_0: "iszero 0"
-by (simp add: iszero_def)
-
-lemma not_iszero_1: "~ iszero 1"
-by (simp add: iszero_def eq_commute)
+  by (simp add: iszero_def)
+
+lemma iszero_Numeral0: "iszero (Numeral0 :: 'a::number_ring)"
+  by (simp add: iszero_0)
+
+lemma not_iszero_1: "\<not> iszero 1"
+  by (simp add: iszero_def)
+
+lemma not_iszero_Numeral1: "\<not> iszero (Numeral1 :: 'a::number_ring)"
+  by (simp add: not_iszero_1)
 
 lemma eq_number_of_eq [simp]:
   "((number_of x::'a::number_ring) = number_of y) =
-   iszero (number_of (x + uminus y) :: 'a)"
+     iszero (number_of (x + uminus y) :: 'a)"
 unfolding iszero_def number_of_add number_of_minus
 by (simp add: algebra_simps)
 
--- a/src/HOL/IsaMakefile	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/IsaMakefile	Thu May 06 19:27:51 2010 +0200
@@ -284,9 +284,7 @@
   Tools/ATP_Manager/atp_manager.ML \
   Tools/ATP_Manager/atp_systems.ML \
   Tools/Groebner_Basis/groebner.ML \
-  Tools/Groebner_Basis/misc.ML \
   Tools/Groebner_Basis/normalizer.ML \
-  Tools/Groebner_Basis/normalizer_data.ML \
   Tools/choice_specification.ML \
   Tools/int_arith.ML \
   Tools/list_code.ML \
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu May 06 19:27:51 2010 +0200
@@ -1195,7 +1195,7 @@
 fun real_nonlinear_prover proof_method ctxt =
  let
   val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+      (the (Normalizer.match ctxt @{cterm "(0::real) + 1"}))
      simple_cterm_ord
   val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
        real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
@@ -1310,7 +1310,7 @@
 fun real_nonlinear_subst_prover prover ctxt =
  let
   val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+      (the (Normalizer.match ctxt @{cterm "(0::real) + 1"}))
      simple_cterm_ord
 
   val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
--- a/src/HOL/Library/normarith.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Library/normarith.ML	Thu May 06 19:27:51 2010 +0200
@@ -167,7 +167,7 @@
   (* FIXME : Should be computed statically!! *)
   val real_poly_conv = 
     Normalizer.semiring_normalize_wrapper ctxt
-     (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+     (the (Normalizer.match ctxt @{cterm "(0::real) + 1"}))
  in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (field_comp_conv then_conv real_poly_conv)))
 end;
 
@@ -278,7 +278,7 @@
    (* FIXME: Should be computed statically!!*)
    val real_poly_conv = 
       Normalizer.semiring_normalize_wrapper ctxt
-       (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"}))
+       (the (Normalizer.match ctxt @{cterm "(0::real) + 1"}))
    val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs
    val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) [] 
    val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" 
@@ -384,7 +384,7 @@
   let 
    val real_poly_neg_conv = #neg
        (Normalizer.semiring_normalizers_ord_wrapper ctxt
-        (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
+        (the (Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord)
    val (th1,th2) = conj_pair(rawrule th)
   in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
   end
--- a/src/HOL/Library/positivstellensatz.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Thu May 06 19:27:51 2010 +0200
@@ -748,7 +748,7 @@
   fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS
   val {add,mul,neg,pow,sub,main} = 
      Normalizer.semiring_normalizers_ord_wrapper ctxt
-      (the (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
+      (the (Normalizer.match ctxt @{cterm "(0::real) + 1"})) 
      simple_cterm_ord
 in gen_real_arith ctxt
    (cterm_of_rat, field_comp_conv, field_comp_conv,field_comp_conv,
--- a/src/HOL/Nat_Numeral.thy	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Nat_Numeral.thy	Thu May 06 19:27:51 2010 +0200
@@ -687,6 +687,20 @@
 lemmas nat_number' =
   nat_number_of_Bit0 nat_number_of_Bit1
 
+lemmas nat_arith =
+  add_nat_number_of
+  diff_nat_number_of
+  mult_nat_number_of
+  eq_nat_number_of
+  less_nat_number_of
+
+lemmas semiring_norm =
+  Let_def arith_simps nat_arith rel_simps neg_simps if_False
+  if_True add_0 add_Suc add_number_of_left mult_number_of_left
+  numeral_1_eq_1 [symmetric] Suc_eq_plus1
+  numeral_0_eq_0 [symmetric] numerals [symmetric]
+  iszero_simps not_iszero_Numeral1
+
 lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)"
   by (fact Let_def)
 
--- a/src/HOL/Tools/Groebner_Basis/groebner.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Tools/Groebner_Basis/groebner.ML	Thu May 06 19:27:51 2010 +0200
@@ -21,7 +21,7 @@
 structure Groebner : GROEBNER =
 struct
 
-open Conv Normalizer Drule Thm;
+open Conv Drule Thm;
 
 fun is_comb ct =
   (case Thm.term_of ct of
@@ -50,11 +50,11 @@
 val lcm_rat = fn x => fn y => Rat.rat_of_int (Integer.lcm (int_of_rat x) (int_of_rat y));
 
 val (eqF_intr, eqF_elim) =
-  let val [th1,th2] = thms "PFalse"
+  let val [th1,th2] = @{thms PFalse}
   in (fn th => th COMP th2, fn th => th COMP th1) end;
 
 val (PFalse, PFalse') =
- let val PFalse_eq = nth (thms "simp_thms") 13
+ let val PFalse_eq = nth @{thms simp_thms} 13
  in (PFalse_eq RS iffD1, PFalse_eq RS iffD2) end;
 
 
@@ -398,7 +398,7 @@
    compose_single(refute_disj rfn (dest_arg tm),2,compose_single(refute_disj rfn (dest_arg1 tm),2,disjE))
   | _ => rfn tm ;
 
-val notnotD = @{thm "notnotD"};
+val notnotD = @{thm notnotD};
 fun mk_binop ct x y = capply (capply ct x) y
 
 val mk_comb = capply;
@@ -440,10 +440,10 @@
 | _ => false;
 
 val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq;
-val bool_simps = @{thms "bool_simps"};
-val nnf_simps = @{thms "nnf_simps"};
+val bool_simps = @{thms bool_simps};
+val nnf_simps = @{thms nnf_simps};
 val nnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps bool_simps addsimps nnf_simps)
-val weak_dnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps @{thms "weak_dnf_simps"});
+val weak_dnf_conv = Simplifier.rewrite (HOL_basic_ss addsimps @{thms weak_dnf_simps});
 val initial_conv =
     Simplifier.rewrite
      (HOL_basic_ss addsimps nnf_simps
@@ -947,29 +947,31 @@
  case try (find_term 0) form of
   NONE => NONE
 | SOME tm =>
-  (case NormalizerData.match ctxt tm of
+  (case Normalizer.match ctxt tm of
     NONE => NONE
   | SOME (res as (theory, {is_const, dest_const, 
           mk_const, conv = ring_eq_conv})) =>
      SOME (ring_and_ideal_conv theory
           dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
-          (semiring_normalize_wrapper ctxt res)))
+          (Normalizer.semiring_normalize_wrapper ctxt res)))
 
 fun ring_solve ctxt form =
   (case try (find_term 0 (* FIXME !? *)) form of
     NONE => reflexive form
   | SOME tm =>
-      (case NormalizerData.match ctxt tm of
+      (case Normalizer.match ctxt tm of
         NONE => reflexive form
       | SOME (res as (theory, {is_const, dest_const, mk_const, conv = ring_eq_conv})) =>
         #ring_conv (ring_and_ideal_conv theory
           dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
-          (semiring_normalize_wrapper ctxt res)) form));
+          (Normalizer.semiring_normalize_wrapper ctxt res)) form));
+
+fun presimplify ctxt add_thms del_thms = asm_full_simp_tac (Simplifier.context ctxt
+  (HOL_basic_ss addsimps (Algebra_Simplification.get ctxt) delsimps del_thms addsimps add_thms));
 
 fun ring_tac add_ths del_ths ctxt =
   Object_Logic.full_atomize_tac
-  THEN' asm_full_simp_tac
-    (Simplifier.context ctxt (fst (NormalizerData.get ctxt)) delsimps del_ths addsimps add_ths)
+  THEN' presimplify ctxt add_ths del_ths
   THEN' CSUBGOAL (fn (p, i) =>
     rtac (let val form = Object_Logic.dest_judgment p
           in case get_ring_ideal_convs ctxt form of
@@ -988,8 +990,7 @@
    | exitac (SOME y) = rtac (instantiate' [SOME (ctyp_of_term y)] [NONE,SOME y] exI) 1
 in 
 fun ideal_tac add_ths del_ths ctxt = 
-  asm_full_simp_tac 
-   (Simplifier.context ctxt (fst (NormalizerData.get ctxt)) delsimps del_ths addsimps add_ths) 
+  presimplify ctxt add_ths del_ths
  THEN'
  CSUBGOAL (fn (p, i) =>
   case get_ring_ideal_convs ctxt p of
--- a/src/HOL/Tools/Groebner_Basis/misc.ML	Thu May 06 10:55:09 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,29 +0,0 @@
-(*  Title:      HOL/Tools/Groebner_Basis/misc.ML
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-
-Very basic stuff for cterms.
-*)
-
-structure Misc =
-struct
-
-fun is_comb ct =
-  (case Thm.term_of ct of
-    _ $ _ => true
-  | _ => false);
-
-val concl = Thm.cprop_of #> Thm.dest_arg;
-
-fun is_binop ct ct' =
-  (case Thm.term_of ct' of
-    c $ _ $ _ => term_of ct aconv c
-  | _ => false);
-
-fun dest_binop ct ct' =
-  if is_binop ct ct' then Thm.dest_binop ct'
-  else raise CTERM ("dest_binop: bad binop", [ct, ct'])
-
-fun inst_thm inst = Thm.instantiate ([], inst);
-
-end;
--- a/src/HOL/Tools/Groebner_Basis/normalizer.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML	Thu May 06 19:27:51 2010 +0200
@@ -1,30 +1,176 @@
 (*  Title:      HOL/Tools/Groebner_Basis/normalizer.ML
     Author:     Amine Chaieb, TU Muenchen
+
+Normalization of expressions in semirings.
 *)
 
 signature NORMALIZER = 
 sig
- val semiring_normalize_conv : Proof.context -> conv
- val semiring_normalize_ord_conv : Proof.context -> (cterm -> cterm -> bool) -> conv
- val semiring_normalize_tac : Proof.context -> int -> tactic
- val semiring_normalize_wrapper :  Proof.context -> NormalizerData.entry -> conv
- val semiring_normalizers_ord_wrapper :  
-     Proof.context -> NormalizerData.entry -> (cterm -> cterm -> bool) ->
+  type entry
+  val get: Proof.context -> (thm * entry) list
+  val match: Proof.context -> cterm -> entry option
+  val del: attribute
+  val add: {semiring: cterm list * thm list, ring: cterm list * thm list,
+    field: cterm list * thm list, idom: thm list, ideal: thm list} -> attribute
+  val funs: thm -> {is_const: morphism -> cterm -> bool,
+    dest_const: morphism -> cterm -> Rat.rat,
+    mk_const: morphism -> ctyp -> Rat.rat -> cterm,
+    conv: morphism -> Proof.context -> cterm -> thm} -> declaration
+
+  val semiring_normalize_conv: Proof.context -> conv
+  val semiring_normalize_ord_conv: Proof.context -> (cterm -> cterm -> bool) -> conv
+  val semiring_normalize_wrapper: Proof.context -> entry -> conv
+  val semiring_normalize_ord_wrapper: Proof.context -> entry
+    -> (cterm -> cterm -> bool) -> conv
+  val semiring_normalizers_conv: cterm list -> cterm list * thm list
+    -> cterm list * thm list -> cterm list * thm list ->
+      (cterm -> bool) * conv * conv * conv -> (cterm -> cterm -> bool) ->
+        {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv}
+  val semiring_normalizers_ord_wrapper:  Proof.context -> entry ->
+    (cterm -> cterm -> bool) ->
       {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv}
- val semiring_normalize_ord_wrapper :  Proof.context -> NormalizerData.entry ->
-   (cterm -> cterm -> bool) -> conv
- val semiring_normalizers_conv :
-     cterm list -> cterm list * thm list -> cterm list * thm list -> cterm list * thm list ->
-     (cterm -> bool) * conv * conv * conv -> (cterm -> cterm -> bool) ->
-       {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv}
+
+  val setup: theory -> theory
 end
 
 structure Normalizer: NORMALIZER = 
 struct
 
-open Conv;
+(** data **)
+
+type entry =
+ {vars: cterm list,
+  semiring: cterm list * thm list,
+  ring: cterm list * thm list,
+  field: cterm list * thm list,
+  idom: thm list,
+  ideal: thm list} *
+ {is_const: cterm -> bool,
+  dest_const: cterm -> Rat.rat,
+  mk_const: ctyp -> Rat.rat -> cterm,
+  conv: Proof.context -> cterm -> thm};
+
+structure Data = Generic_Data
+(
+  type T = (thm * entry) list;
+  val empty = [];
+  val extend = I;
+  val merge = AList.merge Thm.eq_thm (K false);
+);
+
+val get = Data.get o Context.Proof;
+
+fun match ctxt tm =
+  let
+    fun match_inst
+        ({vars, semiring = (sr_ops, sr_rules), 
+          ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal},
+         fns as {is_const, dest_const, mk_const, conv}) pat =
+       let
+        fun h instT =
+          let
+            val substT = Thm.instantiate (instT, []);
+            val substT_cterm = Drule.cterm_rule substT;
+
+            val vars' = map substT_cterm vars;
+            val semiring' = (map substT_cterm sr_ops, map substT sr_rules);
+            val ring' = (map substT_cterm r_ops, map substT r_rules);
+            val field' = (map substT_cterm f_ops, map substT f_rules);
+            val idom' = map substT idom;
+            val ideal' = map substT ideal;
+
+            val result = ({vars = vars', semiring = semiring', 
+                           ring = ring', field = field', idom = idom', ideal = ideal'}, fns);
+          in SOME result end
+      in (case try Thm.match (pat, tm) of
+           NONE => NONE
+         | SOME (instT, _) => h instT)
+      end;
+
+    fun match_struct (_,
+        entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) =
+      get_first (match_inst entry) (sr_ops @ r_ops @ f_ops);
+  in get_first match_struct (get ctxt) end;
+
+
+(* logical content *)
+
+val semiringN = "semiring";
+val ringN = "ring";
+val idomN = "idom";
+val idealN = "ideal";
+val fieldN = "field";
+
+fun undefined _ = raise Match;
 
-(* Very basic stuff for terms *)
+val del = Thm.declaration_attribute (Data.map o AList.delete Thm.eq_thm);
+
+fun add {semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), 
+         field = (f_ops, f_rules), idom, ideal} =
+  Thm.declaration_attribute (fn key => fn context => context |> Data.map
+    let
+      val ctxt = Context.proof_of context;
+
+      fun check kind name xs n =
+        null xs orelse length xs = n orelse
+        error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name);
+      val check_ops = check "operations";
+      val check_rules = check "rules";
+
+      val _ =
+        check_ops semiringN sr_ops 5 andalso
+        check_rules semiringN sr_rules 37 andalso
+        check_ops ringN r_ops 2 andalso
+        check_rules ringN r_rules 2 andalso
+        check_ops fieldN f_ops 2 andalso
+        check_rules fieldN f_rules 2 andalso
+        check_rules idomN idom 2;
+
+      val mk_meta = Local_Defs.meta_rewrite_rule ctxt;
+      val sr_rules' = map mk_meta sr_rules;
+      val r_rules' = map mk_meta r_rules;
+      val f_rules' = map mk_meta f_rules;
+
+      fun rule i = nth sr_rules' (i - 1);
+
+      val (cx, cy) = Thm.dest_binop (hd sr_ops);
+      val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
+      val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
+      val ((clx, crx), (cly, cry)) =
+        rule 13 |> Thm.rhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
+      val ((ca, cb), (cc, cd)) =
+        rule 20 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
+      val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg;
+      val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_arg;
+
+      val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry];
+      val semiring = (sr_ops, sr_rules');
+      val ring = (r_ops, r_rules');
+      val field = (f_ops, f_rules');
+      val ideal' = map (symmetric o mk_meta) ideal
+    in
+      AList.delete Thm.eq_thm key #>
+      cons (key, ({vars = vars, semiring = semiring, 
+                          ring = ring, field = field, idom = idom, ideal = ideal'},
+             {is_const = undefined, dest_const = undefined, mk_const = undefined,
+             conv = undefined}))
+    end);
+
+
+(* extra-logical functions *)
+
+fun funs raw_key {is_const, dest_const, mk_const, conv} phi = 
+ Data.map (fn data =>
+  let
+    val key = Morphism.thm phi raw_key;
+    val _ = AList.defined Thm.eq_thm data key orelse
+      raise THM ("No data entry for structure key", 0, [key]);
+    val fns = {is_const = is_const phi, dest_const = dest_const phi,
+      mk_const = mk_const phi, conv = conv phi};
+  in AList.map_entry Thm.eq_thm key (apsnd (K fns)) data end);
+
+
+(** auxiliary **)
 
 fun is_comb ct =
   (case Thm.term_of ct of
@@ -55,6 +201,7 @@
 val natarith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"},
                 @{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"}, 
                 @{thm "less_nat_number_of"}];
+
 val nat_add_conv = 
  zerone_conv 
   (Simplifier.rewrite 
@@ -64,13 +211,15 @@
                  @{thm add_number_of_left}, @{thm Suc_eq_plus1}]
              @ map (fn th => th RS sym) @{thms numerals}));
 
-val nat_mul_conv = nat_add_conv;
 val zeron_tm = @{cterm "0::nat"};
 val onen_tm  = @{cterm "1::nat"};
 val true_tm = @{cterm "True"};
 
 
-(* The main function! *)
+(** normalizing conversions **)
+
+(* core conversion *)
+
 fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) (f_ops, f_rules)
   (is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) =
 let
@@ -182,7 +331,7 @@
           then
             let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34
                 val (l,r) = Thm.dest_comb(concl th1)
-           in transitive th1 (Drule.arg_cong_rule l (nat_mul_conv r))
+           in transitive th1 (Drule.arg_cong_rule l (nat_add_conv r))
            end
            else
             if opr aconvc mul_tm
@@ -563,7 +712,7 @@
    let val (l,r) = Thm.dest_comb tm in
         if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else
         let val th1 = inst_thm [(cx',r)] neg_mul
-            val th2 = transitive th1 (arg1_conv semiring_mul_conv (concl th1))
+            val th2 = transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1))
         in transitive th2 (polynomial_monomial_mul_conv (concl th2))
         end
    end;
@@ -606,7 +755,7 @@
             then
               let val th1 = combination (Drule.arg_cong_rule opr (polynomial_conv l)) 
                                         (polynomial_conv r)
-                  val th2 = (rewr_conv divide_inverse then_conv polynomial_mul_conv) 
+                  val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv)
                               (Thm.rhs_of th1)
               in transitive th1 th2
               end
@@ -638,11 +787,14 @@
 
 fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
 
+
+(* various normalizing conversions *)
+
 fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, 
                                      {conv, dest_const, mk_const, is_const}) ord =
   let
     val pow_conv =
-      arg_conv (Simplifier.rewrite nat_exp_ss)
+      Conv.arg_conv (Simplifier.rewrite nat_exp_ss)
       then_conv Simplifier.rewrite
         (HOL_basic_ss addsimps [nth (snd semiring) 31, nth (snd semiring) 34])
       then_conv conv ctxt
@@ -656,14 +808,57 @@
   semiring_normalize_ord_wrapper ctxt data simple_cterm_ord;
 
 fun semiring_normalize_ord_conv ctxt ord tm =
-  (case NormalizerData.match ctxt tm of
+  (case match ctxt tm of
     NONE => reflexive tm
   | SOME res => semiring_normalize_ord_wrapper ctxt res ord tm);
  
-
 fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt simple_cterm_ord;
 
-fun semiring_normalize_tac ctxt = SUBGOAL (fn (goal, i) =>
-  rtac (semiring_normalize_conv ctxt
-    (cterm_of (ProofContext.theory_of ctxt) (fst (Logic.dest_equals goal)))) i);
+
+(** Isar setup **)
+
+local
+
+fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
+fun keyword2 k1 k2 = Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.colon) >> K ();
+fun keyword3 k1 k2 k3 =
+  Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.$$$ k3 -- Args.colon) >> K ();
+
+val opsN = "ops";
+val rulesN = "rules";
+
+val normN = "norm";
+val constN = "const";
+val delN = "del";
+
+val any_keyword =
+  keyword2 semiringN opsN || keyword2 semiringN rulesN ||
+  keyword2 ringN opsN || keyword2 ringN rulesN ||
+  keyword2 fieldN opsN || keyword2 fieldN rulesN ||
+  keyword2 idomN rulesN || keyword2 idealN rulesN;
+
+val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
+val terms = thms >> map Drule.dest_term;
+
+fun optional scan = Scan.optional scan [];
+
+in
+
+val setup =
+  Attrib.setup @{binding normalizer}
+    (Scan.lift (Args.$$$ delN >> K del) ||
+      ((keyword2 semiringN opsN |-- terms) --
+       (keyword2 semiringN rulesN |-- thms)) --
+      (optional (keyword2 ringN opsN |-- terms) --
+       optional (keyword2 ringN rulesN |-- thms)) --
+      (optional (keyword2 fieldN opsN |-- terms) --
+       optional (keyword2 fieldN rulesN |-- thms)) --
+      optional (keyword2 idomN rulesN |-- thms) --
+      optional (keyword2 idealN rulesN |-- thms)
+      >> (fn ((((sr, r), f), id), idl) => 
+             add {semiring = sr, ring = r, field = f, idom = id, ideal = idl}))
+    "semiring normalizer data";
+
 end;
+
+end;
--- a/src/HOL/Tools/Groebner_Basis/normalizer_data.ML	Thu May 06 10:55:09 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,227 +0,0 @@
-(*  Title:      HOL/Tools/Groebner_Basis/normalizer_data.ML
-    ID:         $Id$
-    Author:     Amine Chaieb, TU Muenchen
-
-Ring normalization data.
-*)
-
-signature NORMALIZER_DATA =
-sig
-  type entry
-  val get: Proof.context -> simpset * (thm * entry) list
-  val match: Proof.context -> cterm -> entry option
-  val del: attribute
-  val add: {semiring: cterm list * thm list, ring: cterm list * thm list, field: cterm list * thm list, idom: thm list, ideal: thm list}
-    -> attribute
-  val funs: thm -> {is_const: morphism -> cterm -> bool,
-    dest_const: morphism -> cterm -> Rat.rat,
-    mk_const: morphism -> ctyp -> Rat.rat -> cterm,
-    conv: morphism -> Proof.context -> cterm -> thm} -> declaration
-  val setup: theory -> theory
-end;
-
-structure NormalizerData: NORMALIZER_DATA =
-struct
-
-(* data *)
-
-type entry =
- {vars: cterm list,
-  semiring: cterm list * thm list,
-  ring: cterm list * thm list,
-  field: cterm list * thm list,
-  idom: thm list,
-  ideal: thm list} *
- {is_const: cterm -> bool,
-  dest_const: cterm -> Rat.rat,
-  mk_const: ctyp -> Rat.rat -> cterm,
-  conv: Proof.context -> cterm -> thm};
-
-val eq_key = Thm.eq_thm;
-fun eq_data arg = eq_fst eq_key arg;
-
-structure Data = Generic_Data
-(
-  type T = simpset * (thm * entry) list;
-  val empty = (HOL_basic_ss, []);
-  val extend = I;
-  fun merge ((ss, e), (ss', e')) : T =
-    (merge_ss (ss, ss'), AList.merge eq_key (K true) (e, e'));
-);
-
-val get = Data.get o Context.Proof;
-
-
-(* match data *)
-
-fun match ctxt tm =
-  let
-    fun match_inst
-        ({vars, semiring = (sr_ops, sr_rules), 
-          ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal},
-         fns as {is_const, dest_const, mk_const, conv}) pat =
-       let
-        fun h instT =
-          let
-            val substT = Thm.instantiate (instT, []);
-            val substT_cterm = Drule.cterm_rule substT;
-
-            val vars' = map substT_cterm vars;
-            val semiring' = (map substT_cterm sr_ops, map substT sr_rules);
-            val ring' = (map substT_cterm r_ops, map substT r_rules);
-            val field' = (map substT_cterm f_ops, map substT f_rules);
-            val idom' = map substT idom;
-            val ideal' = map substT ideal;
-
-            val result = ({vars = vars', semiring = semiring', 
-                           ring = ring', field = field', idom = idom', ideal = ideal'}, fns);
-          in SOME result end
-      in (case try Thm.match (pat, tm) of
-           NONE => NONE
-         | SOME (instT, _) => h instT)
-      end;
-
-    fun match_struct (_,
-        entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) =
-      get_first (match_inst entry) (sr_ops @ r_ops @ f_ops);
-  in get_first match_struct (snd (get ctxt)) end;
-
-
-(* logical content *)
-
-val semiringN = "semiring";
-val ringN = "ring";
-val idomN = "idom";
-val idealN = "ideal";
-val fieldN = "field";
-
-fun undefined _ = raise Match;
-
-fun del_data key = apsnd (remove eq_data (key, []));
-
-val del = Thm.declaration_attribute (Data.map o del_data);
-val add_ss = Thm.declaration_attribute 
-   (fn th => Data.map (fn (ss,data) => (ss addsimps [th], data)));
-
-val del_ss = Thm.declaration_attribute 
-   (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
-
-fun add {semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), 
-         field = (f_ops, f_rules), idom, ideal} =
-  Thm.declaration_attribute (fn key => fn context => context |> Data.map
-    let
-      val ctxt = Context.proof_of context;
-
-      fun check kind name xs n =
-        null xs orelse length xs = n orelse
-        error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name);
-      val check_ops = check "operations";
-      val check_rules = check "rules";
-
-      val _ =
-        check_ops semiringN sr_ops 5 andalso
-        check_rules semiringN sr_rules 37 andalso
-        check_ops ringN r_ops 2 andalso
-        check_rules ringN r_rules 2 andalso
-        check_ops fieldN f_ops 2 andalso
-        check_rules fieldN f_rules 2 andalso
-        check_rules idomN idom 2;
-
-      val mk_meta = Local_Defs.meta_rewrite_rule ctxt;
-      val sr_rules' = map mk_meta sr_rules;
-      val r_rules' = map mk_meta r_rules;
-      val f_rules' = map mk_meta f_rules;
-
-      fun rule i = nth sr_rules' (i - 1);
-
-      val (cx, cy) = Thm.dest_binop (hd sr_ops);
-      val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
-      val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
-      val ((clx, crx), (cly, cry)) =
-        rule 13 |> Thm.rhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
-      val ((ca, cb), (cc, cd)) =
-        rule 20 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
-      val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg;
-      val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_arg;
-
-      val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry];
-      val semiring = (sr_ops, sr_rules');
-      val ring = (r_ops, r_rules');
-      val field = (f_ops, f_rules');
-      val ideal' = map (symmetric o mk_meta) ideal
-    in
-      del_data key #>
-      apsnd (cons (key, ({vars = vars, semiring = semiring, 
-                          ring = ring, field = field, idom = idom, ideal = ideal'},
-             {is_const = undefined, dest_const = undefined, mk_const = undefined,
-             conv = undefined})))
-    end);
-
-
-(* extra-logical functions *)
-
-fun funs raw_key {is_const, dest_const, mk_const, conv} phi = 
- (Data.map o apsnd) (fn data =>
-  let
-    val key = Morphism.thm phi raw_key;
-    val _ = AList.defined eq_key data key orelse
-      raise THM ("No data entry for structure key", 0, [key]);
-    val fns = {is_const = is_const phi, dest_const = dest_const phi,
-      mk_const = mk_const phi, conv = conv phi};
-  in AList.map_entry eq_key key (apsnd (K fns)) data end);
-
-
-(* concrete syntax *)
-
-local
-
-fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
-fun keyword2 k1 k2 = Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.colon) >> K ();
-fun keyword3 k1 k2 k3 =
-  Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.$$$ k3 -- Args.colon) >> K ();
-
-val opsN = "ops";
-val rulesN = "rules";
-
-val normN = "norm";
-val constN = "const";
-val delN = "del";
-
-val any_keyword =
-  keyword2 semiringN opsN || keyword2 semiringN rulesN ||
-  keyword2 ringN opsN || keyword2 ringN rulesN ||
-  keyword2 fieldN opsN || keyword2 fieldN rulesN ||
-  keyword2 idomN rulesN || keyword2 idealN rulesN;
-
-val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-val terms = thms >> map Drule.dest_term;
-
-fun optional scan = Scan.optional scan [];
-
-in
-
-val normalizer_setup =
-  Attrib.setup @{binding normalizer}
-    (Scan.lift (Args.$$$ delN >> K del) ||
-      ((keyword2 semiringN opsN |-- terms) --
-       (keyword2 semiringN rulesN |-- thms)) --
-      (optional (keyword2 ringN opsN |-- terms) --
-       optional (keyword2 ringN rulesN |-- thms)) --
-      (optional (keyword2 fieldN opsN |-- terms) --
-       optional (keyword2 fieldN rulesN |-- thms)) --
-      optional (keyword2 idomN rulesN |-- thms) --
-      optional (keyword2 idealN rulesN |-- thms)
-      >> (fn ((((sr, r), f), id), idl) => 
-             add {semiring = sr, ring = r, field = f, idom = id, ideal = idl}))
-    "semiring normalizer data";
-
-end;
-
-
-(* theory setup *)
-
-val setup =
-  normalizer_setup #>
-  Attrib.setup @{binding algebra} (Attrib.add_del add_ss del_ss) "pre-simplification for algebra";
-
-end;
--- a/src/HOL/Tools/Qelim/cooper.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Thu May 06 19:27:51 2010 +0200
@@ -3,7 +3,7 @@
 *)
 
 signature COOPER =
- sig
+sig
   val cooper_conv : Proof.context -> conv
   exception COOPER of string * exn
 end;
@@ -12,7 +12,6 @@
 struct
 
 open Conv;
-open Normalizer;
 
 exception COOPER of string * exn;
 fun simp_thms_conv ctxt =
--- a/src/HOL/Tools/Qelim/cooper_data.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper_data.ML	Thu May 06 19:27:51 2010 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Tools/Qelim/cooper_data.ML
-    ID:         $Id$
     Author:     Amine Chaieb, TU Muenchen
 *)
 
@@ -16,8 +15,7 @@
 struct
 
 type entry = simpset * (term list);
-val start_ss = HOL_ss (* addsimps @{thms "Groebner_Basis.comp_arith"}
-               addcongs [if_weak_cong, @{thm "let_weak_cong"}];*)
+
 val allowed_consts = 
   [@{term "op + :: int => _"}, @{term "op + :: nat => _"},
    @{term "op - :: int => _"}, @{term "op - :: nat => _"},
@@ -47,7 +45,7 @@
 structure Data = Generic_Data
 (
   type T = simpset * term list;
-  val empty = (start_ss, allowed_consts);
+  val empty = (HOL_ss, allowed_consts);
   val extend  = I;
   fun merge ((ss1, ts1), (ss2, ts2)) =
     (merge_ss (ss1, ss2), Library.merge (op aconv) (ts1, ts2));
@@ -64,7 +62,7 @@
      (ss delsimps [th], subtract (op aconv) ts' ts ))) 
 
 
-(* concrete syntax *)
+(* theory setup *)
 
 local
 
@@ -79,16 +77,11 @@
 
 in
 
-val presburger_setup =
+val setup =
   Attrib.setup @{binding presburger}
     ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del ||
       optional (keyword constsN |-- terms) >> add) "Cooper data";
 
 end;
 
-
-(* theory setup *)
-
-val setup = presburger_setup;
-
 end;
--- a/src/HOL/Tools/Qelim/presburger.ML	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/Tools/Qelim/presburger.ML	Thu May 06 19:27:51 2010 +0200
@@ -11,7 +11,7 @@
 struct
 
 open Conv;
-val comp_ss = HOL_ss addsimps @{thms "Groebner_Basis.comp_arith"};
+val comp_ss = HOL_ss addsimps @{thms semiring_norm};
 
 fun strip_objimp ct =
   (case Thm.term_of ct of
--- a/src/HOL/ex/Groebner_Examples.thy	Thu May 06 10:55:09 2010 +0200
+++ b/src/HOL/ex/Groebner_Examples.thy	Thu May 06 19:27:51 2010 +0200
@@ -10,18 +10,30 @@
 
 subsection {* Basic examples *}
 
-schematic_lemma "3 ^ 3 == (?X::'a::{number_ring})"
-  by sring_norm
+lemma
+  fixes x :: int
+  shows "x ^ 3 = x ^ 3" 
+  apply (tactic {* ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_conv @{context})))) *})
+  by (rule refl)
 
-schematic_lemma "(x - (-2))^5 == ?X::int"
-  by sring_norm
+lemma
+  fixes x :: int
+  shows "(x - (-2))^5 = x ^ 5 + (10 * x ^ 4 + (40 * x ^ 3 + (80 * x² + (80 * x + 32))))" 
+  apply (tactic {* ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_conv @{context})))) *})
+  by (rule refl)
 
-schematic_lemma "(x - (-2))^5  * (y - 78) ^ 8 == ?X::int"
-  by sring_norm
+schematic_lemma
+  fixes x :: int
+  shows "(x - (-2))^5  * (y - 78) ^ 8 = ?X" 
+  apply (tactic {* ALLGOALS (CONVERSION
+    (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_conv @{context})))) *})
+  by (rule refl)
 
 lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{number_ring})"
   apply (simp only: power_Suc power_0)
-  apply (simp only: comp_arith)
+  apply (simp only: semiring_norm)
   oops
 
 lemma "((x::int) + y)^3 - 1 = (x - z)^2 - 10 \<Longrightarrow> x = z + 3 \<Longrightarrow> x = - y"