# HG changeset patch # User wenzelm # Date 1127218229 -7200 # Node ID 45164074dad4419c6ac0c62223f574cee524a957 # Parent 830bc15e692cf2e7f7be9bcc4281af823bd5c8f7 added Commutative_Ring (from Main HOL); diff -r 830bc15e692c -r 45164074dad4 src/HOL/Library/Commutative_Ring.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Commutative_Ring.thy Tue Sep 20 14:10:29 2005 +0200 @@ -0,0 +1,327 @@ +(* ID: $Id$ + Author: Bernhard Haeupler + +Proving equalities in commutative rings done "right" in Isabelle/HOL. +*) + +header {* Proving equalities in commutative rings *} + +theory Commutative_Ring +imports Main +uses ("comm_ring.ML") +begin + +text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *} + +datatype 'a pol = + Pc 'a + | Pinj nat "'a pol" + | PX "'a pol" nat "'a pol" + +datatype 'a polex = + Pol "'a pol" + | Add "'a polex" "'a polex" + | Sub "'a polex" "'a polex" + | Mul "'a polex" "'a polex" + | Pow "'a polex" nat + | Neg "'a polex" + +text {* Interpretation functions for the shadow syntax. *} + +consts + Ipol :: "'a::{comm_ring,recpower} list \ 'a pol \ 'a" + Ipolex :: "'a::{comm_ring,recpower} list \ 'a polex \ 'a" + +primrec + "Ipol l (Pc c) = c" + "Ipol l (Pinj i P) = Ipol (drop i l) P" + "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q" + +primrec + "Ipolex l (Pol P) = Ipol l P" + "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q" + "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q" + "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q" + "Ipolex l (Pow p n) = Ipolex l p ^ n" + "Ipolex l (Neg P) = - Ipolex l P" + +text {* Create polynomial normalized polynomials given normalized inputs. *} + +constdefs + mkPinj :: "nat \ 'a pol \ 'a pol" + "mkPinj x P \ (case P of + Pc c \ Pc c | + Pinj y P \ Pinj (x + y) P | + PX p1 y p2 \ Pinj x P)" + +constdefs + mkPX :: "'a::{comm_ring,recpower} pol \ nat \ 'a pol \ 'a pol" + "mkPX P i Q == (case P of + Pc c \ (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | + Pinj j R \ PX P i Q | + PX P2 i2 Q2 \ (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" + +text {* Defining the basic ring operations on normalized polynomials *} + +consts + add :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" + mul :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" + neg :: "'a::{comm_ring,recpower} pol \ 'a pol" + sqr :: "'a::{comm_ring,recpower} pol \ 'a pol" + pow :: "'a::{comm_ring,recpower} pol \ nat \ 'a pol" + +text {* Addition *} +recdef add "measure (\(x, y). size x + size y)" + "add (Pc a, Pc b) = Pc (a + b)" + "add (Pc c, Pinj i P) = Pinj i (add (P, Pc c))" + "add (Pinj i P, Pc c) = Pinj i (add (P, Pc c))" + "add (Pc c, PX P i Q) = PX P i (add (Q, Pc c))" + "add (PX P i Q, Pc c) = PX P i (add (Q, Pc c))" + "add (Pinj x P, Pinj y Q) = + (if x=y then mkPinj x (add (P, Q)) + else (if x>y then mkPinj y (add (Pinj (x-y) P, Q)) + else mkPinj x (add (Pinj (y-x) Q, P)) ))" + "add (Pinj x P, PX Q y R) = + (if x=0 then add(P, PX Q y R) + else (if x=1 then PX Q y (add (R, P)) + else PX Q y (add (R, Pinj (x - 1) P))))" + "add (PX P x R, Pinj y Q) = + (if y=0 then add(PX P x R, Q) + else (if y=1 then PX P x (add (R, Q)) + else PX P x (add (R, Pinj (y - 1) Q))))" + "add (PX P1 x P2, PX Q1 y Q2) = + (if x=y then mkPX (add (P1, Q1)) x (add (P2, Q2)) + else (if x>y then mkPX (add (PX P1 (x-y) (Pc 0), Q1)) y (add (P2,Q2)) + else mkPX (add (PX Q1 (y-x) (Pc 0), P1)) x (add (P2,Q2)) ))" + +text {* Multiplication *} +recdef mul "measure (\(x, y). size x + size y)" + "mul (Pc a, Pc b) = Pc (a*b)" + "mul (Pc c, Pinj i P) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))" + "mul (Pinj i P, Pc c) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))" + "mul (Pc c, PX P i Q) = + (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))" + "mul (PX P i Q, Pc c) = + (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))" + "mul (Pinj x P, Pinj y Q) = + (if x=y then mkPinj x (mul (P, Q)) + else (if x>y then mkPinj y (mul (Pinj (x-y) P, Q)) + else mkPinj x (mul (Pinj (y-x) Q, P)) ))" + "mul (Pinj x P, PX Q y R) = + (if x=0 then mul(P, PX Q y R) + else (if x=1 then mkPX (mul (Pinj x P, Q)) y (mul (R, P)) + else mkPX (mul (Pinj x P, Q)) y (mul (R, Pinj (x - 1) P))))" + "mul (PX P x R, Pinj y Q) = + (if y=0 then mul(PX P x R, Q) + else (if y=1 then mkPX (mul (Pinj y Q, P)) x (mul (R, Q)) + else mkPX (mul (Pinj y Q, P)) x (mul (R, Pinj (y - 1) Q))))" + "mul (PX P1 x P2, PX Q1 y Q2) = + add (mkPX (mul (P1, Q1)) (x+y) (mul (P2, Q2)), + add (mkPX (mul (P1, mkPinj 1 Q2)) x (Pc 0), mkPX (mul (Q1, mkPinj 1 P2)) y (Pc 0)) )" +(hints simp add: mkPinj_def split: pol.split) + +text {* Negation*} +primrec + "neg (Pc c) = Pc (-c)" + "neg (Pinj i P) = Pinj i (neg P)" + "neg (PX P x Q) = PX (neg P) x (neg Q)" + +text {* Substraction *} +constdefs + sub :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" + "sub p q \ add (p, neg q)" + +text {* Square for Fast Exponentation *} +primrec + "sqr (Pc c) = Pc (c * c)" + "sqr (Pinj i P) = mkPinj i (sqr P)" + "sqr (PX A x B) = add (mkPX (sqr A) (x + x) (sqr B), + mkPX (mul (mul (Pc (1 + 1), A), mkPinj 1 B)) x (Pc 0))" + +text {* Fast Exponentation *} +lemma pow_wf:"odd n \ (n::nat) div 2 < n" by (cases n) auto +recdef pow "measure (\(x, y). y)" + "pow (p, 0) = Pc 1" + "pow (p, n) = (if even n then (pow (sqr p, n div 2)) else mul (p, pow (sqr p, n div 2)))" +(hints simp add: pow_wf) + +lemma pow_if: + "pow (p,n) = + (if n = 0 then Pc 1 else if even n then pow (sqr p, n div 2) + else mul (p, pow (sqr p, n div 2)))" + by (cases n) simp_all + +(* +lemma number_of_nat_B0: "(number_of (w BIT bit.B0) ::nat) = 2* (number_of w)" +by simp + +lemma number_of_nat_even: "even (number_of (w BIT bit.B0)::nat)" +by simp + +lemma pow_even : "pow (p, number_of(w BIT bit.B0)) = pow (sqr p, number_of w)" + ( is "pow(?p,?n) = pow (_,?n2)") +proof- + have "even ?n" by simp + hence "pow (p, ?n) = pow (sqr p, ?n div 2)" + apply simp + apply (cases "IntDef.neg (number_of w)") + apply simp + done +*) + +text {* Normalization of polynomial expressions *} + +consts norm :: "'a::{comm_ring,recpower} polex \ 'a pol" +primrec + "norm (Pol P) = P" + "norm (Add P Q) = add (norm P, norm Q)" + "norm (Sub p q) = sub (norm p) (norm q)" + "norm (Mul P Q) = mul (norm P, norm Q)" + "norm (Pow p n) = pow (norm p, n)" + "norm (Neg P) = neg (norm P)" + +text {* mkPinj preserve semantics *} +lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" + by (induct B) (auto simp add: mkPinj_def ring_eq_simps) + +text {* mkPX preserves semantics *} +lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)" + by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add ring_eq_simps) + +text {* Correctness theorems for the implemented operations *} + +text {* Negation *} +lemma neg_ci: "\l. Ipol l (neg P) = -(Ipol l P)" + by (induct P) auto + +text {* Addition *} +lemma add_ci: "\l. Ipol l (add (P, Q)) = Ipol l P + Ipol l Q" +proof (induct P Q rule: add.induct) + case (6 x P y Q) + show ?case + proof (rule linorder_cases) + assume "x < y" + with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps) + next + assume "x = y" + with 6 show ?case by (simp add: mkPinj_ci) + next + assume "x > y" + with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps) + qed +next + case (7 x P Q y R) + have "x = 0 \ x = 1 \ x > 1" by arith + moreover + { assume "x = 0" with 7 have ?case by simp } + moreover + { assume "x = 1" with 7 have ?case by (simp add: ring_eq_simps) } + moreover + { assume "x > 1" from 7 have ?case by (cases x) simp_all } + ultimately show ?case by blast +next + case (8 P x R y Q) + have "y = 0 \ y = 1 \ y > 1" by arith + moreover + { assume "y = 0" with 8 have ?case by simp } + moreover + { assume "y = 1" with 8 have ?case by simp } + moreover + { assume "y > 1" with 8 have ?case by simp } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + show ?case + proof (rule linorder_cases) + assume a: "x < y" hence "EX d. d + x = y" by arith + with 9 a show ?case by (auto simp add: mkPX_ci power_add ring_eq_simps) + next + assume a: "y < x" hence "EX d. d + y = x" by arith + with 9 a show ?case by (auto simp add: power_add mkPX_ci ring_eq_simps) + next + assume "x = y" + with 9 show ?case by (simp add: mkPX_ci ring_eq_simps) + qed +qed (auto simp add: ring_eq_simps) + +text {* Multiplication *} +lemma mul_ci: "\l. Ipol l (mul (P, Q)) = Ipol l P * Ipol l Q" + by (induct P Q rule: mul.induct) + (simp_all add: mkPX_ci mkPinj_ci ring_eq_simps add_ci power_add) + +text {* Substraction *} +lemma sub_ci: "Ipol l (sub p q) = Ipol l p - Ipol l q" + by (simp add: add_ci neg_ci sub_def) + +text {* Square *} +lemma sqr_ci:"\ls. Ipol ls (sqr p) = Ipol ls p * Ipol ls p" + by (induct p) (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci ring_eq_simps power_add) + + +text {* Power *} +lemma even_pow:"even n \ pow (p, n) = pow (sqr p, n div 2)" by (induct n) simp_all + +lemma pow_ci: "\p. Ipol ls (pow (p, n)) = (Ipol ls p) ^ n" +proof (induct n rule: nat_less_induct) + case (1 k) + have two:"2 = Suc (Suc 0)" by simp + show ?case + proof (cases k) + case (Suc l) + show ?thesis + proof cases + assume EL: "even l" + have "Suc l div 2 = l div 2" + by (simp add: nat_number even_nat_plus_one_div_two [OF EL]) + moreover + from Suc have "l < k" by simp + with 1 have "\p. Ipol ls (pow (p, l)) = Ipol ls p ^ l" by simp + moreover + note Suc EL even_nat_plus_one_div_two [OF EL] + ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) + next + assume OL: "odd l" + with prems have "\\mp. Ipol ls (pow (p, m)) = Ipol ls p ^ m; k = Suc l; odd l\ \ \p. Ipol ls (sqr p) ^ (Suc l div 2) = Ipol ls p ^ Suc l" + proof(cases l) + case (Suc w) + from prems have EW: "even w" by simp + from two have two_times:"(2 * (w div 2))= w" + by (simp only: even_nat_div_two_times_two[OF EW]) + have A: "\p. (Ipol ls p * Ipol ls p) = (Ipol ls p) ^ (Suc (Suc 0))" + by (simp add: power_Suc) + from A two [symmetric] have "ALL p.(Ipol ls p * Ipol ls p) = (Ipol ls p) ^ 2" + by simp + with prems show ?thesis + by (auto simp add: power_mult[symmetric, of _ 2 _] two_times mul_ci sqr_ci) + qed simp + with prems show ?thesis by simp + qed + next + case 0 + then show ?thesis by simp + qed +qed + +text {* Normalization preserves semantics *} +lemma norm_ci:"Ipolex l Pe = Ipol l (norm Pe)" + by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) + +text {* Reflection lemma: Key to the (incomplete) decision procedure *} +lemma norm_eq: + assumes eq: "norm P1 = norm P2" + shows "Ipolex l P1 = Ipolex l P2" +proof - + from eq have "Ipol l (norm P1) = Ipol l (norm P2)" by simp + thus ?thesis by (simp only: norm_ci) +qed + + +text {* Code generation *} +(* +Does not work, since no generic ring operations implementation is there +generate_code ("ring.ML") test = "norm"*) + +use "comm_ring.ML" +setup "CommRing.setup" + +end diff -r 830bc15e692c -r 45164074dad4 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue Sep 20 14:06:00 2005 +0200 +++ b/src/HOL/Library/Library.thy Tue Sep 20 14:10:29 2005 +0200 @@ -18,6 +18,7 @@ Word Zorn Char_ord + Commutative_Ring begin end (*>*) diff -r 830bc15e692c -r 45164074dad4 src/HOL/Library/comm_ring.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/comm_ring.ML Tue Sep 20 14:10:29 2005 +0200 @@ -0,0 +1,142 @@ +(* ID: $Id$ + Author: Amine Chaieb + +Tactic for solving equalities over commutative rings. +*) + +signature COMM_RING = +sig + val comm_ring_tac : int -> tactic + val comm_ring_method: int -> Proof.method + val algebra_method: int -> Proof.method + val setup : (theory -> theory) list +end + +structure CommRing: COMM_RING = +struct + +(* The Cring exception for erronous uses of cring_tac *) +exception CRing of string; + +(* Zero and One of the commutative ring *) +fun cring_zero T = Const("0",T); +fun cring_one T = Const("1",T); + +(* reification functions *) +(* add two polynom expressions *) +fun polT t = Type ("Commutative_Ring.pol",[t]); +fun polexT t = Type("Commutative_Ring.polex",[t]); +val nT = HOLogic.natT; +fun listT T = Type ("List.list",[T]); + +(* Reification of the constructors *) +(* Nat*) +val succ = Const("Suc",nT --> nT); +val zero = Const("0",nT); +val one = Const("1",nT); + +(* Lists *) +fun reif_list T [] = Const("List.list.Nil",listT T) + | reif_list T (x::xs) = Const("List.list.Cons",[T,listT T] ---> listT T) + $x$(reif_list T xs); + +(* pol*) +fun pol_Pc t = Const("Commutative_Ring.pol.Pc",t --> polT t); +fun pol_Pinj t = Const("Commutative_Ring.pol.Pinj",[nT,polT t] ---> polT t); +fun pol_PX t = Const("Commutative_Ring.pol.PX",[polT t, nT, polT t] ---> polT t); + +(* polex *) +fun polex_add t = Const("Commutative_Ring.polex.Add",[polexT t,polexT t] ---> polexT t); +fun polex_sub t = Const("Commutative_Ring.polex.Sub",[polexT t,polexT t] ---> polexT t); +fun polex_mul t = Const("Commutative_Ring.polex.Mul",[polexT t,polexT t] ---> polexT t); +fun polex_neg t = Const("Commutative_Ring.polex.Neg",polexT t --> polexT t); +fun polex_pol t = Const("Commutative_Ring.polex.Pol",polT t --> polexT t); +fun polex_pow t = Const("Commutative_Ring.polex.Pow",[polexT t, nT] ---> polexT t); +(* reification of natural numbers *) +fun reif_nat n = + if n>0 then succ$(reif_nat (n-1)) + else if n=0 then zero + else raise CRing "ring_tac: reif_nat negative n"; + +(* reification of polynoms : primitive cring expressions *) +fun reif_pol T vs t = + case t of + Free(_,_) => + let val i = find_index_eq t vs + in if i = 0 + then (pol_PX T)$((pol_Pc T)$ (cring_one T)) + $one$((pol_Pc T)$(cring_zero T)) + else (pol_Pinj T)$(reif_nat i)$ + ((pol_PX T)$((pol_Pc T)$ (cring_one T)) + $one$ + ((pol_Pc T)$(cring_zero T))) + end + | _ => (pol_Pc T)$ t; + + +(* reification of polynom expressions *) +fun reif_polex T vs t = + case t of + Const("op +",_)$a$b => (polex_add T) + $ (reif_polex T vs a)$(reif_polex T vs b) + | Const("op -",_)$a$b => (polex_sub T) + $ (reif_polex T vs a)$(reif_polex T vs b) + | Const("op *",_)$a$b => (polex_mul T) + $ (reif_polex T vs a)$ (reif_polex T vs b) + | Const("uminus",_)$a => (polex_neg T) + $ (reif_polex T vs a) + | (Const("Nat.power",_)$a$n) => (polex_pow T) $ (reif_polex T vs a) $ n + + | _ => (polex_pol T) $ (reif_pol T vs t); + +(* reification of the equation *) +val cr_sort = Sign.read_sort (the_context ()) "{comm_ring,recpower}"; +fun reif_eq sg (eq as Const("op =",Type("fun",a::_))$lhs$rhs) = + if Sign.of_sort (the_context()) (a,cr_sort) + then + let val fs = term_frees eq + val cvs = cterm_of sg (reif_list a fs) + val clhs = cterm_of sg (reif_polex a fs lhs) + val crhs = cterm_of sg (reif_polex a fs rhs) + val ca = ctyp_of sg a + in (ca,cvs,clhs, crhs) + end + else raise CRing "reif_eq: not an equation over comm_ring + recpower" + | reif_eq sg _ = raise CRing "reif_eq: not an equation"; + +(*The cring tactic *) +(* Attention: You have to make sure that no t^0 is in the goal!! *) +(* Use simply rewriting t^0 = 1 *) +fun cring_ss sg = simpset_of sg + addsimps + (map thm ["mkPX_def", "mkPinj_def","sub_def", + "power_add","even_def","pow_if"]) + addsimps [sym OF [thm "power_add"]]; + +val norm_eq = thm "norm_eq" +fun comm_ring_tac i =(fn st => + let + val g = List.nth (prems_of st, i - 1) + val sg = sign_of_thm st + val (ca,cvs,clhs,crhs) = reif_eq sg (HOLogic.dest_Trueprop g) + val norm_eq_th = simplify (cring_ss sg) + (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] + norm_eq) + in ((cut_rules_tac [norm_eq_th] i) + THEN (simp_tac (cring_ss sg) i) + THEN (simp_tac (cring_ss sg) i)) st + end); + +fun comm_ring_method i = Method.METHOD (fn facts => + Method.insert_tac facts 1 THEN comm_ring_tac i); +val algebra_method = comm_ring_method; + +val setup = + [Method.add_method ("comm_ring", + Method.no_args (comm_ring_method 1), + "reflective decision procedure for equalities over commutative rings"), + Method.add_method ("algebra", + Method.no_args (algebra_method 1), + "Method for proving algebraic properties: for now only comm_ring")]; + +end;