wenzelm@17516: (* ID: $Id$ wenzelm@17516: Author: Bernhard Haeupler wenzelm@17516: wenzelm@17516: Proving equalities in commutative rings done "right" in Isabelle/HOL. wenzelm@17516: *) wenzelm@17516: wenzelm@17516: header {* Proving equalities in commutative rings *} wenzelm@17516: wenzelm@17516: theory Commutative_Ring wenzelm@21256: imports Main Parity wenzelm@17516: uses ("comm_ring.ML") wenzelm@17516: begin wenzelm@17516: wenzelm@17516: text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *} wenzelm@17516: wenzelm@17516: datatype 'a pol = wenzelm@17516: Pc 'a wenzelm@17516: | Pinj nat "'a pol" wenzelm@17516: | PX "'a pol" nat "'a pol" wenzelm@17516: wenzelm@17516: datatype 'a polex = wenzelm@20622: Pol "'a pol" wenzelm@17516: | Add "'a polex" "'a polex" wenzelm@17516: | Sub "'a polex" "'a polex" wenzelm@17516: | Mul "'a polex" "'a polex" wenzelm@17516: | Pow "'a polex" nat wenzelm@17516: | Neg "'a polex" wenzelm@17516: wenzelm@17516: text {* Interpretation functions for the shadow syntax. *} wenzelm@17516: haftmann@22742: fun wenzelm@17516: Ipol :: "'a::{comm_ring,recpower} list \ 'a pol \ 'a" haftmann@22742: where haftmann@22742: "Ipol l (Pc c) = c" haftmann@22742: | "Ipol l (Pinj i P) = Ipol (drop i l) P" haftmann@22742: | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q" wenzelm@17516: haftmann@22742: fun haftmann@22742: Ipolex :: "'a::{comm_ring,recpower} list \ 'a polex \ 'a" haftmann@22742: where haftmann@22742: "Ipolex l (Pol P) = Ipol l P" haftmann@22742: | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q" haftmann@22742: | "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q" haftmann@22742: | "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q" haftmann@22742: | "Ipolex l (Pow p n) = Ipolex l p ^ n" haftmann@22742: | "Ipolex l (Neg P) = - Ipolex l P" wenzelm@17516: wenzelm@17516: text {* Create polynomial normalized polynomials given normalized inputs. *} wenzelm@17516: wenzelm@19736: definition wenzelm@21404: mkPinj :: "nat \ 'a pol \ 'a pol" where wenzelm@19736: "mkPinj x P = (case P of wenzelm@17516: Pc c \ Pc c | wenzelm@17516: Pinj y P \ Pinj (x + y) P | wenzelm@17516: PX p1 y p2 \ Pinj x P)" wenzelm@17516: wenzelm@19736: definition wenzelm@21404: mkPX :: "'a::{comm_ring,recpower} pol \ nat \ 'a pol \ 'a pol" where wenzelm@19736: "mkPX P i Q = (case P of wenzelm@17516: Pc c \ (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | wenzelm@17516: Pinj j R \ PX P i Q | wenzelm@17516: PX P2 i2 Q2 \ (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" wenzelm@17516: wenzelm@17516: text {* Defining the basic ring operations on normalized polynomials *} wenzelm@17516: haftmann@22742: function haftmann@22742: add :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" (infixl "\" 65) haftmann@22742: where haftmann@22742: "Pc a \ Pc b = Pc (a + b)" haftmann@22742: | "Pc c \ Pinj i P = Pinj i (P \ Pc c)" haftmann@22742: | "Pinj i P \ Pc c = Pinj i (P \ Pc c)" haftmann@22742: | "Pc c \ PX P i Q = PX P i (Q \ Pc c)" haftmann@22742: | "PX P i Q \ Pc c = PX P i (Q \ Pc c)" haftmann@22742: | "Pinj x P \ Pinj y Q = haftmann@22742: (if x = y then mkPinj x (P \ Q) haftmann@22742: else (if x > y then mkPinj y (Pinj (x - y) P \ Q) haftmann@22742: else mkPinj x (Pinj (y - x) Q \ P)))" haftmann@22742: | "Pinj x P \ PX Q y R = haftmann@22742: (if x = 0 then P \ PX Q y R haftmann@22742: else (if x = 1 then PX Q y (R \ P) haftmann@22742: else PX Q y (R \ Pinj (x - 1) P)))" haftmann@22742: | "PX P x R \ Pinj y Q = haftmann@22742: (if y = 0 then PX P x R \ Q haftmann@22742: else (if y = 1 then PX P x (R \ Q) haftmann@22742: else PX P x (R \ Pinj (y - 1) Q)))" haftmann@22742: | "PX P1 x P2 \ PX Q1 y Q2 = haftmann@22742: (if x = y then mkPX (P1 \ Q1) x (P2 \ Q2) haftmann@22742: else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \ Q1) y (P2 \ Q2) haftmann@22742: else mkPX (PX Q1 (y-x) (Pc 0) \ P1) x (P2 \ Q2)))" haftmann@22742: by pat_completeness auto haftmann@22742: termination by (relation "measure (\(x, y). size x + size y)") auto wenzelm@17516: haftmann@22742: function haftmann@22742: mul :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" (infixl "\" 70) haftmann@22742: where haftmann@22742: "Pc a \ Pc b = Pc (a * b)" haftmann@22742: | "Pc c \ Pinj i P = haftmann@22742: (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" haftmann@22742: | "Pinj i P \ Pc c = haftmann@22742: (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" haftmann@22742: | "Pc c \ PX P i Q = haftmann@22742: (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" haftmann@22742: | "PX P i Q \ Pc c = haftmann@22742: (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" haftmann@22742: | "Pinj x P \ Pinj y Q = haftmann@22742: (if x = y then mkPinj x (P \ Q) else haftmann@22742: (if x > y then mkPinj y (Pinj (x-y) P \ Q) haftmann@22742: else mkPinj x (Pinj (y - x) Q \ P)))" haftmann@22742: | "Pinj x P \ PX Q y R = haftmann@22742: (if x = 0 then P \ PX Q y R else haftmann@22742: (if x = 1 then mkPX (Pinj x P \ Q) y (R \ P) haftmann@22742: else mkPX (Pinj x P \ Q) y (R \ Pinj (x - 1) P)))" haftmann@22742: | "PX P x R \ Pinj y Q = haftmann@22742: (if y = 0 then PX P x R \ Q else haftmann@22742: (if y = 1 then mkPX (Pinj y Q \ P) x (R \ Q) haftmann@22742: else mkPX (Pinj y Q \ P) x (R \ Pinj (y - 1) Q)))" haftmann@22742: | "PX P1 x P2 \ PX Q1 y Q2 = haftmann@22742: mkPX (P1 \ Q1) (x + y) (P2 \ Q2) \ haftmann@22742: (mkPX (P1 \ mkPinj 1 Q2) x (Pc 0) \ haftmann@22742: (mkPX (Q1 \ mkPinj 1 P2) y (Pc 0)))" haftmann@22742: by pat_completeness auto haftmann@22742: termination by (relation "measure (\(x, y). size x + size y)") haftmann@22742: (auto simp add: mkPinj_def split: pol.split) wenzelm@17516: wenzelm@17516: text {* Negation*} haftmann@22742: fun haftmann@22742: neg :: "'a::{comm_ring,recpower} pol \ 'a pol" haftmann@22742: where haftmann@22742: "neg (Pc c) = Pc (-c)" haftmann@22742: | "neg (Pinj i P) = Pinj i (neg P)" haftmann@22742: | "neg (PX P x Q) = PX (neg P) x (neg Q)" wenzelm@17516: wenzelm@17516: text {* Substraction *} wenzelm@19736: definition haftmann@22742: sub :: "'a::{comm_ring,recpower} pol \ 'a pol \ 'a pol" (infixl "\" 65) haftmann@22742: where haftmann@22742: "sub P Q = P \ neg Q" wenzelm@17516: wenzelm@17516: text {* Square for Fast Exponentation *} haftmann@22742: fun haftmann@22742: sqr :: "'a::{comm_ring,recpower} pol \ 'a pol" haftmann@22742: where haftmann@22742: "sqr (Pc c) = Pc (c * c)" haftmann@22742: | "sqr (Pinj i P) = mkPinj i (sqr P)" haftmann@22742: | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \ haftmann@22742: mkPX (Pc (1 + 1) \ A \ mkPinj 1 B) x (Pc 0)" wenzelm@17516: wenzelm@17516: text {* Fast Exponentation *} haftmann@22742: fun haftmann@22742: pow :: "nat \ 'a::{comm_ring,recpower} pol \ 'a pol" haftmann@22742: where haftmann@22742: "pow 0 P = Pc 1" haftmann@22742: | "pow n P = (if even n then pow (n div 2) (sqr P) haftmann@22742: else P \ pow (n div 2) (sqr P))" haftmann@22742: wenzelm@17516: lemma pow_if: haftmann@22742: "pow n P = haftmann@22742: (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) haftmann@22742: else P \ pow (n div 2) (sqr P))" wenzelm@17516: by (cases n) simp_all wenzelm@17516: wenzelm@17516: wenzelm@17516: text {* Normalization of polynomial expressions *} wenzelm@17516: haftmann@22742: fun haftmann@22742: norm :: "'a::{comm_ring,recpower} polex \ 'a pol" haftmann@22742: where haftmann@22742: "norm (Pol P) = P" haftmann@22742: | "norm (Add P Q) = norm P \ norm Q" haftmann@22742: | "norm (Sub P Q) = norm P \ norm Q" haftmann@22742: | "norm (Mul P Q) = norm P \ norm Q" haftmann@22742: | "norm (Pow P n) = pow n (norm P)" haftmann@22742: | "norm (Neg P) = neg (norm P)" wenzelm@17516: wenzelm@17516: text {* mkPinj preserve semantics *} wenzelm@17516: lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" wenzelm@17516: by (induct B) (auto simp add: mkPinj_def ring_eq_simps) wenzelm@17516: wenzelm@17516: text {* mkPX preserves semantics *} wenzelm@17516: lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)" wenzelm@17516: by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add ring_eq_simps) wenzelm@17516: wenzelm@17516: text {* Correctness theorems for the implemented operations *} wenzelm@17516: wenzelm@17516: text {* Negation *} wenzelm@20622: lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" wenzelm@20622: by (induct P arbitrary: l) auto wenzelm@17516: wenzelm@17516: text {* Addition *} haftmann@22742: lemma add_ci: "Ipol l (P \ Q) = Ipol l P + Ipol l Q" wenzelm@20622: proof (induct P Q arbitrary: l rule: add.induct) wenzelm@17516: case (6 x P y Q) wenzelm@17516: show ?case wenzelm@17516: proof (rule linorder_cases) wenzelm@17516: assume "x < y" wenzelm@17516: with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps) wenzelm@17516: next wenzelm@17516: assume "x = y" wenzelm@17516: with 6 show ?case by (simp add: mkPinj_ci) wenzelm@17516: next wenzelm@17516: assume "x > y" wenzelm@17516: with 6 show ?case by (simp add: mkPinj_ci ring_eq_simps) wenzelm@17516: qed wenzelm@17516: next wenzelm@17516: case (7 x P Q y R) wenzelm@17516: have "x = 0 \ x = 1 \ x > 1" by arith wenzelm@17516: moreover wenzelm@17516: { assume "x = 0" with 7 have ?case by simp } wenzelm@17516: moreover wenzelm@17516: { assume "x = 1" with 7 have ?case by (simp add: ring_eq_simps) } wenzelm@17516: moreover wenzelm@17516: { assume "x > 1" from 7 have ?case by (cases x) simp_all } wenzelm@17516: ultimately show ?case by blast wenzelm@17516: next wenzelm@17516: case (8 P x R y Q) wenzelm@17516: have "y = 0 \ y = 1 \ y > 1" by arith wenzelm@17516: moreover wenzelm@17516: { assume "y = 0" with 8 have ?case by simp } wenzelm@17516: moreover wenzelm@17516: { assume "y = 1" with 8 have ?case by simp } wenzelm@17516: moreover wenzelm@17516: { assume "y > 1" with 8 have ?case by simp } wenzelm@17516: ultimately show ?case by blast wenzelm@17516: next wenzelm@17516: case (9 P1 x P2 Q1 y Q2) wenzelm@17516: show ?case wenzelm@17516: proof (rule linorder_cases) wenzelm@17516: assume a: "x < y" hence "EX d. d + x = y" by arith wenzelm@17516: with 9 a show ?case by (auto simp add: mkPX_ci power_add ring_eq_simps) wenzelm@17516: next wenzelm@17516: assume a: "y < x" hence "EX d. d + y = x" by arith wenzelm@17516: with 9 a show ?case by (auto simp add: power_add mkPX_ci ring_eq_simps) wenzelm@17516: next wenzelm@17516: assume "x = y" wenzelm@17516: with 9 show ?case by (simp add: mkPX_ci ring_eq_simps) wenzelm@17516: qed wenzelm@17516: qed (auto simp add: ring_eq_simps) wenzelm@17516: wenzelm@17516: text {* Multiplication *} haftmann@22742: lemma mul_ci: "Ipol l (P \ Q) = Ipol l P * Ipol l Q" wenzelm@20622: by (induct P Q arbitrary: l rule: mul.induct) wenzelm@17516: (simp_all add: mkPX_ci mkPinj_ci ring_eq_simps add_ci power_add) wenzelm@17516: wenzelm@17516: text {* Substraction *} haftmann@22742: lemma sub_ci: "Ipol l (P \ Q) = Ipol l P - Ipol l Q" wenzelm@17516: by (simp add: add_ci neg_ci sub_def) wenzelm@17516: wenzelm@17516: text {* Square *} haftmann@22742: lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" haftmann@22742: by (induct P arbitrary: ls) wenzelm@20622: (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci ring_eq_simps power_add) wenzelm@17516: wenzelm@17516: text {* Power *} haftmann@22742: lemma even_pow:"even n \ pow n P = pow (n div 2) (sqr P)" wenzelm@20622: by (induct n) simp_all wenzelm@17516: haftmann@22742: lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" haftmann@22742: proof (induct n arbitrary: P rule: nat_less_induct) wenzelm@17516: case (1 k) wenzelm@17516: show ?case wenzelm@17516: proof (cases k) wenzelm@20622: case 0 wenzelm@20622: then show ?thesis by simp wenzelm@20622: next wenzelm@17516: case (Suc l) wenzelm@17516: show ?thesis wenzelm@17516: proof cases wenzelm@20622: assume "even l" wenzelm@20622: then have "Suc l div 2 = l div 2" wenzelm@20622: by (simp add: nat_number even_nat_plus_one_div_two) wenzelm@17516: moreover wenzelm@17516: from Suc have "l < k" by simp haftmann@22742: with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp wenzelm@17516: moreover wenzelm@20622: note Suc `even l` even_nat_plus_one_div_two wenzelm@17516: ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) wenzelm@17516: next wenzelm@20622: assume "odd l" wenzelm@20622: { wenzelm@20622: fix p haftmann@22742: have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" wenzelm@20622: proof (cases l) wenzelm@20622: case 0 wenzelm@20622: with `odd l` show ?thesis by simp wenzelm@20622: next wenzelm@20622: case (Suc w) wenzelm@20622: with `odd l` have "even w" by simp wenzelm@20678: have two_times: "2 * (w div 2) = w" wenzelm@20678: by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) haftmann@22742: have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" wenzelm@20622: by (simp add: power_Suc) haftmann@22742: then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" wenzelm@20678: by (simp add: numerals) wenzelm@20622: with Suc show ?thesis wenzelm@20678: by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci) wenzelm@20622: qed wenzelm@20622: } with 1 Suc `odd l` show ?thesis by simp wenzelm@17516: qed wenzelm@17516: qed wenzelm@17516: qed wenzelm@17516: wenzelm@17516: text {* Normalization preserves semantics *} wenzelm@20622: lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)" wenzelm@17516: by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) wenzelm@17516: wenzelm@17516: text {* Reflection lemma: Key to the (incomplete) decision procedure *} wenzelm@17516: lemma norm_eq: wenzelm@20622: assumes "norm P1 = norm P2" wenzelm@17516: shows "Ipolex l P1 = Ipolex l P2" wenzelm@17516: proof - wenzelm@20622: from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp wenzelm@20622: then show ?thesis by (simp only: norm_ci) wenzelm@17516: qed wenzelm@17516: wenzelm@17516: wenzelm@17516: use "comm_ring.ML" wenzelm@18708: setup CommRing.setup wenzelm@17516: wenzelm@17516: end