# HG changeset patch # User haftmann # Date 1273237552 -7200 # Node ID 7f1da69cacb3a1bfab8112fd2d7b0d2c8bfb3495 # Parent 912080b2c4495dc7bf99db1618b249b2e42600cd split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Groebner_Basis.thy Fri May 07 15:05:52 2010 +0200 @@ -2,341 +2,14 @@ Author: Amine Chaieb, TU Muenchen *) -header {* Semiring normalization and Groebner Bases *} +header {* Groebner bases *} theory Groebner_Basis -imports Numeral_Simprocs Nat_Transfer +imports Semiring_Normalization uses - "Tools/Groebner_Basis/normalizer.ML" ("Tools/Groebner_Basis/groebner.ML") begin -subsection {* Semiring normalization *} - -setup Normalizer.setup - -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" - and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x" - and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0" - and mul_d:"mul x (add y z) = add (mul x y) (mul x z)" - and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)" -begin - -lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)" -proof (induct p) - case 0 - then show ?case by (auto simp add: pwr_0 mul_1) -next - case Suc - from this [symmetric] show ?case - by (auto simp add: pwr_Suc mul_1 mul_a) -qed - -lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)" -proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1) - fix q x y - assume "\x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)" - have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))" - by (simp add: mul_a) - also have "\ = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c) - also have "\ = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a) - finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) = - mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c) -qed - -lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)" -proof (induct p arbitrary: q) - case 0 - show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto -next - case Suc - thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc) -qed - -lemma semiring_ops: - shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)" - and "TERM r0" and "TERM r1" . - -lemma semiring_rules: - "add (mul a m) (mul b m) = mul (add a b) m" - "add (mul a m) m = mul (add a r1) m" - "add m (mul a m) = mul (add a r1) m" - "add m m = mul (add r1 r1) m" - "add r0 a = a" - "add a r0 = a" - "mul a b = mul b a" - "mul (add a b) c = add (mul a c) (mul b c)" - "mul r0 a = r0" - "mul a r0 = r0" - "mul r1 a = a" - "mul a r1 = a" - "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" - "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" - "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" - "mul (mul lx ly) rx = mul (mul lx rx) ly" - "mul (mul lx ly) rx = mul lx (mul ly rx)" - "mul lx (mul rx ry) = mul (mul lx rx) ry" - "mul lx (mul rx ry) = mul rx (mul lx ry)" - "add (add a b) (add c d) = add (add a c) (add b d)" - "add (add a b) c = add a (add b c)" - "add a (add c d) = add c (add a d)" - "add (add a b) c = add (add a c) b" - "add a c = add c a" - "add a (add c d) = add (add a c) d" - "mul (pwr x p) (pwr x q) = pwr x (p + q)" - "mul x (pwr x q) = pwr x (Suc q)" - "mul (pwr x q) x = pwr x (Suc q)" - "mul x x = pwr x 2" - "pwr (mul x y) q = mul (pwr x q) (pwr y q)" - "pwr (pwr x p) q = pwr x (p * q)" - "pwr x 0 = r1" - "pwr x 1 = x" - "mul x (add y z) = add (mul x y) (mul x z)" - "pwr x (Suc q) = mul x (pwr x q)" - "pwr x (2*n) = mul (pwr x n) (pwr x n)" - "pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))" -proof - - show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp -next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp -next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp -next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp -next show "add r0 a = a" using add_0 by simp -next show "add a r0 = a" using add_0 add_c by simp -next show "mul a b = mul b a" using mul_c by simp -next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp -next show "mul r0 a = r0" using mul_0 by simp -next show "mul a r0 = r0" using mul_0 mul_c by simp -next show "mul r1 a = a" using mul_1 by simp -next show "mul a r1 = a" using mul_1 mul_c by simp -next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" - using mul_c mul_a by simp -next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" - using mul_a by simp -next - have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c) - also have "\ = mul rx (mul ry (mul lx ly))" using mul_a by simp - finally - show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" - using mul_c by simp -next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp -next - show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a) -next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a ) -next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c) -next show "add (add a b) (add c d) = add (add a c) (add b d)" - using add_c add_a by simp -next show "add (add a b) c = add a (add b c)" using add_a by simp -next show "add a (add c d) = add c (add a d)" - apply (simp add: add_a) by (simp only: add_c) -next show "add (add a b) c = add (add a c) b" using add_a add_c by simp -next show "add a c = add c a" by (rule add_c) -next show "add a (add c d) = add (add a c) d" using add_a by simp -next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr) -next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp -next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp -next show "mul x x = pwr x 2" by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) -next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul) -next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr) -next show "pwr x 0 = r1" using pwr_0 . -next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) -next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp -next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp -next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number' mul_pwr) -next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))" - by (simp add: nat_number' pwr_Suc mul_pwr) -qed - - -lemmas normalizing_semiring_axioms' = - normalizing_semiring_axioms [normalizer - semiring ops: semiring_ops - semiring rules: semiring_rules] - -end - -sublocale comm_semiring_1 - < normalizing!: normalizing_semiring plus times power zero one -proof -qed (simp_all add: algebra_simps) - -declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_semiring_axioms'} *} - -locale normalizing_ring = normalizing_semiring + - fixes sub :: "'a \ 'a \ 'a" - and neg :: "'a \ 'a" - assumes neg_mul: "neg x = mul (neg r1) x" - and sub_add: "sub x y = add x (neg y)" -begin - -lemma ring_ops: shows "TERM (sub x y)" and "TERM (neg x)" . - -lemmas ring_rules = neg_mul sub_add - -lemmas normalizing_ring_axioms' = - normalizing_ring_axioms [normalizer - semiring ops: semiring_ops - semiring rules: semiring_rules - ring ops: ring_ops - ring rules: ring_rules] - -end - -sublocale comm_ring_1 - < normalizing!: normalizing_ring plus times power zero one minus uminus -proof -qed (simp_all add: diff_minus) - -declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_axioms'} *} - -locale normalizing_field = normalizing_ring + - fixes divide :: "'a \ 'a \ 'a" - and inverse:: "'a \ 'a" - assumes divide_inverse: "divide x y = mul x (inverse y)" - and inverse_divide: "inverse x = divide r1 x" -begin - -lemma field_ops: shows "TERM (divide x y)" and "TERM (inverse x)" . - -lemmas field_rules = divide_inverse inverse_divide - -lemmas normalizing_field_axioms' = - normalizing_field_axioms [normalizer - semiring ops: semiring_ops - semiring rules: semiring_rules - ring ops: ring_ops - ring rules: ring_rules - field ops: field_ops - field rules: field_rules] - -end - -locale normalizing_semiring_cancel = normalizing_semiring + - assumes add_cancel: "add (x::'a) y = add x z \ y = z" - and add_mul_solve: "add (mul w y) (mul x z) = - add (mul w z) (mul x y) \ w = x \ y = z" -begin - -lemma noteq_reduce: "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" -proof- - have "a \ b \ c \ d \ \ (a = b \ c = d)" by simp - also have "\ \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" - using add_mul_solve by blast - finally show "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" - by simp -qed - -lemma add_scale_eq_noteq: "\r \ r0 ; (a = b) \ ~(c = d)\ - \ add a (mul r c) \ add b (mul r d)" -proof(clarify) - assume nz: "r\ r0" and cnd: "c\d" - and eq: "add b (mul r c) = add b (mul r d)" - hence "mul r c = mul r d" using cnd add_cancel by simp - hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)" - using mul_0 add_cancel by simp - thus "False" using add_mul_solve nz cnd by simp -qed - -lemma add_r0_iff: " x = add x a \ a = r0" -proof- - have "a = r0 \ add x a = add x r0" by (simp add: add_cancel) - thus "x = add x a \ a = r0" by (auto simp add: add_c add_0) -qed - -declare normalizing_semiring_axioms' [normalizer del] - -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 normalizing_ring_cancel = normalizing_semiring_cancel + normalizing_ring + - assumes subr0_iff: "sub x y = r0 \ x = y" -begin - -declare normalizing_ring_axioms' [normalizer del] - -lemmas normalizing_ring_cancel_axioms' = normalizing_ring_cancel_axioms [normalizer - semiring ops: semiring_ops - semiring rules: semiring_rules - ring ops: ring_ops - ring rules: ring_rules - idom rules: noteq_reduce add_scale_eq_noteq - ideal rules: subr0_iff add_r0_iff] - -end - -sublocale idom - < normalizing!: normalizing_ring_cancel plus times power zero one minus uminus -proof - fix w x y z - show "w * y + x * z = w * z + x * y \ w = x \ y = z" - proof - assume "w * y + x * z = w * z + x * y" - then have "w * y + x * z - w * z - x * y = 0" by (simp add: algebra_simps) - then have "w * (y - z) - x * (y - z) = 0" by (simp add: algebra_simps) - then have "(y - z) * (w - x) = 0" by (simp add: algebra_simps) - then have "y - z = 0 \ w - x = 0" by (rule divisors_zero) - then show "w = x \ y = z" by auto - qed (auto simp add: add_ac) -qed (simp_all add: algebra_simps) - -declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_cancel_axioms'} *} - -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" - { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \ z" - hence "y < z \ y > z" by arith - moreover { - assume lt:"y k. z = y + k \ k > 0" by (rule_tac x="z - y" in exI, auto) - then obtain k where kp: "k>0" and yz:"z = y + k" by blast - from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps) - hence "x*k = w*k" by simp - hence "w = x" using kp by simp } - moreover { - assume lt: "y >z" hence "\k. y = z + k \ k>0" by (rule_tac x="y - z" in exI, auto) - then obtain k where kp: "k>0" and yz:"y = z + k" by blast - from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps) - hence "w*k = x*k" by simp - hence "w = x" using kp by simp } - ultimately have "w=x" by blast } - thus "(w * y + x * z = w * z + x * y) = (w = x \ y = z)" by auto -qed - -declaration {* Normalizer.semiring_funs @{thm normalizing_nat.normalizing_semiring_cancel_axioms'} *} - -locale normalizing_field_cancel = normalizing_ring_cancel + normalizing_field -begin - -declare normalizing_field_axioms' [normalizer del] - -lemmas normalizing_field_cancel_axioms' = normalizing_field_cancel_axioms [normalizer - semiring ops: semiring_ops - semiring rules: semiring_rules - ring ops: ring_ops - ring rules: ring_rules - field ops: field_ops - field rules: field_rules - idom rules: noteq_reduce add_scale_eq_noteq - ideal rules: subr0_iff add_r0_iff] - -end - -sublocale field - < normalizing!: normalizing_field_cancel plus times power zero one minus uminus divide inverse -proof -qed (simp_all add: divide_inverse) - -declaration {* Normalizer.field_funs @{thm normalizing.normalizing_field_cancel_axioms'} *} - - subsection {* Groebner Bases *} lemmas bool_simps = simp_thms(1-34) @@ -367,6 +40,11 @@ setup Algebra_Simplification.setup +use "Tools/Groebner_Basis/groebner.ML" + +method_setup algebra = Groebner.algebra_method + "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] @@ -395,9 +73,4 @@ declare zmod_eq_dvd_iff[algebra] declare nat_mod_eq_iff[algebra] -use "Tools/Groebner_Basis/groebner.ML" - -method_setup algebra = Groebner.algebra_method - "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases" - end diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/IsaMakefile Fri May 07 15:05:52 2010 +0200 @@ -271,6 +271,7 @@ Random.thy \ Random_Sequence.thy \ Recdef.thy \ + Semiring_Normalization.thy \ SetInterval.thy \ Sledgehammer.thy \ String.thy \ diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML --- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Fri May 07 15:05:52 2010 +0200 @@ -1222,7 +1222,7 @@ in (let val th = tryfind trivial_axiom (keq @ klep @ kltp) in - (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv Normalizer.field_comp_conv) th, RealArith.Trivial) + (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv Numeral_Simprocs.field_comp_conv) th, RealArith.Trivial) end) handle Failure _ => (let val proof = diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Library/normarith.ML --- a/src/HOL/Library/normarith.ML Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Library/normarith.ML Fri May 07 15:05:52 2010 +0200 @@ -168,7 +168,7 @@ val real_poly_conv = Normalizer.semiring_normalize_wrapper ctxt (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 (Normalizer.field_comp_conv then_conv real_poly_conv))) + in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv then_conv real_poly_conv))) end; fun absc cv ct = case term_of ct of @@ -190,8 +190,8 @@ val apply_pth5 = rewr_conv @{thm pth_5}; val apply_pth6 = rewr_conv @{thm pth_6}; val apply_pth7 = rewrs_conv @{thms pth_7}; - val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv Normalizer.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left}))); - val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Normalizer.field_comp_conv); + val apply_pth8 = rewr_conv @{thm pth_8} then_conv arg1_conv Numeral_Simprocs.field_comp_conv then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left}))); + val apply_pth9 = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv Numeral_Simprocs.field_comp_conv); val apply_ptha = rewr_conv @{thm pth_a}; val apply_pthb = rewrs_conv @{thms pth_b}; val apply_pthc = rewrs_conv @{thms pth_c}; @@ -204,7 +204,7 @@ | _ => error "headvector: non-canonical term" fun vector_cmul_conv ct = - ((apply_pth5 then_conv arg1_conv Normalizer.field_comp_conv) else_conv + ((apply_pth5 then_conv arg1_conv Numeral_Simprocs.field_comp_conv) else_conv (apply_pth6 then_conv binop_conv vector_cmul_conv)) ct fun vector_add_conv ct = apply_pth7 ct @@ -396,7 +396,7 @@ fun init_conv ctxt = Simplifier.rewrite (Simplifier.context ctxt (HOL_basic_ss addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm diff_0_right}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths}))) - then_conv Normalizer.field_comp_conv + then_conv Numeral_Simprocs.field_comp_conv then_conv nnf_conv fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt); diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Library/positivstellensatz.ML --- a/src/HOL/Library/positivstellensatz.ML Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Library/positivstellensatz.ML Fri May 07 15:05:52 2010 +0200 @@ -751,7 +751,7 @@ (the (Normalizer.match ctxt @{cterm "(0::real) + 1"})) simple_cterm_ord in gen_real_arith ctxt - (cterm_of_rat, Normalizer.field_comp_conv, Normalizer.field_comp_conv, Normalizer.field_comp_conv, + (cterm_of_rat, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv, Numeral_Simprocs.field_comp_conv, main,neg,add,mul, prover) end; diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Semiring_Normalization.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Semiring_Normalization.thy Fri May 07 15:05:52 2010 +0200 @@ -0,0 +1,336 @@ +(* Title: HOL/Semiring_Normalization.thy + Author: Amine Chaieb, TU Muenchen +*) + +header {* Semiring normalization *} + +theory Semiring_Normalization +imports Numeral_Simprocs Nat_Transfer +uses + "Tools/Groebner_Basis/normalizer.ML" +begin + +setup Normalizer.setup + +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" + and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x" + and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0" + and mul_d:"mul x (add y z) = add (mul x y) (mul x z)" + and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)" +begin + +lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)" +proof (induct p) + case 0 + then show ?case by (auto simp add: pwr_0 mul_1) +next + case Suc + from this [symmetric] show ?case + by (auto simp add: pwr_Suc mul_1 mul_a) +qed + +lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)" +proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1) + fix q x y + assume "\x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)" + have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))" + by (simp add: mul_a) + also have "\ = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c) + also have "\ = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a) + finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) = + mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c) +qed + +lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)" +proof (induct p arbitrary: q) + case 0 + show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto +next + case Suc + thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc) +qed + +lemma semiring_ops: + shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)" + and "TERM r0" and "TERM r1" . + +lemma semiring_rules: + "add (mul a m) (mul b m) = mul (add a b) m" + "add (mul a m) m = mul (add a r1) m" + "add m (mul a m) = mul (add a r1) m" + "add m m = mul (add r1 r1) m" + "add r0 a = a" + "add a r0 = a" + "mul a b = mul b a" + "mul (add a b) c = add (mul a c) (mul b c)" + "mul r0 a = r0" + "mul a r0 = r0" + "mul r1 a = a" + "mul a r1 = a" + "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" + "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" + "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" + "mul (mul lx ly) rx = mul (mul lx rx) ly" + "mul (mul lx ly) rx = mul lx (mul ly rx)" + "mul lx (mul rx ry) = mul (mul lx rx) ry" + "mul lx (mul rx ry) = mul rx (mul lx ry)" + "add (add a b) (add c d) = add (add a c) (add b d)" + "add (add a b) c = add a (add b c)" + "add a (add c d) = add c (add a d)" + "add (add a b) c = add (add a c) b" + "add a c = add c a" + "add a (add c d) = add (add a c) d" + "mul (pwr x p) (pwr x q) = pwr x (p + q)" + "mul x (pwr x q) = pwr x (Suc q)" + "mul (pwr x q) x = pwr x (Suc q)" + "mul x x = pwr x 2" + "pwr (mul x y) q = mul (pwr x q) (pwr y q)" + "pwr (pwr x p) q = pwr x (p * q)" + "pwr x 0 = r1" + "pwr x 1 = x" + "mul x (add y z) = add (mul x y) (mul x z)" + "pwr x (Suc q) = mul x (pwr x q)" + "pwr x (2*n) = mul (pwr x n) (pwr x n)" + "pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))" +proof - + show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp +next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp +next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp +next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp +next show "add r0 a = a" using add_0 by simp +next show "add a r0 = a" using add_0 add_c by simp +next show "mul a b = mul b a" using mul_c by simp +next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp +next show "mul r0 a = r0" using mul_0 by simp +next show "mul a r0 = r0" using mul_0 mul_c by simp +next show "mul r1 a = a" using mul_1 by simp +next show "mul a r1 = a" using mul_1 mul_c by simp +next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" + using mul_c mul_a by simp +next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" + using mul_a by simp +next + have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c) + also have "\ = mul rx (mul ry (mul lx ly))" using mul_a by simp + finally + show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" + using mul_c by simp +next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp +next + show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a) +next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a ) +next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c) +next show "add (add a b) (add c d) = add (add a c) (add b d)" + using add_c add_a by simp +next show "add (add a b) c = add a (add b c)" using add_a by simp +next show "add a (add c d) = add c (add a d)" + apply (simp add: add_a) by (simp only: add_c) +next show "add (add a b) c = add (add a c) b" using add_a add_c by simp +next show "add a c = add c a" by (rule add_c) +next show "add a (add c d) = add (add a c) d" using add_a by simp +next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr) +next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp +next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp +next show "mul x x = pwr x 2" by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) +next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul) +next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr) +next show "pwr x 0 = r1" using pwr_0 . +next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number' pwr_Suc pwr_0 mul_1 mul_c) +next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp +next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp +next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number' mul_pwr) +next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))" + by (simp add: nat_number' pwr_Suc mul_pwr) +qed + + +lemmas normalizing_semiring_axioms' = + normalizing_semiring_axioms [normalizer + semiring ops: semiring_ops + semiring rules: semiring_rules] + +end + +sublocale comm_semiring_1 + < normalizing!: normalizing_semiring plus times power zero one +proof +qed (simp_all add: algebra_simps) + +declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_semiring_axioms'} *} + +locale normalizing_ring = normalizing_semiring + + fixes sub :: "'a \ 'a \ 'a" + and neg :: "'a \ 'a" + assumes neg_mul: "neg x = mul (neg r1) x" + and sub_add: "sub x y = add x (neg y)" +begin + +lemma ring_ops: shows "TERM (sub x y)" and "TERM (neg x)" . + +lemmas ring_rules = neg_mul sub_add + +lemmas normalizing_ring_axioms' = + normalizing_ring_axioms [normalizer + semiring ops: semiring_ops + semiring rules: semiring_rules + ring ops: ring_ops + ring rules: ring_rules] + +end + +sublocale comm_ring_1 + < normalizing!: normalizing_ring plus times power zero one minus uminus +proof +qed (simp_all add: diff_minus) + +declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_axioms'} *} + +locale normalizing_field = normalizing_ring + + fixes divide :: "'a \ 'a \ 'a" + and inverse:: "'a \ 'a" + assumes divide_inverse: "divide x y = mul x (inverse y)" + and inverse_divide: "inverse x = divide r1 x" +begin + +lemma field_ops: shows "TERM (divide x y)" and "TERM (inverse x)" . + +lemmas field_rules = divide_inverse inverse_divide + +lemmas normalizing_field_axioms' = + normalizing_field_axioms [normalizer + semiring ops: semiring_ops + semiring rules: semiring_rules + ring ops: ring_ops + ring rules: ring_rules + field ops: field_ops + field rules: field_rules] + +end + +locale normalizing_semiring_cancel = normalizing_semiring + + assumes add_cancel: "add (x::'a) y = add x z \ y = z" + and add_mul_solve: "add (mul w y) (mul x z) = + add (mul w z) (mul x y) \ w = x \ y = z" +begin + +lemma noteq_reduce: "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" +proof- + have "a \ b \ c \ d \ \ (a = b \ c = d)" by simp + also have "\ \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" + using add_mul_solve by blast + finally show "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" + by simp +qed + +lemma add_scale_eq_noteq: "\r \ r0 ; (a = b) \ ~(c = d)\ + \ add a (mul r c) \ add b (mul r d)" +proof(clarify) + assume nz: "r\ r0" and cnd: "c\d" + and eq: "add b (mul r c) = add b (mul r d)" + hence "mul r c = mul r d" using cnd add_cancel by simp + hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)" + using mul_0 add_cancel by simp + thus "False" using add_mul_solve nz cnd by simp +qed + +lemma add_r0_iff: " x = add x a \ a = r0" +proof- + have "a = r0 \ add x a = add x r0" by (simp add: add_cancel) + thus "x = add x a \ a = r0" by (auto simp add: add_c add_0) +qed + +declare normalizing_semiring_axioms' [normalizer del] + +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 normalizing_ring_cancel = normalizing_semiring_cancel + normalizing_ring + + assumes subr0_iff: "sub x y = r0 \ x = y" +begin + +declare normalizing_ring_axioms' [normalizer del] + +lemmas normalizing_ring_cancel_axioms' = normalizing_ring_cancel_axioms [normalizer + semiring ops: semiring_ops + semiring rules: semiring_rules + ring ops: ring_ops + ring rules: ring_rules + idom rules: noteq_reduce add_scale_eq_noteq + ideal rules: subr0_iff add_r0_iff] + +end + +sublocale idom + < normalizing!: normalizing_ring_cancel plus times power zero one minus uminus +proof + fix w x y z + show "w * y + x * z = w * z + x * y \ w = x \ y = z" + proof + assume "w * y + x * z = w * z + x * y" + then have "w * y + x * z - w * z - x * y = 0" by (simp add: algebra_simps) + then have "w * (y - z) - x * (y - z) = 0" by (simp add: algebra_simps) + then have "(y - z) * (w - x) = 0" by (simp add: algebra_simps) + then have "y - z = 0 \ w - x = 0" by (rule divisors_zero) + then show "w = x \ y = z" by auto + qed (auto simp add: add_ac) +qed (simp_all add: algebra_simps) + +declaration {* Normalizer.semiring_funs @{thm normalizing.normalizing_ring_cancel_axioms'} *} + +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" + { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \ z" + hence "y < z \ y > z" by arith + moreover { + assume lt:"y k. z = y + k \ k > 0" by (rule_tac x="z - y" in exI, auto) + then obtain k where kp: "k>0" and yz:"z = y + k" by blast + from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps) + hence "x*k = w*k" by simp + hence "w = x" using kp by simp } + moreover { + assume lt: "y >z" hence "\k. y = z + k \ k>0" by (rule_tac x="y - z" in exI, auto) + then obtain k where kp: "k>0" and yz:"y = z + k" by blast + from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps) + hence "w*k = x*k" by simp + hence "w = x" using kp by simp } + ultimately have "w=x" by blast } + thus "(w * y + x * z = w * z + x * y) = (w = x \ y = z)" by auto +qed + +declaration {* Normalizer.semiring_funs @{thm normalizing_nat.normalizing_semiring_cancel_axioms'} *} + +locale normalizing_field_cancel = normalizing_ring_cancel + normalizing_field +begin + +declare normalizing_field_axioms' [normalizer del] + +lemmas normalizing_field_cancel_axioms' = normalizing_field_cancel_axioms [normalizer + semiring ops: semiring_ops + semiring rules: semiring_rules + ring ops: ring_ops + ring rules: ring_rules + field ops: field_ops + field rules: field_rules + idom rules: noteq_reduce add_scale_eq_noteq + ideal rules: subr0_iff add_r0_iff] + +end + +sublocale field + < normalizing!: normalizing_field_cancel plus times power zero one minus uminus divide inverse +proof +qed (simp_all add: divide_inverse) + +declaration {* Normalizer.field_funs @{thm normalizing.normalizing_field_cancel_axioms'} *} + +end diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Tools/Groebner_Basis/normalizer.ML --- a/src/HOL/Tools/Groebner_Basis/normalizer.ML Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML Fri May 07 15:05:52 2010 +0200 @@ -31,7 +31,6 @@ val semiring_normalizers_ord_wrapper: Proof.context -> entry -> (cterm -> cterm -> bool) -> {add: conv, mul: conv, neg: conv, main: conv, pow: conv, sub: conv} - val field_comp_conv: conv val setup: theory -> theory end @@ -41,156 +40,6 @@ (** some conversion **) -local - val zr = @{cpat "0"} - val zT = ctyp_of_term zr - val geq = @{cpat "op ="} - val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd - val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"} - val add_frac_num = mk_meta_eq @{thm "add_frac_num"} - val add_num_frac = mk_meta_eq @{thm "add_num_frac"} - - fun prove_nz ss T t = - let - val z = instantiate_cterm ([(zT,T)],[]) zr - val eq = instantiate_cterm ([(eqT,T)],[]) geq - val th = Simplifier.rewrite (ss addsimps @{thms simp_thms}) - (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} - (Thm.capply (Thm.capply eq t) z))) - in equal_elim (symmetric th) TrueI - end - - fun proc phi ss ct = - let - val ((x,y),(w,z)) = - (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct - val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] - val T = ctyp_of_term x - val [y_nz, z_nz] = map (prove_nz ss T) [y, z] - val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq - in SOME (implies_elim (implies_elim th y_nz) z_nz) - end - handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE - - fun proc2 phi ss ct = - let - val (l,r) = Thm.dest_binop ct - val T = ctyp_of_term l - in (case (term_of l, term_of r) of - (Const(@{const_name Rings.divide},_)$_$_, _) => - let val (x,y) = Thm.dest_binop l val z = r - val _ = map (HOLogic.dest_number o term_of) [x,y,z] - val ynz = prove_nz ss T y - in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) - end - | (_, Const (@{const_name Rings.divide},_)$_$_) => - let val (x,y) = Thm.dest_binop r val z = l - val _ = map (HOLogic.dest_number o term_of) [x,y,z] - val ynz = prove_nz ss T y - in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) - end - | _ => NONE) - end - handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE - - fun is_number (Const(@{const_name Rings.divide},_)$a$b) = is_number a andalso is_number b - | is_number t = can HOLogic.dest_number t - - val is_number = is_number o term_of - - fun proc3 phi ss ct = - (case term_of ct of - Const(@{const_name Orderings.less},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => - let - val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"} - in SOME (mk_meta_eq th) end - | Const(@{const_name Orderings.less_eq},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => - let - val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"} - in SOME (mk_meta_eq th) end - | Const("op =",_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => - let - val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"} - in SOME (mk_meta_eq th) end - | Const(@{const_name Orderings.less},_)$_$(Const(@{const_name Rings.divide},_)$_$_) => - let - val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"} - in SOME (mk_meta_eq th) end - | Const(@{const_name Orderings.less_eq},_)$_$(Const(@{const_name Rings.divide},_)$_$_) => - let - val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"} - in SOME (mk_meta_eq th) end - | Const("op =",_)$_$(Const(@{const_name Rings.divide},_)$_$_) => - let - val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop - val _ = map is_number [a,b,c] - val T = ctyp_of_term c - val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"} - in SOME (mk_meta_eq th) end - | _ => NONE) - handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE - -val add_frac_frac_simproc = - make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}], - name = "add_frac_frac_simproc", - proc = proc, identifier = []} - -val add_frac_num_simproc = - make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}], - name = "add_frac_num_simproc", - proc = proc2, identifier = []} - -val ord_frac_simproc = - make_simproc - {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"}, - @{cpat "(?a::(?'a::{field, ord}))/?b <= ?c"}, - @{cpat "?c < (?a::(?'a::{field, ord}))/?b"}, - @{cpat "?c <= (?a::(?'a::{field, ord}))/?b"}, - @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"}, - @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}], - name = "ord_frac_simproc", proc = proc3, identifier = []} - -val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, - @{thm "divide_Numeral1"}, - @{thm "divide_zero"}, @{thm "divide_Numeral0"}, - @{thm "divide_divide_eq_left"}, - @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"}, - @{thm "times_divide_times_eq"}, - @{thm "divide_divide_eq_right"}, - @{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}, - Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute})))) - (@{thm field_divide_inverse} RS sym)] - -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, - ord_frac_simproc] - 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 (** data **) @@ -365,7 +214,7 @@ {is_const = K numeral_is_const, dest_const = K dest_const, mk_const = mk_const, - conv = K (K field_comp_conv)} + conv = K (K Numeral_Simprocs.field_comp_conv)} end; diff -r 912080b2c449 -r 7f1da69cacb3 src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Fri May 07 10:00:24 2010 +0200 +++ b/src/HOL/Tools/numeral_simprocs.ML Fri May 07 15:05:52 2010 +0200 @@ -1,7 +1,7 @@ (* Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 2000 University of Cambridge -Simprocs for the integer numerals. +Simprocs for the (integer) numerals. *) (*To quote from Provers/Arith/cancel_numeral_factor.ML: @@ -24,6 +24,7 @@ val field_combine_numerals: simproc val field_cancel_numeral_factors: simproc list val num_ss: simpset + val field_comp_conv: conv end; structure Numeral_Simprocs : NUMERAL_SIMPROCS = @@ -602,6 +603,157 @@ "(l::'a::field_inverse_zero) / (m * n)"], K DivideCancelFactor.proc)]; +local + val zr = @{cpat "0"} + val zT = ctyp_of_term zr + val geq = @{cpat "op ="} + val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd + val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"} + val add_frac_num = mk_meta_eq @{thm "add_frac_num"} + val add_num_frac = mk_meta_eq @{thm "add_num_frac"} + + fun prove_nz ss T t = + let + val z = instantiate_cterm ([(zT,T)],[]) zr + val eq = instantiate_cterm ([(eqT,T)],[]) geq + val th = Simplifier.rewrite (ss addsimps @{thms simp_thms}) + (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} + (Thm.capply (Thm.capply eq t) z))) + in equal_elim (symmetric th) TrueI + end + + fun proc phi ss ct = + let + val ((x,y),(w,z)) = + (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct + val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] + val T = ctyp_of_term x + val [y_nz, z_nz] = map (prove_nz ss T) [y, z] + val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq + in SOME (implies_elim (implies_elim th y_nz) z_nz) + end + handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE + + fun proc2 phi ss ct = + let + val (l,r) = Thm.dest_binop ct + val T = ctyp_of_term l + in (case (term_of l, term_of r) of + (Const(@{const_name Rings.divide},_)$_$_, _) => + let val (x,y) = Thm.dest_binop l val z = r + val _ = map (HOLogic.dest_number o term_of) [x,y,z] + val ynz = prove_nz ss T y + in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) + end + | (_, Const (@{const_name Rings.divide},_)$_$_) => + let val (x,y) = Thm.dest_binop r val z = l + val _ = map (HOLogic.dest_number o term_of) [x,y,z] + val ynz = prove_nz ss T y + in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) + end + | _ => NONE) + end + handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE + + fun is_number (Const(@{const_name Rings.divide},_)$a$b) = is_number a andalso is_number b + | is_number t = can HOLogic.dest_number t + + val is_number = is_number o term_of + + fun proc3 phi ss ct = + (case term_of ct of + Const(@{const_name Orderings.less},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => + let + val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"} + in SOME (mk_meta_eq th) end + | Const(@{const_name Orderings.less_eq},_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => + let + val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"} + in SOME (mk_meta_eq th) end + | Const("op =",_)$(Const(@{const_name Rings.divide},_)$_$_)$_ => + let + val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"} + in SOME (mk_meta_eq th) end + | Const(@{const_name Orderings.less},_)$_$(Const(@{const_name Rings.divide},_)$_$_) => + let + val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"} + in SOME (mk_meta_eq th) end + | Const(@{const_name Orderings.less_eq},_)$_$(Const(@{const_name Rings.divide},_)$_$_) => + let + val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"} + in SOME (mk_meta_eq th) end + | Const("op =",_)$_$(Const(@{const_name Rings.divide},_)$_$_) => + let + val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop + val _ = map is_number [a,b,c] + val T = ctyp_of_term c + val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"} + in SOME (mk_meta_eq th) end + | _ => NONE) + handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE + +val add_frac_frac_simproc = + make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}], + name = "add_frac_frac_simproc", + proc = proc, identifier = []} + +val add_frac_num_simproc = + make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}], + name = "add_frac_num_simproc", + proc = proc2, identifier = []} + +val ord_frac_simproc = + make_simproc + {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"}, + @{cpat "(?a::(?'a::{field, ord}))/?b <= ?c"}, + @{cpat "?c < (?a::(?'a::{field, ord}))/?b"}, + @{cpat "?c <= (?a::(?'a::{field, ord}))/?b"}, + @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"}, + @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}], + name = "ord_frac_simproc", proc = proc3, identifier = []} + +val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, + @{thm "divide_Numeral1"}, + @{thm "divide_zero"}, @{thm "divide_Numeral0"}, + @{thm "divide_divide_eq_left"}, + @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"}, + @{thm "times_divide_times_eq"}, + @{thm "divide_divide_eq_right"}, + @{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}, + Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute})))) + (@{thm field_divide_inverse} RS sym)] + +in + +val field_comp_conv = (Simplifier.rewrite +(HOL_basic_ss addsimps @{thms "semiring_norm"} + addsimps ths addsimps @{thms simp_thms} + addsimprocs field_cancel_numeral_factors + addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, + ord_frac_simproc] + 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 + end; Addsimprocs Numeral_Simprocs.cancel_numerals;