# HG changeset patch # User haftmann # Date 1256907762 -3600 # Node ID 85adf83af7ce1c93b82984ee31f17eae0694110d # Parent ee5b5ef5e970abcdf3d14dc0ea6130499824b97c# Parent 1f18de40b43f8166a2ceeb85ff9f43aa7cac14fb merged diff -r ee5b5ef5e970 -r 85adf83af7ce NEWS --- a/NEWS Fri Oct 30 11:31:34 2009 +0100 +++ b/NEWS Fri Oct 30 14:02:42 2009 +0100 @@ -37,6 +37,11 @@ *** HOL *** +* Combined former theories Divides and IntDiv to one theory Divides +in the spirit of other number theory theories in HOL; some constants +(and to a lesser extent also facts) have been suffixed with _nat und _int +respectively. INCOMPATIBILITY. + * Most rules produced by inductive and datatype package have mandatory prefixes. INCOMPATIBILITY. diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Decision_Procs/Commutative_Ring.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy Fri Oct 30 14:02:42 2009 +0100 @@ -0,0 +1,319 @@ +(* Author: Bernhard Haeupler + +Proving equalities in commutative rings done "right" in Isabelle/HOL. +*) + +header {* Proving equalities in commutative rings *} + +theory Commutative_Ring +imports Main Parity +uses ("commutative_ring_tac.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. *} + +primrec + Ipol :: "'a::{comm_ring_1} list \ 'a pol \ 'a" +where + "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 :: "'a::{comm_ring_1} list \ 'a polex \ 'a" +where + "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. *} + +definition + mkPinj :: "nat \ 'a pol \ 'a pol" where + "mkPinj x P = (case P of + Pc c \ Pc c | + Pinj y P \ Pinj (x + y) P | + PX p1 y p2 \ Pinj x P)" + +definition + mkPX :: "'a::{comm_ring} pol \ nat \ 'a pol \ 'a pol" where + "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 *} + +function + add :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) +where + "Pc a \ Pc b = Pc (a + b)" + | "Pc c \ Pinj i P = Pinj i (P \ Pc c)" + | "Pinj i P \ Pc c = Pinj i (P \ Pc c)" + | "Pc c \ PX P i Q = PX P i (Q \ Pc c)" + | "PX P i Q \ Pc c = PX P i (Q \ Pc c)" + | "Pinj x P \ Pinj y Q = + (if x = y then mkPinj x (P \ Q) + else (if x > y then mkPinj y (Pinj (x - y) P \ Q) + else mkPinj x (Pinj (y - x) Q \ P)))" + | "Pinj x P \ PX Q y R = + (if x = 0 then P \ PX Q y R + else (if x = 1 then PX Q y (R \ P) + else PX Q y (R \ Pinj (x - 1) P)))" + | "PX P x R \ Pinj y Q = + (if y = 0 then PX P x R \ Q + else (if y = 1 then PX P x (R \ Q) + else PX P x (R \ Pinj (y - 1) Q)))" + | "PX P1 x P2 \ PX Q1 y Q2 = + (if x = y then mkPX (P1 \ Q1) x (P2 \ Q2) + else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \ Q1) y (P2 \ Q2) + else mkPX (PX Q1 (y-x) (Pc 0) \ P1) x (P2 \ Q2)))" +by pat_completeness auto +termination by (relation "measure (\(x, y). size x + size y)") auto + +function + mul :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 70) +where + "Pc a \ Pc b = Pc (a * b)" + | "Pc c \ Pinj i P = + (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" + | "Pinj i P \ Pc c = + (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" + | "Pc c \ PX P i Q = + (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" + | "PX P i Q \ Pc c = + (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" + | "Pinj x P \ Pinj y Q = + (if x = y then mkPinj x (P \ Q) else + (if x > y then mkPinj y (Pinj (x-y) P \ Q) + else mkPinj x (Pinj (y - x) Q \ P)))" + | "Pinj x P \ PX Q y R = + (if x = 0 then P \ PX Q y R else + (if x = 1 then mkPX (Pinj x P \ Q) y (R \ P) + else mkPX (Pinj x P \ Q) y (R \ Pinj (x - 1) P)))" + | "PX P x R \ Pinj y Q = + (if y = 0 then PX P x R \ Q else + (if y = 1 then mkPX (Pinj y Q \ P) x (R \ Q) + else mkPX (Pinj y Q \ P) x (R \ Pinj (y - 1) Q)))" + | "PX P1 x P2 \ PX Q1 y Q2 = + mkPX (P1 \ Q1) (x + y) (P2 \ Q2) \ + (mkPX (P1 \ mkPinj 1 Q2) x (Pc 0) \ + (mkPX (Q1 \ mkPinj 1 P2) y (Pc 0)))" +by pat_completeness auto +termination by (relation "measure (\(x, y). size x + size y)") + (auto simp add: mkPinj_def split: pol.split) + +text {* Negation*} +primrec + neg :: "'a::{comm_ring} pol \ 'a pol" +where + "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 *} +definition + sub :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) +where + "sub P Q = P \ neg Q" + +text {* Square for Fast Exponentation *} +primrec + sqr :: "'a::{comm_ring_1} pol \ 'a pol" +where + "sqr (Pc c) = Pc (c * c)" + | "sqr (Pinj i P) = mkPinj i (sqr P)" + | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \ + mkPX (Pc (1 + 1) \ A \ mkPinj 1 B) x (Pc 0)" + +text {* Fast Exponentation *} +fun + pow :: "nat \ 'a::{comm_ring_1} pol \ 'a pol" +where + "pow 0 P = Pc 1" + | "pow n P = (if even n then pow (n div 2) (sqr P) + else P \ pow (n div 2) (sqr P))" + +lemma pow_if: + "pow n P = + (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) + else P \ pow (n div 2) (sqr P))" + by (cases n) simp_all + + +text {* Normalization of polynomial expressions *} + +primrec + norm :: "'a::{comm_ring_1} polex \ 'a pol" +where + "norm (Pol P) = P" + | "norm (Add P Q) = norm P \ norm Q" + | "norm (Sub P Q) = norm P \ norm Q" + | "norm (Mul P Q) = norm P \ norm Q" + | "norm (Pow P n) = pow n (norm P)" + | "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 algebra_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 algebra_simps) + +text {* Correctness theorems for the implemented operations *} + +text {* Negation *} +lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" + by (induct P arbitrary: l) auto + +text {* Addition *} +lemma add_ci: "Ipol l (P \ Q) = Ipol l P + Ipol l Q" +proof (induct P Q arbitrary: l 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 algebra_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 algebra_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: algebra_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 algebra_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 algebra_simps) + next + assume "x = y" + with 9 show ?case by (simp add: mkPX_ci algebra_simps) + qed +qed (auto simp add: algebra_simps) + +text {* Multiplication *} +lemma mul_ci: "Ipol l (P \ Q) = Ipol l P * Ipol l Q" + by (induct P Q arbitrary: l rule: mul.induct) + (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add) + +text {* Substraction *} +lemma sub_ci: "Ipol l (P \ Q) = Ipol l P - Ipol l Q" + by (simp add: add_ci neg_ci sub_def) + +text {* Square *} +lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" + by (induct P arbitrary: ls) + (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add) + +text {* Power *} +lemma even_pow:"even n \ pow n P = pow (n div 2) (sqr P)" + by (induct n) simp_all + +lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" +proof (induct n arbitrary: P rule: nat_less_induct) + case (1 k) + show ?case + proof (cases k) + case 0 + then show ?thesis by simp + next + case (Suc l) + show ?thesis + proof cases + assume "even l" + then have "Suc l div 2 = l div 2" + by (simp add: nat_number even_nat_plus_one_div_two) + moreover + from Suc have "l < k" by simp + with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp + moreover + note Suc `even l` even_nat_plus_one_div_two + ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) + next + assume "odd l" + { + fix p + have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" + proof (cases l) + case 0 + with `odd l` show ?thesis by simp + next + case (Suc w) + with `odd l` have "even w" by simp + have two_times: "2 * (w div 2) = w" + by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) + have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" + by (simp add: power_Suc) + then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" + by (simp add: numerals) + with Suc show ?thesis + by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci + simp del: power_Suc) + qed + } with 1 Suc `odd l` show ?thesis by simp + qed + 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 "norm P1 = norm P2" + shows "Ipolex l P1 = Ipolex l P2" +proof - + from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp + then show ?thesis by (simp only: norm_ci) +qed + + +use "commutative_ring_tac.ML" +setup Commutative_Ring_Tac.setup + +end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Decision_Procs/Commutative_Ring_Complete.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy Fri Oct 30 14:02:42 2009 +0100 @@ -0,0 +1,391 @@ +(* Author: Bernhard Haeupler + +This theory is about of the relative completeness of method comm-ring +method. As long as the reified atomic polynomials of type 'a pol are +in normal form, the cring method is complete. +*) + +header {* Proof of the relative completeness of method comm-ring *} + +theory Commutative_Ring_Complete +imports Commutative_Ring +begin + +text {* Formalization of normal form *} +fun + isnorm :: "('a::{comm_ring}) pol \ bool" +where + "isnorm (Pc c) \ True" + | "isnorm (Pinj i (Pc c)) \ False" + | "isnorm (Pinj i (Pinj j Q)) \ False" + | "isnorm (Pinj 0 P) \ False" + | "isnorm (Pinj i (PX Q1 j Q2)) \ isnorm (PX Q1 j Q2)" + | "isnorm (PX P 0 Q) \ False" + | "isnorm (PX (Pc c) i Q) \ c \ 0 \ isnorm Q" + | "isnorm (PX (PX P1 j (Pc c)) i Q) \ c \ 0 \ isnorm (PX P1 j (Pc c)) \ isnorm Q" + | "isnorm (PX P i Q) \ isnorm P \ isnorm Q" + +(* Some helpful lemmas *) +lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False" +by(cases P, auto) + +lemma norm_PX_0_False:"isnorm (PX (Pc 0) i Q) = False" +by(cases i, auto) + +lemma norm_Pinj:"isnorm (Pinj i Q) \ isnorm Q" +by(cases i,simp add: norm_Pinj_0_False norm_PX_0_False,cases Q) auto + +lemma norm_PX2:"isnorm (PX P i Q) \ isnorm Q" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma norm_PX1:"isnorm (PX P i Q) \ isnorm P" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma mkPinj_cn:"\y~=0; isnorm Q\ \ isnorm (mkPinj y Q)" +apply(auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) +apply(case_tac nat, auto simp add: norm_Pinj_0_False) +by(case_tac pol, auto) (case_tac y, auto) + +lemma norm_PXtrans: + assumes A:"isnorm (PX P x Q)" and "isnorm Q2" + shows "isnorm (PX P x Q2)" +proof(cases P) + case (PX p1 y p2) from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc from prems show ?thesis by(cases x, auto) +next + case Pinj from prems show ?thesis by(cases x, auto) +qed + +lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)" +proof(cases P) + case (PX p1 y p2) + from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc + from prems show ?thesis by(cases x, auto) +next + case Pinj + from prems show ?thesis by(cases x, auto) +qed + +text {* mkPX conserves normalizedness (@{text "_cn"}) *} +lemma mkPX_cn: + assumes "x \ 0" and "isnorm P" and "isnorm Q" + shows "isnorm (mkPX P x Q)" +proof(cases P) + case (Pc c) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (Pinj i Q) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (PX P1 y P2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) + with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto) +qed + +text {* add conserves normalizedness *} +lemma add_cn:"isnorm P \ isnorm Q \ isnorm (P \ Q)" +proof(induct P Q rule: add.induct) + case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (3 i P2 c) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (6 x P2 y Q2) + from prems have Y0:"y>0" by (cases y, auto simp add: norm_Pinj_0_False) + from prems have X0:"x>0" by (cases x, auto simp add: norm_Pinj_0_False) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "xy" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems Y0 + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + have "x = 0 \ x = 1 \ x > 1" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have X0:"x>0" by(cases x, auto) + from prems have NP1:"isnorm P1" and NP2:"isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) + from prems have NQ1:"isnorm Q1" and NQ2:"isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) + have "y < x \ x = y \ x < y" by arith + moreover + {assume sm1:"y < x" hence "EX d. x=d+y" by arith + then obtain d where sm2:"x=d+y".. + note prems NQ1 NP1 NP2 NQ2 sm1 sm2 + moreover + have "isnorm (PX P1 d (Pc 0))" + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases d, simp,cases p2, auto) + next case Pc from prems show ?thesis by(cases d, auto) + next case Pinj from prems show ?thesis by(cases d, auto) + qed + ultimately have "isnorm (P2 \ Q2)" "isnorm (PX P1 (x - y) (Pc 0) \ Q1)" by auto + with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)} + moreover + {assume "x=y" + from prems NP1 NP2 NQ1 NQ2 have "isnorm (P2 \ Q2)" "isnorm (P1 \ Q1)" by auto + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + {assume sm1:"x Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \ P1)" by auto + with X0 sm1 sm2 have ?case by (simp add: mkPX_cn)} + ultimately show ?case by blast +qed simp + +text {* mul concerves normalizedness *} +lemma mul_cn :"isnorm P \ isnorm Q \ isnorm (P \ Q)" +proof(induct P Q rule: mul.induct) + case (2 c i P2) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (3 i P2 c) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (6 x P2 y Q2) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "x0" by (cases x, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + moreover + { assume "x=y" + moreover + from prems have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + note prems + moreover + ultimately have ?case by (simp add: mkPinj_cn) } + moreover + { assume "x>y" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems + moreover + from prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn) } + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn)} + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto + with Y0 X have ?case by (simp add: mkPX_cn)} + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto + with Y0 X have ?case by (simp add: mkPX_cn) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have X0:"x>0" by(cases x, auto) + from prems have Y0:"y>0" by(cases y, auto) + note prems + moreover + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + moreover + from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) + ultimately have "isnorm (P1 \ Q1)" "isnorm (P2 \ Q2)" + "isnorm (P1 \ mkPinj 1 Q2)" "isnorm (Q1 \ mkPinj 1 P2)" + by (auto simp add: mkPinj_cn) + with prems X0 Y0 have + "isnorm (mkPX (P1 \ Q1) (x + y) (P2 \ Q2))" + "isnorm (mkPX (P1 \ mkPinj (Suc 0) Q2) x (Pc 0))" + "isnorm (mkPX (Q1 \ mkPinj (Suc 0) P2) y (Pc 0))" + by (auto simp add: mkPX_cn) + thus ?case by (simp add: add_cn) +qed(simp) + +text {* neg conserves normalizedness *} +lemma neg_cn: "isnorm P \ isnorm (neg P)" +proof (induct P) + case (Pinj i P2) + from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2]) + with prems show ?case by(cases P2, auto, cases i, auto) +next + case (PX P1 x P2) + from prems have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems show ?case + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases x, auto, cases p2, auto) + next + case Pinj + with prems show ?thesis by(cases x, auto) + qed(cases x, auto) +qed(simp) + +text {* sub conserves normalizedness *} +lemma sub_cn:"isnorm p \ isnorm q \ isnorm (p \ q)" +by (simp add: sub_def add_cn neg_cn) + +text {* sqr conserves normalizizedness *} +lemma sqr_cn:"isnorm P \ isnorm (sqr P)" +proof(induct P) + case (Pinj i Q) + from prems show ?case by(cases Q, auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn) +next + case (PX P1 x P2) + from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems have + "isnorm (mkPX (Pc (1 + 1) \ P1 \ mkPinj (Suc 0) P2) x (Pc 0))" + and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" + by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) + thus ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) +qed simp + +text {* pow conserves normalizedness *} +lemma pow_cn:"isnorm P \ isnorm (pow n P)" +proof (induct n arbitrary: P rule: nat_less_induct) + case (1 k) + show ?case + proof (cases "k=0") + case False + then have K2:"k div 2 < k" by (cases k, auto) + from prems have "isnorm (sqr P)" by (simp add: sqr_cn) + with prems K2 show ?thesis + by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn) + qed simp +qed + +end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Decision_Procs/Decision_Procs.thy --- a/src/HOL/Decision_Procs/Decision_Procs.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Decision_Procs/Decision_Procs.thy Fri Oct 30 14:02:42 2009 +0100 @@ -1,7 +1,10 @@ -header {* Various decision procedures. typically involving reflection *} +header {* Various decision procedures, typically involving reflection *} theory Decision_Procs -imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" Parametric_Ferrante_Rackoff +imports + Commutative_Ring Cooper Ferrack MIR Approximation Dense_Linear_Order Parametric_Ferrante_Rackoff + Commutative_Ring_Complete + "ex/Commutative_Ring_Ex" "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" begin end \ No newline at end of file diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Decision_Procs/commutative_ring_tac.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Fri Oct 30 14:02:42 2009 +0100 @@ -0,0 +1,104 @@ +(* Author: Amine Chaieb + +Tactic for solving equalities over commutative rings. +*) + +signature COMMUTATIVE_RING_TAC = +sig + val tac: Proof.context -> int -> tactic + val setup: theory -> theory +end + +structure Commutative_Ring_Tac: COMMUTATIVE_RING_TAC = +struct + +(* Zero and One of the commutative ring *) +fun cring_zero T = Const (@{const_name HOL.zero}, T); +fun cring_one T = Const (@{const_name HOL.one}, T); + +(* reification functions *) +(* add two polynom expressions *) +fun polT t = Type (@{type_name Commutative_Ring.pol}, [t]); +fun polexT t = Type (@{type_name Commutative_Ring.polex}, [t]); + +(* pol *) +fun pol_Pc t = Const (@{const_name Commutative_Ring.pol.Pc}, t --> polT t); +fun pol_Pinj t = Const (@{const_name Commutative_Ring.pol.Pinj}, HOLogic.natT --> polT t --> polT t); +fun pol_PX t = Const (@{const_name Commutative_Ring.pol.PX}, polT t --> HOLogic.natT --> polT t --> polT t); + +(* polex *) +fun polex_add t = Const (@{const_name Commutative_Ring.polex.Add}, polexT t --> polexT t --> polexT t); +fun polex_sub t = Const (@{const_name Commutative_Ring.polex.Sub}, polexT t --> polexT t --> polexT t); +fun polex_mul t = Const (@{const_name Commutative_Ring.polex.Mul}, polexT t --> polexT t --> polexT t); +fun polex_neg t = Const (@{const_name Commutative_Ring.polex.Neg}, polexT t --> polexT t); +fun polex_pol t = Const (@{const_name Commutative_Ring.polex.Pol}, polT t --> polexT t); +fun polex_pow t = Const (@{const_name Commutative_Ring.polex.Pow}, polexT t --> HOLogic.natT --> polexT t); + +(* reification of polynoms : primitive cring expressions *) +fun reif_pol T vs (t as Free _) = + let + val one = @{term "1::nat"}; + val i = find_index (fn t' => t' = 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 $ HOLogic.mk_nat i + $ (pol_PX T $ (pol_Pc T $ cring_one T) + $ one $ (pol_Pc T $ cring_zero T)) + end + | reif_pol T vs t = pol_Pc T $ t; + +(* reification of polynom expressions *) +fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) = + polex_add T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) = + polex_sub T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) = + polex_mul T $ reif_polex T vs a $ reif_polex T vs b + | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) = + polex_neg T $ reif_polex T vs a + | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) = + polex_pow T $ reif_polex T vs a $ n + | reif_polex T vs t = polex_pol T $ reif_pol T vs t; + +(* reification of the equation *) +val cr_sort = @{sort "comm_ring_1"}; + +fun reif_eq thy (eq as Const(@{const_name "op ="}, Type("fun", [T, _])) $ lhs $ rhs) = + if Sign.of_sort thy (T, cr_sort) then + let + val fs = OldTerm.term_frees eq; + val cvs = cterm_of thy (HOLogic.mk_list T fs); + val clhs = cterm_of thy (reif_polex T fs lhs); + val crhs = cterm_of thy (reif_polex T fs rhs); + val ca = ctyp_of thy T; + in (ca, cvs, clhs, crhs) end + else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort) + | reif_eq _ _ = error "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 *) +val cring_simps = + [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add}, + @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]]; + +fun tac ctxt = SUBGOAL (fn (g, i) => + let + val thy = ProofContext.theory_of ctxt; + val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) + addsimps cring_simps; + val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) + val norm_eq_th = + simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) + in + cut_rules_tac [norm_eq_th] i + THEN (simp_tac cring_ss i) + THEN (simp_tac cring_ss i) + end); + +val setup = + Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o tac)) + "reflective decision procedure for equalities over commutative rings"; + +end; diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/ex/Commutative_Ring_Ex.thy Fri Oct 30 14:02:42 2009 +0100 @@ -0,0 +1,48 @@ +(* Author: Bernhard Haeupler *) + +header {* Some examples demonstrating the comm-ring method *} + +theory Commutative_Ring_Ex +imports Commutative_Ring +begin + +lemma "4*(x::int)^5*y^3*x^2*3 + x*z + 3^5 = 12*x^7*y^3 + z*x + 243" +by comm_ring + +lemma "((x::int) + y)^2 = x^2 + y^2 + 2*x*y" +by comm_ring + +lemma "((x::int) + y)^3 = x^3 + y^3 + 3*x^2*y + 3*y^2*x" +by comm_ring + +lemma "((x::int) - y)^3 = x^3 + 3*x*y^2 + (-3)*y*x^2 - y^3" +by comm_ring + +lemma "((x::int) - y)^2 = x^2 + y^2 - 2*x*y" +by comm_ring + +lemma " ((a::int) + b + c)^2 = a^2 + b^2 + c^2 + 2*a*b + 2*b*c + 2*a*c" +by comm_ring + +lemma "((a::int) - b - c)^2 = a^2 + b^2 + c^2 - 2*a*b + 2*b*c - 2*a*c" +by comm_ring + +lemma "(a::int)*b + a*c = a*(b+c)" +by comm_ring + +lemma "(a::int)^2 - b^2 = (a - b) * (a + b)" +by comm_ring + +lemma "(a::int)^3 - b^3 = (a - b) * (a^2 + a*b + b^2)" +by comm_ring + +lemma "(a::int)^3 + b^3 = (a + b) * (a^2 - a*b + b^2)" +by comm_ring + +lemma "(a::int)^4 - b^4 = (a - b) * (a + b)*(a^2 + b^2)" +by comm_ring + +lemma "(a::int)^10 - b^10 = (a - b) * (a^9 + a^8*b + a^7*b^2 + a^6*b^3 + a^5*b^4 + a^4*b^5 + a^3*b^6 + a^2*b^7 + a*b^8 + b^9 )" +by comm_ring + +end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Divides.thy --- a/src/HOL/Divides.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Divides.thy Fri Oct 30 14:02:42 2009 +0100 @@ -634,6 +634,11 @@ end +lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else + let (q, r) = divmod_nat (m - n) n in (Suc q, r))" +by (simp add: divmod_nat_zero divmod_nat_base divmod_nat_step) + (simp add: divmod_nat_div_mod) + text {* Simproc for cancelling @{const div} and @{const mod} *} ML {* @@ -666,22 +671,6 @@ end *} -text {* code generator setup *} - -lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \ m < n then (0, m) else - let (q, r) = divmod_nat (m - n) n in (Suc q, r))" -by (simp add: divmod_nat_zero divmod_nat_base divmod_nat_step) - (simp add: divmod_nat_div_mod) - -code_modulename SML - Divides Nat - -code_modulename OCaml - Divides Nat - -code_modulename Haskell - Divides Nat - subsubsection {* Quotient *} @@ -1136,4 +1125,1339 @@ Suc_mod_eq_add3_mod [of _ "number_of v", standard] declare Suc_mod_eq_add3_mod_number_of [simp] + +lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" +apply (induct "m") +apply (simp_all add: mod_Suc) +done + +declare Suc_times_mod_eq [of "number_of w", standard, simp] + +lemma [simp]: "n div k \ (Suc n) div k" +by (simp add: div_le_mono) + +lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" +by (cases n) simp_all + +lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" +using Suc_n_div_2_gt_zero [of "n - 1"] by simp + + (* Potential use of algebra : Equality modulo n*) +lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" +by (simp add: mult_ac add_ac) + +lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" +proof - + have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp + also have "... = Suc m mod n" by (rule mod_mult_self3) + finally show ?thesis . +qed + +lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" +apply (subst mod_Suc [of m]) +apply (subst mod_Suc [of "m mod n"], simp) +done + + +subsection {* Division on @{typ int} *} + +definition divmod_int_rel :: "int \ int \ int \ int \ bool" where + --{*definition of quotient and remainder*} + [code]: "divmod_int_rel a b = (\(q, r). a = b * q + r \ + (if 0 < b then 0 \ r \ r < b else b < r \ r \ 0))" + +definition adjust :: "int \ int \ int \ int \ int" where + --{*for the division algorithm*} + [code]: "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) + else (2 * q, r))" + +text{*algorithm for the case @{text "a\0, b>0"}*} +function posDivAlg :: "int \ int \ int \ int" where + "posDivAlg a b = (if a < b \ b \ 0 then (0, a) + else adjust b (posDivAlg a (2 * b)))" +by auto +termination by (relation "measure (\(a, b). nat (a - b + 1))") + (auto simp add: mult_2) + +text{*algorithm for the case @{text "a<0, b>0"}*} +function negDivAlg :: "int \ int \ int \ int" where + "negDivAlg a b = (if 0 \a + b \ b \ 0 then (-1, a + b) + else adjust b (negDivAlg a (2 * b)))" +by auto +termination by (relation "measure (\(a, b). nat (- a - b))") + (auto simp add: mult_2) + +text{*algorithm for the general case @{term "b\0"}*} +definition negateSnd :: "int \ int \ int \ int" where + [code_unfold]: "negateSnd = apsnd uminus" + +definition divmod_int :: "int \ int \ int \ int" where + --{*The full division algorithm considers all possible signs for a, b + including the special case @{text "a=0, b<0"} because + @{term negDivAlg} requires @{term "a<0"}.*} + "divmod_int a b = (if 0 \ a then if 0 \ b then posDivAlg a b + else if a = 0 then (0, 0) + else negateSnd (negDivAlg (-a) (-b)) + else + if 0 < b then negDivAlg a b + else negateSnd (posDivAlg (-a) (-b)))" + +instantiation int :: Divides.div +begin + +definition + "a div b = fst (divmod_int a b)" + +definition + "a mod b = snd (divmod_int a b)" + +instance .. + end + +lemma divmod_int_mod_div: + "divmod_int p q = (p div q, p mod q)" + by (auto simp add: div_int_def mod_int_def) + +text{* +Here is the division algorithm in ML: + +\begin{verbatim} + fun posDivAlg (a,b) = + if ar-b then (2*q+1, r-b) else (2*q, r) + end + + fun negDivAlg (a,b) = + if 0\a+b then (~1,a+b) + else let val (q,r) = negDivAlg(a, 2*b) + in if 0\r-b then (2*q+1, r-b) else (2*q, r) + end; + + fun negateSnd (q,r:int) = (q,~r); + + fun divmod (a,b) = if 0\a then + if b>0 then posDivAlg (a,b) + else if a=0 then (0,0) + else negateSnd (negDivAlg (~a,~b)) + else + if 0 b*q + r; 0 \ r'; r' < b; r < b |] + ==> q' \ (q::int)" +apply (subgoal_tac "r' + b * (q'-q) \ r") + prefer 2 apply (simp add: right_diff_distrib) +apply (subgoal_tac "0 < b * (1 + q - q') ") +apply (erule_tac [2] order_le_less_trans) + prefer 2 apply (simp add: right_diff_distrib right_distrib) +apply (subgoal_tac "b * q' < b * (1 + q) ") + prefer 2 apply (simp add: right_diff_distrib right_distrib) +apply (simp add: mult_less_cancel_left) +done + +lemma unique_quotient_lemma_neg: + "[| b*q' + r' \ b*q + r; r \ 0; b < r; b < r' |] + ==> q \ (q'::int)" +by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, + auto) + +lemma unique_quotient: + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + ==> q = q'" +apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) +apply (blast intro: order_antisym + dest: order_eq_refl [THEN unique_quotient_lemma] + order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ +done + + +lemma unique_remainder: + "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] + ==> r = r'" +apply (subgoal_tac "q = q'") + apply (simp add: divmod_int_rel_def) +apply (blast intro: unique_quotient) +done + + +subsubsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*} + +text{*And positive divisors*} + +lemma adjust_eq [simp]: + "adjust b (q,r) = + (let diff = r-b in + if 0 \ diff then (2*q + 1, diff) + else (2*q, r))" +by (simp add: Let_def adjust_def) + +declare posDivAlg.simps [simp del] + +text{*use with a simproc to avoid repeatedly proving the premise*} +lemma posDivAlg_eqn: + "0 < b ==> + posDivAlg a b = (if a a" and "0 < b" + shows "divmod_int_rel a b (posDivAlg a b)" +using prems apply (induct a b rule: posDivAlg.induct) +apply auto +apply (simp add: divmod_int_rel_def) +apply (subst posDivAlg_eqn, simp add: right_distrib) +apply (case_tac "a < b") +apply simp_all +apply (erule splitE) +apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) +done + + +subsubsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} + +text{*And positive divisors*} + +declare negDivAlg.simps [simp del] + +text{*use with a simproc to avoid repeatedly proving the premise*} +lemma negDivAlg_eqn: + "0 < b ==> + negDivAlg a b = + (if 0\a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" +by (rule negDivAlg.simps [THEN trans], simp) + +(*Correctness of negDivAlg: it computes quotients correctly + It doesn't work if a=0 because the 0/b equals 0, not -1*) +lemma negDivAlg_correct: + assumes "a < 0" and "b > 0" + shows "divmod_int_rel a b (negDivAlg a b)" +using prems apply (induct a b rule: negDivAlg.induct) +apply (auto simp add: linorder_not_le) +apply (simp add: divmod_int_rel_def) +apply (subst negDivAlg_eqn, assumption) +apply (case_tac "a + b < (0\int)") +apply simp_all +apply (erule splitE) +apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) +done + + +subsubsection{*Existence Shown by Proving the Division Algorithm to be Correct*} + +(*the case a=0*) +lemma divmod_int_rel_0: "b \ 0 ==> divmod_int_rel 0 b (0, 0)" +by (auto simp add: divmod_int_rel_def linorder_neq_iff) + +lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" +by (subst posDivAlg.simps, auto) + +lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" +by (subst negDivAlg.simps, auto) + +lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" +by (simp add: negateSnd_def) + +lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)" +by (auto simp add: split_ifs divmod_int_rel_def) + +lemma divmod_int_correct: "b \ 0 ==> divmod_int_rel a b (divmod_int a b)" +by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg + posDivAlg_correct negDivAlg_correct) + +text{*Arbitrary definitions for division by zero. Useful to simplify + certain equations.*} + +lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" +by (simp add: div_int_def mod_int_def divmod_int_def posDivAlg.simps) + + +text{*Basic laws about division and remainder*} + +lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" +apply (case_tac "b = 0", simp) +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +done + +lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" +by(simp add: zmod_zdiv_equality[symmetric]) + +lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" +by(simp add: mult_commute zmod_zdiv_equality[symmetric]) + +text {* Tool setup *} + +ML {* +local + +structure CancelDivMod = CancelDivModFun(struct + + val div_name = @{const_name div}; + val mod_name = @{const_name mod}; + val mk_binop = HOLogic.mk_binop; + val mk_sum = Arith_Data.mk_sum HOLogic.intT; + val dest_sum = Arith_Data.dest_sum; + + val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; + + val trans = trans; + + val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac + (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) + +end) + +in + +val cancel_div_mod_int_proc = Simplifier.simproc @{theory} + "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); + +val _ = Addsimprocs [cancel_div_mod_int_proc]; + +end +*} + +lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def mod_int_def) +done + +lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] + and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] + +lemma neg_mod_conj : "b < (0::int) ==> a mod b \ 0 & b < a mod b" +apply (cut_tac a = a and b = b in divmod_int_correct) +apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) +done + +lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] + and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] + + + +subsubsection{*General Properties of div and mod*} + +lemma divmod_int_rel_div_mod: "b \ 0 ==> divmod_int_rel a b (a div b, a mod b)" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (force simp add: divmod_int_rel_def linorder_neq_iff) +done + +lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a div b = q" +by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) + +lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a mod b = r" +by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) + +lemma div_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a div b = 0" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +lemma div_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a div b = 0" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a div b = -1" +apply (rule divmod_int_rel_div) +apply (auto simp add: divmod_int_rel_def) +done + +(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) + +lemma mod_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a mod b = a" +apply (rule_tac q = 0 in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +lemma mod_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a mod b = a" +apply (rule_tac q = 0 in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a mod b = a+b" +apply (rule_tac q = "-1" in divmod_int_rel_mod) +apply (auto simp add: divmod_int_rel_def) +done + +text{*There is no @{text mod_neg_pos_trivial}.*} + + +(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) +lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" +apply (case_tac "b = 0", simp) +apply (simp add: divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, + THEN divmod_int_rel_div, THEN sym]) + +done + +(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) +lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" +apply (case_tac "b = 0", simp) +apply (subst divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, THEN divmod_int_rel_mod], + auto) +done + + +subsubsection{*Laws for div and mod with Unary Minus*} + +lemma zminus1_lemma: + "divmod_int_rel a b (q, r) + ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, + if r=0 then 0 else b-r)" +by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib) + + +lemma zdiv_zminus1_eq_if: + "b \ (0::int) + ==> (-a) div b = + (if a mod b = 0 then - (a div b) else - (a div b) - 1)" +by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_div]) + +lemma zmod_zminus1_eq_if: + "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" +apply (case_tac "b = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_mod]) +done + +lemma zmod_zminus1_not_zero: + fixes k l :: int + shows "- k mod l \ 0 \ k mod l \ 0" + unfolding zmod_zminus1_eq_if by auto + +lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" +by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) + +lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" +by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) + +lemma zdiv_zminus2_eq_if: + "b \ (0::int) + ==> a div (-b) = + (if a mod b = 0 then - (a div b) else - (a div b) - 1)" +by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) + +lemma zmod_zminus2_eq_if: + "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" +by (simp add: zmod_zminus1_eq_if zmod_zminus2) + +lemma zmod_zminus2_not_zero: + fixes k l :: int + shows "k mod - l \ 0 \ k mod l \ 0" + unfolding zmod_zminus2_eq_if by auto + + +subsubsection{*Division of a Number by Itself*} + +lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \ q" +apply (subgoal_tac "0 < a*q") + apply (simp add: zero_less_mult_iff, arith) +done + +lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \ r |] ==> q \ 1" +apply (subgoal_tac "0 \ a* (1-q) ") + apply (simp add: zero_le_mult_iff) +apply (simp add: right_diff_distrib) +done + +lemma self_quotient: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> q = 1" +apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff) +apply (rule order_antisym, safe, simp_all) +apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) +apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) +apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ +done + +lemma self_remainder: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> r = 0" +apply (frule self_quotient, assumption) +apply (simp add: divmod_int_rel_def) +done + +lemma zdiv_self [simp]: "a \ 0 ==> a div a = (1::int)" +by (simp add: divmod_int_rel_div_mod [THEN self_quotient]) + +(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) +lemma zmod_self [simp]: "a mod a = (0::int)" +apply (case_tac "a = 0", simp) +apply (simp add: divmod_int_rel_div_mod [THEN self_remainder]) +done + + +subsubsection{*Computation of Division and Remainder*} + +lemma zdiv_zero [simp]: "(0::int) div b = 0" +by (simp add: div_int_def divmod_int_def) + +lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" +by (simp add: div_int_def divmod_int_def) + +lemma zmod_zero [simp]: "(0::int) mod b = 0" +by (simp add: mod_int_def divmod_int_def) + +lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" +by (simp add: mod_int_def divmod_int_def) + +text{*a positive, b positive *} + +lemma div_pos_pos: "[| 0 < a; 0 \ b |] ==> a div b = fst (posDivAlg a b)" +by (simp add: div_int_def divmod_int_def) + +lemma mod_pos_pos: "[| 0 < a; 0 \ b |] ==> a mod b = snd (posDivAlg a b)" +by (simp add: mod_int_def divmod_int_def) + +text{*a negative, b positive *} + +lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" +by (simp add: div_int_def divmod_int_def) + +lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" +by (simp add: mod_int_def divmod_int_def) + +text{*a positive, b negative *} + +lemma div_pos_neg: + "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" +by (simp add: div_int_def divmod_int_def) + +lemma mod_pos_neg: + "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" +by (simp add: mod_int_def divmod_int_def) + +text{*a negative, b negative *} + +lemma div_neg_neg: + "[| a < 0; b \ 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" +by (simp add: div_int_def divmod_int_def) + +lemma mod_neg_neg: + "[| a < 0; b \ 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" +by (simp add: mod_int_def divmod_int_def) + +text {*Simplify expresions in which div and mod combine numerical constants*} + +lemma divmod_int_relI: + "\a == b * q + r; if 0 < b then 0 \ r \ r < b else b < r \ r \ 0\ + \ divmod_int_rel a b (q, r)" + unfolding divmod_int_rel_def by simp + +lemmas divmod_int_rel_div_eq = divmod_int_relI [THEN divmod_int_rel_div, THEN eq_reflection] +lemmas divmod_int_rel_mod_eq = divmod_int_relI [THEN divmod_int_rel_mod, THEN eq_reflection] +lemmas arithmetic_simps = + arith_simps + add_special + OrderedGroup.add_0_left + OrderedGroup.add_0_right + mult_zero_left + mult_zero_right + mult_1_left + mult_1_right + +(* simprocs adapted from HOL/ex/Binary.thy *) +ML {* +local + val mk_number = HOLogic.mk_number HOLogic.intT; + fun mk_cert u k l = @{term "plus :: int \ int \ int"} $ + (@{term "times :: int \ int \ int"} $ u $ mk_number k) $ + mk_number l; + fun prove ctxt prop = Goal.prove ctxt [] [] prop + (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps})))); + fun binary_proc proc ss ct = + (case Thm.term_of ct of + _ $ t $ u => + (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of + SOME args => proc (Simplifier.the_context ss) args + | NONE => NONE) + | _ => NONE); +in + fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => + if n = 0 then NONE + else let val (k, l) = Integer.div_mod m n; + in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); +end +*} + +simproc_setup binary_int_div ("number_of m div number_of n :: int") = + {* K (divmod_proc (@{thm divmod_int_rel_div_eq})) *} + +simproc_setup binary_int_mod ("number_of m mod number_of n :: int") = + {* K (divmod_proc (@{thm divmod_int_rel_mod_eq})) *} + +lemmas posDivAlg_eqn_number_of [simp] = + posDivAlg_eqn [of "number_of v" "number_of w", standard] + +lemmas negDivAlg_eqn_number_of [simp] = + negDivAlg_eqn [of "number_of v" "number_of w", standard] + + +text{*Special-case simplification *} + +lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" +apply (cut_tac a = a and b = "-1" in neg_mod_sign) +apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) +apply (auto simp del: neg_mod_sign neg_mod_bound) +done + +lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" +by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) + +(** The last remaining special cases for constant arithmetic: + 1 div z and 1 mod z **) + +lemmas div_pos_pos_1_number_of [simp] = + div_pos_pos [OF int_0_less_1, of "number_of w", standard] + +lemmas div_pos_neg_1_number_of [simp] = + div_pos_neg [OF int_0_less_1, of "number_of w", standard] + +lemmas mod_pos_pos_1_number_of [simp] = + mod_pos_pos [OF int_0_less_1, of "number_of w", standard] + +lemmas mod_pos_neg_1_number_of [simp] = + mod_pos_neg [OF int_0_less_1, of "number_of w", standard] + + +lemmas posDivAlg_eqn_1_number_of [simp] = + posDivAlg_eqn [of concl: 1 "number_of w", standard] + +lemmas negDivAlg_eqn_1_number_of [simp] = + negDivAlg_eqn [of concl: 1 "number_of w", standard] + + + +subsubsection{*Monotonicity in the First Argument (Dividend)*} + +lemma zdiv_mono1: "[| a \ a'; 0 < (b::int) |] ==> a div b \ a' div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a' and b = b in zmod_zdiv_equality) +apply (rule unique_quotient_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + +lemma zdiv_mono1_neg: "[| a \ a'; (b::int) < 0 |] ==> a' div b \ a div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a' and b = b in zmod_zdiv_equality) +apply (rule unique_quotient_lemma_neg) +apply (erule subst) +apply (erule subst, simp_all) +done + + +subsubsection{*Monotonicity in the Second Argument (Divisor)*} + +lemma q_pos_lemma: + "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" +apply (subgoal_tac "0 < b'* (q' + 1) ") + apply (simp add: zero_less_mult_iff) +apply (simp add: right_distrib) +done + +lemma zdiv_mono2_lemma: + "[| b*q + r = b'*q' + r'; 0 \ b'*q' + r'; + r' < b'; 0 \ r; 0 < b'; b' \ b |] + ==> q \ (q'::int)" +apply (frule q_pos_lemma, assumption+) +apply (subgoal_tac "b*q < b* (q' + 1) ") + apply (simp add: mult_less_cancel_left) +apply (subgoal_tac "b*q = r' - r + b'*q'") + prefer 2 apply simp +apply (simp (no_asm_simp) add: right_distrib) +apply (subst add_commute, rule zadd_zless_mono, arith) +apply (rule mult_right_mono, auto) +done + +lemma zdiv_mono2: + "[| (0::int) \ a; 0 < b'; b' \ b |] ==> a div b \ a div b'" +apply (subgoal_tac "b \ 0") + prefer 2 apply arith +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a and b = b' in zmod_zdiv_equality) +apply (rule zdiv_mono2_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + +lemma q_neg_lemma: + "[| b'*q' + r' < 0; 0 \ r'; 0 < b' |] ==> q' \ (0::int)" +apply (subgoal_tac "b'*q' < 0") + apply (simp add: mult_less_0_iff, arith) +done + +lemma zdiv_mono2_neg_lemma: + "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; + r < b; 0 \ r'; 0 < b'; b' \ b |] + ==> q' \ (q::int)" +apply (frule q_neg_lemma, assumption+) +apply (subgoal_tac "b*q' < b* (q + 1) ") + apply (simp add: mult_less_cancel_left) +apply (simp add: right_distrib) +apply (subgoal_tac "b*q' \ b'*q'") + prefer 2 apply (simp add: mult_right_mono_neg, arith) +done + +lemma zdiv_mono2_neg: + "[| a < (0::int); 0 < b'; b' \ b |] ==> a div b' \ a div b" +apply (cut_tac a = a and b = b in zmod_zdiv_equality) +apply (cut_tac a = a and b = b' in zmod_zdiv_equality) +apply (rule zdiv_mono2_neg_lemma) +apply (erule subst) +apply (erule subst, simp_all) +done + + +subsubsection{*More Algebraic Laws for div and mod*} + +text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *} + +lemma zmult1_lemma: + "[| divmod_int_rel b c (q, r); c \ 0 |] + ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" +by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) + +lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" +apply (case_tac "c = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_div]) +done + +lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" +apply (case_tac "c = 0", simp) +apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_mod]) +done + +lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" +apply (case_tac "b = 0", simp) +apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) +done + +text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *} + +lemma zadd1_lemma: + "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br); c \ 0 |] + ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" +by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) + +(*NOT suitable for rewriting: the RHS has an instance of the LHS*) +lemma zdiv_zadd1_eq: + "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" +apply (case_tac "c = 0", simp) +apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] divmod_int_rel_div) +done + +instance int :: ring_div +proof + fix a b c :: int + assume not0: "b \ 0" + show "(a + c * b) div b = c + a div b" + unfolding zdiv_zadd1_eq [of a "c * b"] using not0 + by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) +next + fix a b c :: int + assume "a \ 0" + then show "(a * b) div (a * c) = b div c" + proof (cases "b \ 0 \ c \ 0") + case False then show ?thesis by auto + next + case True then have "b \ 0" and "c \ 0" by auto + with `a \ 0` + have "\q r. divmod_int_rel b c (q, r) \ divmod_int_rel (a * b) (a * c) (q, a * r)" + apply (auto simp add: divmod_int_rel_def) + apply (auto simp add: algebra_simps) + apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff mult_commute [of a] mult_less_cancel_right) + done + moreover with `c \ 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto + ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" . + moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp + ultimately show ?thesis by (rule divmod_int_rel_div) + qed +qed auto + +lemma posDivAlg_div_mod: + assumes "k \ 0" + and "l \ 0" + shows "posDivAlg k l = (k div l, k mod l)" +proof (cases "l = 0") + case True then show ?thesis by (simp add: posDivAlg.simps) +next + case False with assms posDivAlg_correct + have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" + by simp + from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + show ?thesis by simp +qed + +lemma negDivAlg_div_mod: + assumes "k < 0" + and "l > 0" + shows "negDivAlg k l = (k div l, k mod l)" +proof - + from assms have "l \ 0" by simp + from assms negDivAlg_correct + have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" + by simp + from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] + show ?thesis by simp +qed + +lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" +by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) + +(* REVISIT: should this be generalized to all semiring_div types? *) +lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] + + +subsubsection{*Proving @{term "a div (b*c) = (a div b) div c"} *} + +(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but + 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems + to cause particular problems.*) + +text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *} + +lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \ 0 |] ==> b*c < b*(q mod c) + r" +apply (subgoal_tac "b * (c - q mod c) < r * 1") + apply (simp add: algebra_simps) +apply (rule order_le_less_trans) + apply (erule_tac [2] mult_strict_right_mono) + apply (rule mult_left_mono_neg) + using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound) + apply (simp) +apply (simp) +done + +lemma zmult2_lemma_aux2: + "[| (0::int) < c; b < r; r \ 0 |] ==> b * (q mod c) + r \ 0" +apply (subgoal_tac "b * (q mod c) \ 0") + apply arith +apply (simp add: mult_le_0_iff) +done + +lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \ r; r < b |] ==> 0 \ b * (q mod c) + r" +apply (subgoal_tac "0 \ b * (q mod c) ") +apply arith +apply (simp add: zero_le_mult_iff) +done + +lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \ r; r < b |] ==> b * (q mod c) + r < b * c" +apply (subgoal_tac "r * 1 < b * (c - q mod c) ") + apply (simp add: right_diff_distrib) +apply (rule order_less_le_trans) + apply (erule mult_strict_right_mono) + apply (rule_tac [2] mult_left_mono) + apply simp + using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound) +apply simp +done + +lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); b \ 0; 0 < c |] + ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" +by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff + zero_less_mult_iff right_distrib [symmetric] + zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) + +lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" +apply (case_tac "b = 0", simp) +apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_div]) +done + +lemma zmod_zmult2_eq: + "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" +apply (case_tac "b = 0", simp) +apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod]) +done + + +subsubsection {*Splitting Rules for div and mod*} + +text{*The proofs of the two lemmas below are essentially identical*} + +lemma split_pos_lemma: + "0 + P(n div k :: int)(n mod k) = (\i j. 0\j & j P i j)" +apply (rule iffI, clarify) + apply (erule_tac P="P ?x ?y" in rev_mp) + apply (subst mod_add_eq) + apply (subst zdiv_zadd1_eq) + apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) +txt{*converse direction*} +apply (drule_tac x = "n div k" in spec) +apply (drule_tac x = "n mod k" in spec, simp) +done + +lemma split_neg_lemma: + "k<0 ==> + P(n div k :: int)(n mod k) = (\i j. k0 & n = k*i + j --> P i j)" +apply (rule iffI, clarify) + apply (erule_tac P="P ?x ?y" in rev_mp) + apply (subst mod_add_eq) + apply (subst zdiv_zadd1_eq) + apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) +txt{*converse direction*} +apply (drule_tac x = "n div k" in spec) +apply (drule_tac x = "n mod k" in spec, simp) +done + +lemma split_zdiv: + "P(n div k :: int) = + ((k = 0 --> P 0) & + (0 (\i j. 0\j & j P i)) & + (k<0 --> (\i j. k0 & n = k*i + j --> P i)))" +apply (case_tac "k=0", simp) +apply (simp only: linorder_neq_iff) +apply (erule disjE) + apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] + split_neg_lemma [of concl: "%x y. P x"]) +done + +lemma split_zmod: + "P(n mod k :: int) = + ((k = 0 --> P n) & + (0 (\i j. 0\j & j P j)) & + (k<0 --> (\i j. k0 & n = k*i + j --> P j)))" +apply (case_tac "k=0", simp) +apply (simp only: linorder_neq_iff) +apply (erule disjE) + apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] + split_neg_lemma [of concl: "%x y. P y"]) +done + +(* Enable arith to deal with div 2 and mod 2: *) +declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split] +declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split] + + +subsubsection{*Speeding up the Division Algorithm with Shifting*} + +text{*computing div by shifting *} + +lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" +proof cases + assume "a=0" + thus ?thesis by simp +next + assume "a\0" and le_a: "0\a" + hence a_pos: "1 \ a" by arith + hence one_less_a2: "1 < 2 * a" by arith + hence le_2a: "2 * (1 + b mod a) \ 2 * a" + unfolding mult_le_cancel_left + by (simp add: add1_zle_eq add_commute [of 1]) + with a_pos have "0 \ b mod a" by simp + hence le_addm: "0 \ 1 mod (2*a) + 2*(b mod a)" + by (simp add: mod_pos_pos_trivial one_less_a2) + with le_2a + have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" + by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 + right_distrib) + thus ?thesis + by (subst zdiv_zadd1_eq, + simp add: mod_mult_mult1 one_less_a2 + div_pos_pos_trivial) +qed + +lemma neg_zdiv_mult_2: "a \ (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" +apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") +apply (rule_tac [2] pos_zdiv_mult_2) +apply (auto simp add: right_diff_distrib) +apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") +apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric]) +apply (simp_all add: algebra_simps) +apply (simp only: ab_diff_minus minus_add_distrib [symmetric] number_of_Min zdiv_zminus_zminus) +done + +lemma zdiv_number_of_Bit0 [simp]: + "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = + number_of v div (number_of w :: int)" +by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric]) + +lemma zdiv_number_of_Bit1 [simp]: + "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = + (if (0::int) \ number_of w + then number_of v div (number_of w) + else (number_of v + (1::int)) div (number_of w))" +apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) +apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric]) +done + + +subsubsection{*Computing mod by Shifting (proofs resemble those for div)*} + +lemma pos_zmod_mult_2: + fixes a b :: int + assumes "0 \ a" + shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" +proof (cases "0 < a") + case False with assms show ?thesis by simp +next + case True + then have "b mod a < a" by (rule pos_mod_bound) + then have "1 + b mod a \ a" by simp + then have A: "2 * (1 + b mod a) \ 2 * a" by simp + from `0 < a` have "0 \ b mod a" by (rule pos_mod_sign) + then have B: "0 \ 1 + 2 * (b mod a)" by simp + have "((1\int) mod ((2\int) * a) + (2\int) * b mod ((2\int) * a)) mod ((2\int) * a) = (1\int) + (2\int) * (b mod a)" + using `0 < a` and A + by (auto simp add: mod_mult_mult1 mod_pos_pos_trivial ring_distribs intro!: mod_pos_pos_trivial B) + then show ?thesis by (subst mod_add_eq) +qed + +lemma neg_zmod_mult_2: + fixes a b :: int + assumes "a \ 0" + shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" +proof - + from assms have "0 \ - a" by auto + then have "(1 + 2 * (- b - 1)) mod (2 * (- a)) = 1 + 2 * ((- b - 1) mod (- a))" + by (rule pos_zmod_mult_2) + then show ?thesis by (simp add: zmod_zminus2 algebra_simps) + (simp add: diff_minus add_ac) +qed + +lemma zmod_number_of_Bit0 [simp]: + "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = + (2::int) * (number_of v mod number_of w)" +apply (simp only: number_of_eq numeral_simps) +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 + neg_zmod_mult_2 add_ac mult_2 [symmetric]) +done + +lemma zmod_number_of_Bit1 [simp]: + "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = + (if (0::int) \ number_of w + then 2 * (number_of v mod number_of w) + 1 + else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" +apply (simp only: number_of_eq numeral_simps) +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 + neg_zmod_mult_2 add_ac mult_2 [symmetric]) +done + + +subsubsection{*Quotients of Signs*} + +lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" +apply (subgoal_tac "a div b \ -1", force) +apply (rule order_trans) +apply (rule_tac a' = "-1" in zdiv_mono1) +apply (auto simp add: div_eq_minus1) +done + +lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" +by (drule zdiv_mono1_neg, auto) + +lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" +by (drule zdiv_mono1, auto) + +lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \ a div b) = (0 \ a)" +apply auto +apply (drule_tac [2] zdiv_mono1) +apply (auto simp add: linorder_neq_iff) +apply (simp (no_asm_use) add: linorder_not_less [symmetric]) +apply (blast intro: div_neg_pos_less0) +done + +lemma neg_imp_zdiv_nonneg_iff: + "b < (0::int) ==> (0 \ a div b) = (a \ (0::int))" +apply (subst zdiv_zminus_zminus [symmetric]) +apply (subst pos_imp_zdiv_nonneg_iff, auto) +done + +(*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) +lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" +by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) + +(*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) +lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" +by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) + + +subsubsection {* The Divides Relation *} + +lemmas zdvd_iff_zmod_eq_0_number_of [simp] = + dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] + +lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" + by (rule dvd_mod) (* TODO: remove *) + +lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" + by (rule dvd_mod_imp_dvd) (* TODO: remove *) + +lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" + using zmod_zdiv_equality[where a="m" and b="n"] + by (simp add: algebra_simps) + +lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" +apply (induct "y", auto) +apply (rule zmod_zmult1_eq [THEN trans]) +apply (simp (no_asm_simp)) +apply (rule mod_mult_eq [symmetric]) +done + +lemma zdiv_int: "int (a div b) = (int a) div (int b)" +apply (subst split_div, auto) +apply (subst split_zdiv, auto) +apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in unique_quotient) +apply (auto simp add: divmod_int_rel_def of_nat_mult) +done + +lemma zmod_int: "int (a mod b) = (int a) mod (int b)" +apply (subst split_mod, auto) +apply (subst split_zmod, auto) +apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia + in unique_remainder) +apply (auto simp add: divmod_int_rel_def of_nat_mult) +done + +lemma abs_div: "(y::int) dvd x \ abs (x div y) = abs x div abs y" +by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) + +lemma zdvd_mult_div_cancel:"(n::int) dvd m \ n * (m div n) = m" +apply (subgoal_tac "m mod n = 0") + apply (simp add: zmult_div_cancel) +apply (simp only: dvd_eq_mod_eq_0) +done + +text{*Suggested by Matthias Daum*} +lemma int_power_div_base: + "\0 < m; 0 < k\ \ k ^ m div k = (k::int) ^ (m - Suc 0)" +apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") + apply (erule ssubst) + apply (simp only: power_add) + apply simp_all +done + +text {* by Brian Huffman *} +lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" +by (rule mod_minus_eq [symmetric]) + +lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" +by (rule mod_diff_left_eq [symmetric]) + +lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" +by (rule mod_diff_right_eq [symmetric]) + +lemmas zmod_simps = + mod_add_left_eq [symmetric] + mod_add_right_eq [symmetric] + zmod_zmult1_eq [symmetric] + mod_mult_left_eq [symmetric] + zpower_zmod + zminus_zmod zdiff_zmod_left zdiff_zmod_right + +text {* Distributive laws for function @{text nat}. *} + +lemma nat_div_distrib: "0 \ x \ nat (x div y) = nat x div nat y" +apply (rule linorder_cases [of y 0]) +apply (simp add: div_nonneg_neg_le0) +apply simp +apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) +done + +(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) +lemma nat_mod_distrib: + "\0 \ x; 0 \ y\ \ nat (x mod y) = nat x mod nat y" +apply (case_tac "y = 0", simp) +apply (simp add: nat_eq_iff zmod_int) +done + +text {* transfer setup *} + +lemma transfer_nat_int_functions: + "(x::int) >= 0 \ y >= 0 \ (nat x) div (nat y) = nat (x div y)" + "(x::int) >= 0 \ y >= 0 \ (nat x) mod (nat y) = nat (x mod y)" + by (auto simp add: nat_div_distrib nat_mod_distrib) + +lemma transfer_nat_int_function_closures: + "(x::int) >= 0 \ y >= 0 \ x div y >= 0" + "(x::int) >= 0 \ y >= 0 \ x mod y >= 0" + apply (cases "y = 0") + apply (auto simp add: pos_imp_zdiv_nonneg_iff) + apply (cases "y = 0") + apply auto +done + +declare TransferMorphism_nat_int [transfer add return: + transfer_nat_int_functions + transfer_nat_int_function_closures +] + +lemma transfer_int_nat_functions: + "(int x) div (int y) = int (x div y)" + "(int x) mod (int y) = int (x mod y)" + by (auto simp add: zdiv_int zmod_int) + +lemma transfer_int_nat_function_closures: + "is_nat x \ is_nat y \ is_nat (x div y)" + "is_nat x \ is_nat y \ is_nat (x mod y)" + by (simp_all only: is_nat_def transfer_nat_int_function_closures) + +declare TransferMorphism_int_nat [transfer add return: + transfer_int_nat_functions + transfer_int_nat_function_closures +] + +text{*Suggested by Matthias Daum*} +lemma int_div_less_self: "\0 < x; 1 < k\ \ x div k < (x::int)" +apply (subgoal_tac "nat x div nat k < nat x") + apply (simp (asm_lr) add: nat_div_distrib [symmetric]) +apply (rule Divides.div_less_dividend, simp_all) +done + +text {* code generator setup *} + +context ring_1 +begin + +lemma of_int_num [code]: + "of_int k = (if k = 0 then 0 else if k < 0 then + - of_int (- k) else let + (l, m) = divmod_int k 2; + l' = of_int l + in if m = 0 then l' + l' else l' + l' + 1)" +proof - + have aux1: "k mod (2\int) \ (0\int) \ + of_int k = of_int (k div 2 * 2 + 1)" + proof - + have "k mod 2 < 2" by (auto intro: pos_mod_bound) + moreover have "0 \ k mod 2" by (auto intro: pos_mod_sign) + moreover assume "k mod 2 \ 0" + ultimately have "k mod 2 = 1" by arith + moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp + ultimately show ?thesis by auto + qed + have aux2: "\x. of_int 2 * x = x + x" + proof - + fix x + have int2: "(2::int) = 1 + 1" by arith + show "of_int 2 * x = x + x" + unfolding int2 of_int_add left_distrib by simp + qed + have aux3: "\x. x * of_int 2 = x + x" + proof - + fix x + have int2: "(2::int) = 1 + 1" by arith + show "x * of_int 2 = x + x" + unfolding int2 of_int_add right_distrib by simp + qed + from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3) +qed + +end + +lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \ n dvd x - y" +proof + assume H: "x mod n = y mod n" + hence "x mod n - y mod n = 0" by simp + hence "(x mod n - y mod n) mod n = 0" by simp + hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) + thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) +next + assume H: "n dvd x - y" + then obtain k where k: "x-y = n*k" unfolding dvd_def by blast + hence "x = n*k + y" by simp + hence "x mod n = (n*k + y) mod n" by simp + thus "x mod n = y mod n" by (simp add: mod_add_left_eq) +qed + +lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \ x" + shows "\q. x = y + n * q" +proof- + from xy have th: "int x - int y = int (x - y)" by simp + from xyn have "int x mod int n = int y mod int n" + by (simp add: zmod_int[symmetric]) + hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) + hence "n dvd x - y" by (simp add: th zdvd_int) + then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith +qed + +lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" + (is "?lhs = ?rhs") +proof + assume H: "x mod n = y mod n" + {assume xy: "x \ y" + from H have th: "y mod n = x mod n" by simp + from nat_mod_eq_lemma[OF th xy] have ?rhs + apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} + moreover + {assume xy: "y \ x" + from nat_mod_eq_lemma[OF H xy] have ?rhs + apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} + ultimately show ?rhs using linear[of x y] by blast +next + assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast + hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp + thus ?lhs by simp +qed + +lemma div_nat_number_of [simp]: + "(number_of v :: nat) div number_of v' = + (if neg (number_of v :: int) then 0 + else nat (number_of v div number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_div_distrib) + +lemma one_div_nat_number_of [simp]: + "Suc 0 div number_of v' = nat (1 div number_of v')" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + +lemma mod_nat_number_of [simp]: + "(number_of v :: nat) mod number_of v' = + (if neg (number_of v :: int) then 0 + else if neg (number_of v' :: int) then number_of v + else nat (number_of v mod number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_mod_distrib) + +lemma one_mod_nat_number_of [simp]: + "Suc 0 mod number_of v' = + (if neg (number_of v' :: int) then Suc 0 + else nat (1 mod number_of v'))" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + +lemmas dvd_eq_mod_eq_0_number_of = + dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] + +declare dvd_eq_mod_eq_0_number_of [simp] + + +subsubsection {* Code generation *} + +definition pdivmod :: "int \ int \ int \ int" where + "pdivmod k l = (\k\ div \l\, \k\ mod \l\)" + +lemma pdivmod_posDivAlg [code]: + "pdivmod k l = (if l = 0 then (0, \k\) else posDivAlg \k\ \l\)" +by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) + +lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else + apsnd ((op *) (sgn l)) (if 0 < l \ 0 \ k \ l < 0 \ k < 0 + then pdivmod k l + else (let (r, s) = pdivmod k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have aux: "\q::int. - k = l * q \ k = l * - q" by auto + show ?thesis + by (simp add: divmod_int_mod_div pdivmod_def) + (auto simp add: aux not_less not_le zdiv_zminus1_eq_if + zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) +qed + +lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else + apsnd ((op *) (sgn l)) (if sgn k = sgn l + then pdivmod k l + else (let (r, s) = pdivmod k l in + if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" +proof - + have "k \ 0 \ l \ 0 \ 0 < l \ 0 \ k \ l < 0 \ k < 0 \ sgn k = sgn l" + by (auto simp add: not_less sgn_if) + then show ?thesis by (simp add: divmod_int_pdivmod) +qed + +end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Groebner_Basis.thy Fri Oct 30 14:02:42 2009 +0100 @@ -5,7 +5,7 @@ header {* Semiring normalization and Groebner Bases *} theory Groebner_Basis -imports IntDiv +imports Numeral_Simprocs uses "Tools/Groebner_Basis/misc.ML" "Tools/Groebner_Basis/normalizer_data.ML" diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/IntDiv.thy --- a/src/HOL/IntDiv.thy Fri Oct 30 11:31:34 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1474 +0,0 @@ -(* Title: HOL/IntDiv.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge -*) - -header{* The Division Operators div and mod *} - -theory IntDiv -imports Int Divides FunDef -uses - "~~/src/Provers/Arith/assoc_fold.ML" - "~~/src/Provers/Arith/cancel_numerals.ML" - "~~/src/Provers/Arith/combine_numerals.ML" - "~~/src/Provers/Arith/cancel_numeral_factor.ML" - "~~/src/Provers/Arith/extract_common_term.ML" - ("Tools/numeral_simprocs.ML") - ("Tools/nat_numeral_simprocs.ML") -begin - -definition divmod_int_rel :: "int \ int \ int \ int \ bool" where - --{*definition of quotient and remainder*} - [code]: "divmod_int_rel a b = (\(q, r). a = b * q + r \ - (if 0 < b then 0 \ r \ r < b else b < r \ r \ 0))" - -definition adjust :: "int \ int \ int \ int \ int" where - --{*for the division algorithm*} - [code]: "adjust b = (\(q, r). if 0 \ r - b then (2 * q + 1, r - b) - else (2 * q, r))" - -text{*algorithm for the case @{text "a\0, b>0"}*} -function posDivAlg :: "int \ int \ int \ int" where - "posDivAlg a b = (if a < b \ b \ 0 then (0, a) - else adjust b (posDivAlg a (2 * b)))" -by auto -termination by (relation "measure (\(a, b). nat (a - b + 1))") - (auto simp add: mult_2) - -text{*algorithm for the case @{text "a<0, b>0"}*} -function negDivAlg :: "int \ int \ int \ int" where - "negDivAlg a b = (if 0 \a + b \ b \ 0 then (-1, a + b) - else adjust b (negDivAlg a (2 * b)))" -by auto -termination by (relation "measure (\(a, b). nat (- a - b))") - (auto simp add: mult_2) - -text{*algorithm for the general case @{term "b\0"}*} -definition negateSnd :: "int \ int \ int \ int" where - [code_unfold]: "negateSnd = apsnd uminus" - -definition divmod_int :: "int \ int \ int \ int" where - --{*The full division algorithm considers all possible signs for a, b - including the special case @{text "a=0, b<0"} because - @{term negDivAlg} requires @{term "a<0"}.*} - "divmod_int a b = (if 0 \ a then if 0 \ b then posDivAlg a b - else if a = 0 then (0, 0) - else negateSnd (negDivAlg (-a) (-b)) - else - if 0 < b then negDivAlg a b - else negateSnd (posDivAlg (-a) (-b)))" - -instantiation int :: Divides.div -begin - -definition - "a div b = fst (divmod_int a b)" - -definition - "a mod b = snd (divmod_int a b)" - -instance .. - -end - -lemma divmod_int_mod_div: - "divmod_int p q = (p div q, p mod q)" - by (auto simp add: div_int_def mod_int_def) - -text{* -Here is the division algorithm in ML: - -\begin{verbatim} - fun posDivAlg (a,b) = - if ar-b then (2*q+1, r-b) else (2*q, r) - end - - fun negDivAlg (a,b) = - if 0\a+b then (~1,a+b) - else let val (q,r) = negDivAlg(a, 2*b) - in if 0\r-b then (2*q+1, r-b) else (2*q, r) - end; - - fun negateSnd (q,r:int) = (q,~r); - - fun divmod (a,b) = if 0\a then - if b>0 then posDivAlg (a,b) - else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) - else - if 0 b*q + r; 0 \ r'; r' < b; r < b |] - ==> q' \ (q::int)" -apply (subgoal_tac "r' + b * (q'-q) \ r") - prefer 2 apply (simp add: right_diff_distrib) -apply (subgoal_tac "0 < b * (1 + q - q') ") -apply (erule_tac [2] order_le_less_trans) - prefer 2 apply (simp add: right_diff_distrib right_distrib) -apply (subgoal_tac "b * q' < b * (1 + q) ") - prefer 2 apply (simp add: right_diff_distrib right_distrib) -apply (simp add: mult_less_cancel_left) -done - -lemma unique_quotient_lemma_neg: - "[| b*q' + r' \ b*q + r; r \ 0; b < r; b < r' |] - ==> q \ (q'::int)" -by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, - auto) - -lemma unique_quotient: - "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] - ==> q = q'" -apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) -apply (blast intro: order_antisym - dest: order_eq_refl [THEN unique_quotient_lemma] - order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ -done - - -lemma unique_remainder: - "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \ 0 |] - ==> r = r'" -apply (subgoal_tac "q = q'") - apply (simp add: divmod_int_rel_def) -apply (blast intro: unique_quotient) -done - - -subsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*} - -text{*And positive divisors*} - -lemma adjust_eq [simp]: - "adjust b (q,r) = - (let diff = r-b in - if 0 \ diff then (2*q + 1, diff) - else (2*q, r))" -by (simp add: Let_def adjust_def) - -declare posDivAlg.simps [simp del] - -text{*use with a simproc to avoid repeatedly proving the premise*} -lemma posDivAlg_eqn: - "0 < b ==> - posDivAlg a b = (if a a" and "0 < b" - shows "divmod_int_rel a b (posDivAlg a b)" -using prems apply (induct a b rule: posDivAlg.induct) -apply auto -apply (simp add: divmod_int_rel_def) -apply (subst posDivAlg_eqn, simp add: right_distrib) -apply (case_tac "a < b") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) -done - - -subsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*} - -text{*And positive divisors*} - -declare negDivAlg.simps [simp del] - -text{*use with a simproc to avoid repeatedly proving the premise*} -lemma negDivAlg_eqn: - "0 < b ==> - negDivAlg a b = - (if 0\a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" -by (rule negDivAlg.simps [THEN trans], simp) - -(*Correctness of negDivAlg: it computes quotients correctly - It doesn't work if a=0 because the 0/b equals 0, not -1*) -lemma negDivAlg_correct: - assumes "a < 0" and "b > 0" - shows "divmod_int_rel a b (negDivAlg a b)" -using prems apply (induct a b rule: negDivAlg.induct) -apply (auto simp add: linorder_not_le) -apply (simp add: divmod_int_rel_def) -apply (subst negDivAlg_eqn, assumption) -apply (case_tac "a + b < (0\int)") -apply simp_all -apply (erule splitE) -apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) -done - - -subsection{*Existence Shown by Proving the Division Algorithm to be Correct*} - -(*the case a=0*) -lemma divmod_int_rel_0: "b \ 0 ==> divmod_int_rel 0 b (0, 0)" -by (auto simp add: divmod_int_rel_def linorder_neq_iff) - -lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" -by (subst posDivAlg.simps, auto) - -lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" -by (subst negDivAlg.simps, auto) - -lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" -by (simp add: negateSnd_def) - -lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)" -by (auto simp add: split_ifs divmod_int_rel_def) - -lemma divmod_int_correct: "b \ 0 ==> divmod_int_rel a b (divmod_int a b)" -by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg - posDivAlg_correct negDivAlg_correct) - -text{*Arbitrary definitions for division by zero. Useful to simplify - certain equations.*} - -lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" -by (simp add: div_int_def mod_int_def divmod_int_def posDivAlg.simps) - - -text{*Basic laws about division and remainder*} - -lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" -apply (case_tac "b = 0", simp) -apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) -done - -lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" -by(simp add: zmod_zdiv_equality[symmetric]) - -lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" -by(simp add: mult_commute zmod_zdiv_equality[symmetric]) - -text {* Tool setup *} - -ML {* -local - -fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; - -fun find_first_numeral past (t::terms) = - ((snd (HOLogic.dest_number t), rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; - -fun mk_minus t = - let val T = Term.fastype_of t - in Const (@{const_name HOL.uminus}, T --> T) $ t end; - -(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) -fun mk_sum T [] = mk_number T 0 - | mk_sum T [t,u] = mk_plus (t, u) - | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum T [] = mk_number T 0 - | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT; - -(*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (not pos, u, ts)) - | dest_summing (pos, t, ts) = - if pos then t::ts else mk_minus t :: ts; - -fun dest_sum t = dest_summing (true, t, []); - -structure CancelDivMod = CancelDivModFun(struct - - val div_name = @{const_name div}; - val mod_name = @{const_name mod}; - val mk_binop = HOLogic.mk_binop; - val mk_sum = mk_sum HOLogic.intT; - val dest_sum = dest_sum; - - val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; - - val trans = trans; - - val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac - (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) - -end) - -in - -val cancel_div_mod_int_proc = Simplifier.simproc @{theory} - "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); - -val _ = Addsimprocs [cancel_div_mod_int_proc]; - -end -*} - -lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" -apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def mod_int_def) -done - -lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] - and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] - -lemma neg_mod_conj : "b < (0::int) ==> a mod b \ 0 & b < a mod b" -apply (cut_tac a = a and b = b in divmod_int_correct) -apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) -done - -lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] - and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] - - - -subsection{*General Properties of div and mod*} - -lemma divmod_int_rel_div_mod: "b \ 0 ==> divmod_int_rel a b (a div b, a mod b)" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (force simp add: divmod_int_rel_def linorder_neq_iff) -done - -lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a div b = q" -by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) - -lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r); b \ 0 |] ==> a mod b = r" -by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) - -lemma div_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a div b = 0" -apply (rule divmod_int_rel_div) -apply (auto simp add: divmod_int_rel_def) -done - -lemma div_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a div b = 0" -apply (rule divmod_int_rel_div) -apply (auto simp add: divmod_int_rel_def) -done - -lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a div b = -1" -apply (rule divmod_int_rel_div) -apply (auto simp add: divmod_int_rel_def) -done - -(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) - -lemma mod_pos_pos_trivial: "[| (0::int) \ a; a < b |] ==> a mod b = a" -apply (rule_tac q = 0 in divmod_int_rel_mod) -apply (auto simp add: divmod_int_rel_def) -done - -lemma mod_neg_neg_trivial: "[| a \ (0::int); b < a |] ==> a mod b = a" -apply (rule_tac q = 0 in divmod_int_rel_mod) -apply (auto simp add: divmod_int_rel_def) -done - -lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \ 0 |] ==> a mod b = a+b" -apply (rule_tac q = "-1" in divmod_int_rel_mod) -apply (auto simp add: divmod_int_rel_def) -done - -text{*There is no @{text mod_neg_pos_trivial}.*} - - -(*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) -lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" -apply (case_tac "b = 0", simp) -apply (simp add: divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, - THEN divmod_int_rel_div, THEN sym]) - -done - -(*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) -lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" -apply (case_tac "b = 0", simp) -apply (subst divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, THEN divmod_int_rel_mod], - auto) -done - - -subsection{*Laws for div and mod with Unary Minus*} - -lemma zminus1_lemma: - "divmod_int_rel a b (q, r) - ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, - if r=0 then 0 else b-r)" -by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib) - - -lemma zdiv_zminus1_eq_if: - "b \ (0::int) - ==> (-a) div b = - (if a mod b = 0 then - (a div b) else - (a div b) - 1)" -by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_div]) - -lemma zmod_zminus1_eq_if: - "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" -apply (case_tac "b = 0", simp) -apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_mod]) -done - -lemma zmod_zminus1_not_zero: - fixes k l :: int - shows "- k mod l \ 0 \ k mod l \ 0" - unfolding zmod_zminus1_eq_if by auto - -lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" -by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) - -lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" -by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) - -lemma zdiv_zminus2_eq_if: - "b \ (0::int) - ==> a div (-b) = - (if a mod b = 0 then - (a div b) else - (a div b) - 1)" -by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) - -lemma zmod_zminus2_eq_if: - "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" -by (simp add: zmod_zminus1_eq_if zmod_zminus2) - -lemma zmod_zminus2_not_zero: - fixes k l :: int - shows "k mod - l \ 0 \ k mod l \ 0" - unfolding zmod_zminus2_eq_if by auto - - -subsection{*Division of a Number by Itself*} - -lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \ q" -apply (subgoal_tac "0 < a*q") - apply (simp add: zero_less_mult_iff, arith) -done - -lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \ r |] ==> q \ 1" -apply (subgoal_tac "0 \ a* (1-q) ") - apply (simp add: zero_le_mult_iff) -apply (simp add: right_diff_distrib) -done - -lemma self_quotient: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> q = 1" -apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff) -apply (rule order_antisym, safe, simp_all) -apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) -apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) -apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ -done - -lemma self_remainder: "[| divmod_int_rel a a (q, r); a \ (0::int) |] ==> r = 0" -apply (frule self_quotient, assumption) -apply (simp add: divmod_int_rel_def) -done - -lemma zdiv_self [simp]: "a \ 0 ==> a div a = (1::int)" -by (simp add: divmod_int_rel_div_mod [THEN self_quotient]) - -(*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) -lemma zmod_self [simp]: "a mod a = (0::int)" -apply (case_tac "a = 0", simp) -apply (simp add: divmod_int_rel_div_mod [THEN self_remainder]) -done - - -subsection{*Computation of Division and Remainder*} - -lemma zdiv_zero [simp]: "(0::int) div b = 0" -by (simp add: div_int_def divmod_int_def) - -lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" -by (simp add: div_int_def divmod_int_def) - -lemma zmod_zero [simp]: "(0::int) mod b = 0" -by (simp add: mod_int_def divmod_int_def) - -lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" -by (simp add: mod_int_def divmod_int_def) - -text{*a positive, b positive *} - -lemma div_pos_pos: "[| 0 < a; 0 \ b |] ==> a div b = fst (posDivAlg a b)" -by (simp add: div_int_def divmod_int_def) - -lemma mod_pos_pos: "[| 0 < a; 0 \ b |] ==> a mod b = snd (posDivAlg a b)" -by (simp add: mod_int_def divmod_int_def) - -text{*a negative, b positive *} - -lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" -by (simp add: div_int_def divmod_int_def) - -lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" -by (simp add: mod_int_def divmod_int_def) - -text{*a positive, b negative *} - -lemma div_pos_neg: - "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" -by (simp add: div_int_def divmod_int_def) - -lemma mod_pos_neg: - "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" -by (simp add: mod_int_def divmod_int_def) - -text{*a negative, b negative *} - -lemma div_neg_neg: - "[| a < 0; b \ 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" -by (simp add: div_int_def divmod_int_def) - -lemma mod_neg_neg: - "[| a < 0; b \ 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" -by (simp add: mod_int_def divmod_int_def) - -text {*Simplify expresions in which div and mod combine numerical constants*} - -lemma divmod_int_relI: - "\a == b * q + r; if 0 < b then 0 \ r \ r < b else b < r \ r \ 0\ - \ divmod_int_rel a b (q, r)" - unfolding divmod_int_rel_def by simp - -lemmas divmod_int_rel_div_eq = divmod_int_relI [THEN divmod_int_rel_div, THEN eq_reflection] -lemmas divmod_int_rel_mod_eq = divmod_int_relI [THEN divmod_int_rel_mod, THEN eq_reflection] -lemmas arithmetic_simps = - arith_simps - add_special - OrderedGroup.add_0_left - OrderedGroup.add_0_right - mult_zero_left - mult_zero_right - mult_1_left - mult_1_right - -(* simprocs adapted from HOL/ex/Binary.thy *) -ML {* -local - val mk_number = HOLogic.mk_number HOLogic.intT; - fun mk_cert u k l = @{term "plus :: int \ int \ int"} $ - (@{term "times :: int \ int \ int"} $ u $ mk_number k) $ - mk_number l; - fun prove ctxt prop = Goal.prove ctxt [] [] prop - (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps})))); - fun binary_proc proc ss ct = - (case Thm.term_of ct of - _ $ t $ u => - (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of - SOME args => proc (Simplifier.the_context ss) args - | NONE => NONE) - | _ => NONE); -in - fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => - if n = 0 then NONE - else let val (k, l) = Integer.div_mod m n; - in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); -end -*} - -simproc_setup binary_int_div ("number_of m div number_of n :: int") = - {* K (divmod_proc (@{thm divmod_int_rel_div_eq})) *} - -simproc_setup binary_int_mod ("number_of m mod number_of n :: int") = - {* K (divmod_proc (@{thm divmod_int_rel_mod_eq})) *} - -lemmas posDivAlg_eqn_number_of [simp] = - posDivAlg_eqn [of "number_of v" "number_of w", standard] - -lemmas negDivAlg_eqn_number_of [simp] = - negDivAlg_eqn [of "number_of v" "number_of w", standard] - - -text{*Special-case simplification *} - -lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" -apply (cut_tac a = a and b = "-1" in neg_mod_sign) -apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) -apply (auto simp del: neg_mod_sign neg_mod_bound) -done - -lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" -by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) - -(** The last remaining special cases for constant arithmetic: - 1 div z and 1 mod z **) - -lemmas div_pos_pos_1_number_of [simp] = - div_pos_pos [OF int_0_less_1, of "number_of w", standard] - -lemmas div_pos_neg_1_number_of [simp] = - div_pos_neg [OF int_0_less_1, of "number_of w", standard] - -lemmas mod_pos_pos_1_number_of [simp] = - mod_pos_pos [OF int_0_less_1, of "number_of w", standard] - -lemmas mod_pos_neg_1_number_of [simp] = - mod_pos_neg [OF int_0_less_1, of "number_of w", standard] - - -lemmas posDivAlg_eqn_1_number_of [simp] = - posDivAlg_eqn [of concl: 1 "number_of w", standard] - -lemmas negDivAlg_eqn_1_number_of [simp] = - negDivAlg_eqn [of concl: 1 "number_of w", standard] - - - -subsection{*Monotonicity in the First Argument (Dividend)*} - -lemma zdiv_mono1: "[| a \ a'; 0 < (b::int) |] ==> a div b \ a' div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a' and b = b in zmod_zdiv_equality) -apply (rule unique_quotient_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - -lemma zdiv_mono1_neg: "[| a \ a'; (b::int) < 0 |] ==> a' div b \ a div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a' and b = b in zmod_zdiv_equality) -apply (rule unique_quotient_lemma_neg) -apply (erule subst) -apply (erule subst, simp_all) -done - - -subsection{*Monotonicity in the Second Argument (Divisor)*} - -lemma q_pos_lemma: - "[| 0 \ b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \ (q'::int)" -apply (subgoal_tac "0 < b'* (q' + 1) ") - apply (simp add: zero_less_mult_iff) -apply (simp add: right_distrib) -done - -lemma zdiv_mono2_lemma: - "[| b*q + r = b'*q' + r'; 0 \ b'*q' + r'; - r' < b'; 0 \ r; 0 < b'; b' \ b |] - ==> q \ (q'::int)" -apply (frule q_pos_lemma, assumption+) -apply (subgoal_tac "b*q < b* (q' + 1) ") - apply (simp add: mult_less_cancel_left) -apply (subgoal_tac "b*q = r' - r + b'*q'") - prefer 2 apply simp -apply (simp (no_asm_simp) add: right_distrib) -apply (subst add_commute, rule zadd_zless_mono, arith) -apply (rule mult_right_mono, auto) -done - -lemma zdiv_mono2: - "[| (0::int) \ a; 0 < b'; b' \ b |] ==> a div b \ a div b'" -apply (subgoal_tac "b \ 0") - prefer 2 apply arith -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a and b = b' in zmod_zdiv_equality) -apply (rule zdiv_mono2_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - -lemma q_neg_lemma: - "[| b'*q' + r' < 0; 0 \ r'; 0 < b' |] ==> q' \ (0::int)" -apply (subgoal_tac "b'*q' < 0") - apply (simp add: mult_less_0_iff, arith) -done - -lemma zdiv_mono2_neg_lemma: - "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; - r < b; 0 \ r'; 0 < b'; b' \ b |] - ==> q' \ (q::int)" -apply (frule q_neg_lemma, assumption+) -apply (subgoal_tac "b*q' < b* (q + 1) ") - apply (simp add: mult_less_cancel_left) -apply (simp add: right_distrib) -apply (subgoal_tac "b*q' \ b'*q'") - prefer 2 apply (simp add: mult_right_mono_neg, arith) -done - -lemma zdiv_mono2_neg: - "[| a < (0::int); 0 < b'; b' \ b |] ==> a div b' \ a div b" -apply (cut_tac a = a and b = b in zmod_zdiv_equality) -apply (cut_tac a = a and b = b' in zmod_zdiv_equality) -apply (rule zdiv_mono2_neg_lemma) -apply (erule subst) -apply (erule subst, simp_all) -done - - -subsection{*More Algebraic Laws for div and mod*} - -text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *} - -lemma zmult1_lemma: - "[| divmod_int_rel b c (q, r); c \ 0 |] - ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" -by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) - -lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" -apply (case_tac "c = 0", simp) -apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_div]) -done - -lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" -apply (case_tac "c = 0", simp) -apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_mod]) -done - -lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" -apply (case_tac "b = 0", simp) -apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) -done - -text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *} - -lemma zadd1_lemma: - "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br); c \ 0 |] - ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" -by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) - -(*NOT suitable for rewriting: the RHS has an instance of the LHS*) -lemma zdiv_zadd1_eq: - "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" -apply (case_tac "c = 0", simp) -apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] divmod_int_rel_div) -done - -instance int :: ring_div -proof - fix a b c :: int - assume not0: "b \ 0" - show "(a + c * b) div b = c + a div b" - unfolding zdiv_zadd1_eq [of a "c * b"] using not0 - by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) -next - fix a b c :: int - assume "a \ 0" - then show "(a * b) div (a * c) = b div c" - proof (cases "b \ 0 \ c \ 0") - case False then show ?thesis by auto - next - case True then have "b \ 0" and "c \ 0" by auto - with `a \ 0` - have "\q r. divmod_int_rel b c (q, r) \ divmod_int_rel (a * b) (a * c) (q, a * r)" - apply (auto simp add: divmod_int_rel_def) - apply (auto simp add: algebra_simps) - apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff mult_commute [of a] mult_less_cancel_right) - done - moreover with `c \ 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto - ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" . - moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp - ultimately show ?thesis by (rule divmod_int_rel_div) - qed -qed auto - -lemma posDivAlg_div_mod: - assumes "k \ 0" - and "l \ 0" - shows "posDivAlg k l = (k div l, k mod l)" -proof (cases "l = 0") - case True then show ?thesis by (simp add: posDivAlg.simps) -next - case False with assms posDivAlg_correct - have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" - by simp - from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] - show ?thesis by simp -qed - -lemma negDivAlg_div_mod: - assumes "k < 0" - and "l > 0" - shows "negDivAlg k l = (k div l, k mod l)" -proof - - from assms have "l \ 0" by simp - from assms negDivAlg_correct - have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" - by simp - from divmod_int_rel_div [OF this `l \ 0`] divmod_int_rel_mod [OF this `l \ 0`] - show ?thesis by simp -qed - -lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" -by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) - -(* REVISIT: should this be generalized to all semiring_div types? *) -lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] - - -subsection{*Proving @{term "a div (b*c) = (a div b) div c"} *} - -(*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but - 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems - to cause particular problems.*) - -text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *} - -lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \ 0 |] ==> b*c < b*(q mod c) + r" -apply (subgoal_tac "b * (c - q mod c) < r * 1") - apply (simp add: algebra_simps) -apply (rule order_le_less_trans) - apply (erule_tac [2] mult_strict_right_mono) - apply (rule mult_left_mono_neg) - using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps pos_mod_bound) - apply (simp) -apply (simp) -done - -lemma zmult2_lemma_aux2: - "[| (0::int) < c; b < r; r \ 0 |] ==> b * (q mod c) + r \ 0" -apply (subgoal_tac "b * (q mod c) \ 0") - apply arith -apply (simp add: mult_le_0_iff) -done - -lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \ r; r < b |] ==> 0 \ b * (q mod c) + r" -apply (subgoal_tac "0 \ b * (q mod c) ") -apply arith -apply (simp add: zero_le_mult_iff) -done - -lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \ r; r < b |] ==> b * (q mod c) + r < b * c" -apply (subgoal_tac "r * 1 < b * (c - q mod c) ") - apply (simp add: right_diff_distrib) -apply (rule order_less_le_trans) - apply (erule mult_strict_right_mono) - apply (rule_tac [2] mult_left_mono) - apply simp - using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps pos_mod_bound) -apply simp -done - -lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); b \ 0; 0 < c |] - ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" -by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff - zero_less_mult_iff right_distrib [symmetric] - zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) - -lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" -apply (case_tac "b = 0", simp) -apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_div]) -done - -lemma zmod_zmult2_eq: - "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" -apply (case_tac "b = 0", simp) -apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod]) -done - - -subsection {*Splitting Rules for div and mod*} - -text{*The proofs of the two lemmas below are essentially identical*} - -lemma split_pos_lemma: - "0 - P(n div k :: int)(n mod k) = (\i j. 0\j & j P i j)" -apply (rule iffI, clarify) - apply (erule_tac P="P ?x ?y" in rev_mp) - apply (subst mod_add_eq) - apply (subst zdiv_zadd1_eq) - apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) -txt{*converse direction*} -apply (drule_tac x = "n div k" in spec) -apply (drule_tac x = "n mod k" in spec, simp) -done - -lemma split_neg_lemma: - "k<0 ==> - P(n div k :: int)(n mod k) = (\i j. k0 & n = k*i + j --> P i j)" -apply (rule iffI, clarify) - apply (erule_tac P="P ?x ?y" in rev_mp) - apply (subst mod_add_eq) - apply (subst zdiv_zadd1_eq) - apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) -txt{*converse direction*} -apply (drule_tac x = "n div k" in spec) -apply (drule_tac x = "n mod k" in spec, simp) -done - -lemma split_zdiv: - "P(n div k :: int) = - ((k = 0 --> P 0) & - (0 (\i j. 0\j & j P i)) & - (k<0 --> (\i j. k0 & n = k*i + j --> P i)))" -apply (case_tac "k=0", simp) -apply (simp only: linorder_neq_iff) -apply (erule disjE) - apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] - split_neg_lemma [of concl: "%x y. P x"]) -done - -lemma split_zmod: - "P(n mod k :: int) = - ((k = 0 --> P n) & - (0 (\i j. 0\j & j P j)) & - (k<0 --> (\i j. k0 & n = k*i + j --> P j)))" -apply (case_tac "k=0", simp) -apply (simp only: linorder_neq_iff) -apply (erule disjE) - apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] - split_neg_lemma [of concl: "%x y. P y"]) -done - -(* Enable arith to deal with div 2 and mod 2: *) -declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split] -declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split] - - -subsection{*Speeding up the Division Algorithm with Shifting*} - -text{*computing div by shifting *} - -lemma pos_zdiv_mult_2: "(0::int) \ a ==> (1 + 2*b) div (2*a) = b div a" -proof cases - assume "a=0" - thus ?thesis by simp -next - assume "a\0" and le_a: "0\a" - hence a_pos: "1 \ a" by arith - hence one_less_a2: "1 < 2 * a" by arith - hence le_2a: "2 * (1 + b mod a) \ 2 * a" - unfolding mult_le_cancel_left - by (simp add: add1_zle_eq add_commute [of 1]) - with a_pos have "0 \ b mod a" by simp - hence le_addm: "0 \ 1 mod (2*a) + 2*(b mod a)" - by (simp add: mod_pos_pos_trivial one_less_a2) - with le_2a - have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" - by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 - right_distrib) - thus ?thesis - by (subst zdiv_zadd1_eq, - simp add: mod_mult_mult1 one_less_a2 - div_pos_pos_trivial) -qed - -lemma neg_zdiv_mult_2: "a \ (0::int) ==> (1 + 2*b) div (2*a) = (b+1) div a" -apply (subgoal_tac " (1 + 2* (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a) ") -apply (rule_tac [2] pos_zdiv_mult_2) -apply (auto simp add: right_diff_distrib) -apply (subgoal_tac " (-1 - (2 * b)) = - (1 + (2 * b))") -apply (simp only: zdiv_zminus_zminus diff_minus minus_add_distrib [symmetric]) -apply (simp_all add: algebra_simps) -apply (simp only: ab_diff_minus minus_add_distrib [symmetric] number_of_Min zdiv_zminus_zminus) -done - -lemma zdiv_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = - number_of v div (number_of w :: int)" -by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric]) - -lemma zdiv_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then number_of v div (number_of w) - else (number_of v + (1::int)) div (number_of w))" -apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) -apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric]) -done - - -subsection{*Computing mod by Shifting (proofs resemble those for div)*} - -lemma pos_zmod_mult_2: - fixes a b :: int - assumes "0 \ a" - shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" -proof (cases "0 < a") - case False with assms show ?thesis by simp -next - case True - then have "b mod a < a" by (rule pos_mod_bound) - then have "1 + b mod a \ a" by simp - then have A: "2 * (1 + b mod a) \ 2 * a" by simp - from `0 < a` have "0 \ b mod a" by (rule pos_mod_sign) - then have B: "0 \ 1 + 2 * (b mod a)" by simp - have "((1\int) mod ((2\int) * a) + (2\int) * b mod ((2\int) * a)) mod ((2\int) * a) = (1\int) + (2\int) * (b mod a)" - using `0 < a` and A - by (auto simp add: mod_mult_mult1 mod_pos_pos_trivial ring_distribs intro!: mod_pos_pos_trivial B) - then show ?thesis by (subst mod_add_eq) -qed - -lemma neg_zmod_mult_2: - fixes a b :: int - assumes "a \ 0" - shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" -proof - - from assms have "0 \ - a" by auto - then have "(1 + 2 * (- b - 1)) mod (2 * (- a)) = 1 + 2 * ((- b - 1) mod (- a))" - by (rule pos_zmod_mult_2) - then show ?thesis by (simp add: zmod_zminus2 algebra_simps) - (simp add: diff_minus add_ac) -qed - -lemma zmod_number_of_Bit0 [simp]: - "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = - (2::int) * (number_of v mod number_of w)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac mult_2 [symmetric]) -done - -lemma zmod_number_of_Bit1 [simp]: - "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = - (if (0::int) \ number_of w - then 2 * (number_of v mod number_of w) + 1 - else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" -apply (simp only: number_of_eq numeral_simps) -apply (simp add: mod_mult_mult1 pos_zmod_mult_2 - neg_zmod_mult_2 add_ac mult_2 [symmetric]) -done - - -subsection{*Quotients of Signs*} - -lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" -apply (subgoal_tac "a div b \ -1", force) -apply (rule order_trans) -apply (rule_tac a' = "-1" in zdiv_mono1) -apply (auto simp add: div_eq_minus1) -done - -lemma div_nonneg_neg_le0: "[| (0::int) \ a; b < 0 |] ==> a div b \ 0" -by (drule zdiv_mono1_neg, auto) - -lemma div_nonpos_pos_le0: "[| (a::int) \ 0; b > 0 |] ==> a div b \ 0" -by (drule zdiv_mono1, auto) - -lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \ a div b) = (0 \ a)" -apply auto -apply (drule_tac [2] zdiv_mono1) -apply (auto simp add: linorder_neq_iff) -apply (simp (no_asm_use) add: linorder_not_less [symmetric]) -apply (blast intro: div_neg_pos_less0) -done - -lemma neg_imp_zdiv_nonneg_iff: - "b < (0::int) ==> (0 \ a div b) = (a \ (0::int))" -apply (subst zdiv_zminus_zminus [symmetric]) -apply (subst pos_imp_zdiv_nonneg_iff, auto) -done - -(*But not (a div b \ 0 iff a\0); consider a=1, b=2 when a div b = 0.*) -lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" -by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) - -(*Again the law fails for \: consider a = -1, b = -2 when a div b = 0*) -lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" -by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) - - -subsection {* The Divides Relation *} - -lemmas zdvd_iff_zmod_eq_0_number_of [simp] = - dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] - -lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" - by (rule dvd_mod) (* TODO: remove *) - -lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" - by (rule dvd_mod_imp_dvd) (* TODO: remove *) - -lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" - using zmod_zdiv_equality[where a="m" and b="n"] - by (simp add: algebra_simps) - -lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" -apply (induct "y", auto) -apply (rule zmod_zmult1_eq [THEN trans]) -apply (simp (no_asm_simp)) -apply (rule mod_mult_eq [symmetric]) -done - -lemma zdiv_int: "int (a div b) = (int a) div (int b)" -apply (subst split_div, auto) -apply (subst split_zdiv, auto) -apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient) -apply (auto simp add: IntDiv.divmod_int_rel_def of_nat_mult) -done - -lemma zmod_int: "int (a mod b) = (int a) mod (int b)" -apply (subst split_mod, auto) -apply (subst split_zmod, auto) -apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia - in unique_remainder) -apply (auto simp add: IntDiv.divmod_int_rel_def of_nat_mult) -done - -lemma abs_div: "(y::int) dvd x \ abs (x div y) = abs x div abs y" -by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) - -lemma zdvd_mult_div_cancel:"(n::int) dvd m \ n * (m div n) = m" -apply (subgoal_tac "m mod n = 0") - apply (simp add: zmult_div_cancel) -apply (simp only: dvd_eq_mod_eq_0) -done - -text{*Suggested by Matthias Daum*} -lemma int_power_div_base: - "\0 < m; 0 < k\ \ k ^ m div k = (k::int) ^ (m - Suc 0)" -apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") - apply (erule ssubst) - apply (simp only: power_add) - apply simp_all -done - -text {* by Brian Huffman *} -lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" -by (rule mod_minus_eq [symmetric]) - -lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" -by (rule mod_diff_left_eq [symmetric]) - -lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" -by (rule mod_diff_right_eq [symmetric]) - -lemmas zmod_simps = - mod_add_left_eq [symmetric] - mod_add_right_eq [symmetric] - zmod_zmult1_eq [symmetric] - mod_mult_left_eq [symmetric] - zpower_zmod - zminus_zmod zdiff_zmod_left zdiff_zmod_right - -text {* Distributive laws for function @{text nat}. *} - -lemma nat_div_distrib: "0 \ x \ nat (x div y) = nat x div nat y" -apply (rule linorder_cases [of y 0]) -apply (simp add: div_nonneg_neg_le0) -apply simp -apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) -done - -(*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) -lemma nat_mod_distrib: - "\0 \ x; 0 \ y\ \ nat (x mod y) = nat x mod nat y" -apply (case_tac "y = 0", simp add: DIVISION_BY_ZERO) -apply (simp add: nat_eq_iff zmod_int) -done - -text{*Suggested by Matthias Daum*} -lemma int_div_less_self: "\0 < x; 1 < k\ \ x div k < (x::int)" -apply (subgoal_tac "nat x div nat k < nat x") - apply (simp (asm_lr) add: nat_div_distrib [symmetric]) -apply (rule Divides.div_less_dividend, simp_all) -done - -text {* code generator setup *} - -context ring_1 -begin - -lemma of_int_num [code]: - "of_int k = (if k = 0 then 0 else if k < 0 then - - of_int (- k) else let - (l, m) = divmod_int k 2; - l' = of_int l - in if m = 0 then l' + l' else l' + l' + 1)" -proof - - have aux1: "k mod (2\int) \ (0\int) \ - of_int k = of_int (k div 2 * 2 + 1)" - proof - - have "k mod 2 < 2" by (auto intro: pos_mod_bound) - moreover have "0 \ k mod 2" by (auto intro: pos_mod_sign) - moreover assume "k mod 2 \ 0" - ultimately have "k mod 2 = 1" by arith - moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp - ultimately show ?thesis by auto - qed - have aux2: "\x. of_int 2 * x = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "of_int 2 * x = x + x" - unfolding int2 of_int_add left_distrib by simp - qed - have aux3: "\x. x * of_int 2 = x + x" - proof - - fix x - have int2: "(2::int) = 1 + 1" by arith - show "x * of_int 2 = x + x" - unfolding int2 of_int_add right_distrib by simp - qed - from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3) -qed - -end - -lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \ n dvd x - y" -proof - assume H: "x mod n = y mod n" - hence "x mod n - y mod n = 0" by simp - hence "(x mod n - y mod n) mod n = 0" by simp - hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) - thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) -next - assume H: "n dvd x - y" - then obtain k where k: "x-y = n*k" unfolding dvd_def by blast - hence "x = n*k + y" by simp - hence "x mod n = (n*k + y) mod n" by simp - thus "x mod n = y mod n" by (simp add: mod_add_left_eq) -qed - -lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \ x" - shows "\q. x = y + n * q" -proof- - from xy have th: "int x - int y = int (x - y)" by simp - from xyn have "int x mod int n = int y mod int n" - by (simp add: zmod_int[symmetric]) - hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) - hence "n dvd x - y" by (simp add: th zdvd_int) - then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith -qed - -lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \ (\q1 q2. x + n * q1 = y + n * q2)" - (is "?lhs = ?rhs") -proof - assume H: "x mod n = y mod n" - {assume xy: "x \ y" - from H have th: "y mod n = x mod n" by simp - from nat_mod_eq_lemma[OF th xy] have ?rhs - apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} - moreover - {assume xy: "y \ x" - from nat_mod_eq_lemma[OF H xy] have ?rhs - apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} - ultimately show ?rhs using linear[of x y] by blast -next - assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast - hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp - thus ?lhs by simp -qed - -lemma div_nat_number_of [simp]: - "(number_of v :: nat) div number_of v' = - (if neg (number_of v :: int) then 0 - else nat (number_of v div number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_div_distrib) - -lemma one_div_nat_number_of [simp]: - "Suc 0 div number_of v' = nat (1 div number_of v')" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - -lemma mod_nat_number_of [simp]: - "(number_of v :: nat) mod number_of v' = - (if neg (number_of v :: int) then 0 - else if neg (number_of v' :: int) then number_of v - else nat (number_of v mod number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_mod_distrib) - -lemma one_mod_nat_number_of [simp]: - "Suc 0 mod number_of v' = - (if neg (number_of v' :: int) then Suc 0 - else nat (1 mod number_of v'))" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - -lemmas dvd_eq_mod_eq_0_number_of = - dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] - -declare dvd_eq_mod_eq_0_number_of [simp] - - -subsection {* Transfer setup *} - -lemma transfer_nat_int_functions: - "(x::int) >= 0 \ y >= 0 \ (nat x) div (nat y) = nat (x div y)" - "(x::int) >= 0 \ y >= 0 \ (nat x) mod (nat y) = nat (x mod y)" - by (auto simp add: nat_div_distrib nat_mod_distrib) - -lemma transfer_nat_int_function_closures: - "(x::int) >= 0 \ y >= 0 \ x div y >= 0" - "(x::int) >= 0 \ y >= 0 \ x mod y >= 0" - apply (cases "y = 0") - apply (auto simp add: pos_imp_zdiv_nonneg_iff) - apply (cases "y = 0") - apply auto -done - -declare TransferMorphism_nat_int[transfer add return: - transfer_nat_int_functions - transfer_nat_int_function_closures -] - -lemma transfer_int_nat_functions: - "(int x) div (int y) = int (x div y)" - "(int x) mod (int y) = int (x mod y)" - by (auto simp add: zdiv_int zmod_int) - -lemma transfer_int_nat_function_closures: - "is_nat x \ is_nat y \ is_nat (x div y)" - "is_nat x \ is_nat y \ is_nat (x mod y)" - by (simp_all only: is_nat_def transfer_nat_int_function_closures) - -declare TransferMorphism_int_nat[transfer add return: - transfer_int_nat_functions - transfer_int_nat_function_closures -] - - -subsection {* Code generation *} - -definition pdivmod :: "int \ int \ int \ int" where - "pdivmod k l = (\k\ div \l\, \k\ mod \l\)" - -lemma pdivmod_posDivAlg [code]: - "pdivmod k l = (if l = 0 then (0, \k\) else posDivAlg \k\ \l\)" -by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) - -lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else - apsnd ((op *) (sgn l)) (if 0 < l \ 0 \ k \ l < 0 \ k < 0 - then pdivmod k l - else (let (r, s) = pdivmod k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" -proof - - have aux: "\q::int. - k = l * q \ k = l * - q" by auto - show ?thesis - by (simp add: divmod_int_mod_div pdivmod_def) - (auto simp add: aux not_less not_le zdiv_zminus1_eq_if - zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) -qed - -lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else - apsnd ((op *) (sgn l)) (if sgn k = sgn l - then pdivmod k l - else (let (r, s) = pdivmod k l in - if s = 0 then (- r, 0) else (- r - 1, \l\ - s))))" -proof - - have "k \ 0 \ l \ 0 \ 0 < l \ 0 \ k \ l < 0 \ k < 0 \ sgn k = sgn l" - by (auto simp add: not_less sgn_if) - then show ?thesis by (simp add: divmod_int_pdivmod) -qed - -code_modulename SML - IntDiv Integer - -code_modulename OCaml - IntDiv Integer - -code_modulename Haskell - IntDiv Integer - - - -subsection {* Proof Tools setup; Combination and Cancellation Simprocs *} - -declare split_div[of _ _ "number_of k", standard, arith_split] -declare split_mod[of _ _ "number_of k", standard, arith_split] - - -subsubsection{*For @{text combine_numerals}*} - -lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" -by (simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numerals}*} - -lemma nat_diff_add_eq1: - "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_diff_add_eq2: - "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_eq_add_iff1: - "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_eq_add_iff2: - "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff1: - "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff2: - "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff1: - "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff2: - "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numeral_factors} *} - -lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" -by auto - -lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" -by auto - -lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" -by auto - -lemma nat_mult_dvd_cancel_disj[simp]: - "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))" -by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric]) - -lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" -by(auto) - - -subsubsection{*For @{text cancel_factor} *} - -lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)" -by auto - -lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, - @{thm nat_0}, @{thm nat_1}, - @{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}, @{thm le_number_of_eq_not_less}, - @{thm le_Suc_number_of}, @{thm le_number_of_Suc}, - @{thm less_Suc_number_of}, @{thm less_number_of_Suc}, - @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc}, - @{thm mult_Suc}, @{thm mult_Suc_right}, - @{thm add_Suc}, @{thm add_Suc_right}, - @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of}, - @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, - @{thm if_True}, @{thm if_False}]) - #> Lin_Arith.add_simprocs (Numeral_Simprocs.assoc_fold_simproc - :: Numeral_Simprocs.combine_numerals - :: Numeral_Simprocs.cancel_numerals) - #> Lin_Arith.add_simprocs (Nat_Numeral_Simprocs.combine_numerals :: Nat_Numeral_Simprocs.cancel_numerals)) -*} - -end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/IsaMakefile Fri Oct 30 14:02:42 2009 +0100 @@ -248,12 +248,12 @@ Groebner_Basis.thy \ Hilbert_Choice.thy \ Int.thy \ - IntDiv.thy \ List.thy \ Main.thy \ Map.thy \ Nat_Numeral.thy \ Nat_Transfer.thy \ + Numeral_Simprocs.thy \ Presburger.thy \ Predicate_Compile.thy \ Quickcheck.thy \ @@ -382,7 +382,6 @@ Library/While_Combinator.thy Library/Product_ord.thy \ Library/Char_nat.thy Library/Char_ord.thy Library/Option_ord.thy \ Library/Sublist_Order.thy Library/List_lexord.thy \ - Library/Commutative_Ring.thy Library/comm_ring.ML \ Library/Coinductive_List.thy Library/AssocList.thy \ Library/Formal_Power_Series.thy Library/Binomial.thy \ Library/Eval_Witness.thy Library/Code_Char.thy \ @@ -785,6 +784,9 @@ $(LOG)/HOL-Decision_Procs.gz: $(OUT)/HOL \ Decision_Procs/Approximation.thy \ + Decision_Procs/Commutative_Ring.thy \ + Decision_Procs/Commutative_Ring_Complete.thy \ + Decision_Procs/commutative_ring_tac.ML \ Decision_Procs/Cooper.thy \ Decision_Procs/cooper_tac.ML \ Decision_Procs/Dense_Linear_Order.thy \ @@ -795,6 +797,7 @@ Decision_Procs/Decision_Procs.thy \ Decision_Procs/ex/Dense_Linear_Order_Ex.thy \ Decision_Procs/ex/Approximation_Ex.thy \ + Decision_Procs/ex/Commutative_Ring_Ex.thy \ Decision_Procs/ROOT.ML @$(ISABELLE_TOOL) usedir $(OUT)/HOL Decision_Procs @@ -937,7 +940,7 @@ HOL-ex: HOL $(LOG)/HOL-ex.gz -$(LOG)/HOL-ex.gz: $(OUT)/HOL Library/Commutative_Ring.thy \ +$(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \ Number_Theory/Primes.thy \ Tools/Predicate_Compile/predicate_compile_core.ML \ ex/Abstract_NAT.thy ex/Antiquote.thy ex/Arith_Examples.thy \ @@ -945,8 +948,8 @@ ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy \ ex/CodegenSML_Test.thy ex/Codegenerator_Candidates.thy \ ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy \ - ex/Codegenerator_Test.thy ex/Coherent.thy ex/Commutative_RingEx.thy \ - ex/Commutative_Ring_Complete.thy ex/Efficient_Nat_examples.thy \ + ex/Codegenerator_Test.thy ex/Coherent.thy \ + ex/Efficient_Nat_examples.thy \ ex/Eval_Examples.thy ex/Fundefs.thy ex/Groebner_Examples.thy \ ex/Guess.thy ex/HarmonicSeries.thy ex/Hebrew.thy \ ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy \ diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Library/Commutative_Ring.thy --- a/src/HOL/Library/Commutative_Ring.thy Fri Oct 30 11:31:34 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,319 +0,0 @@ -(* Author: Bernhard Haeupler - -Proving equalities in commutative rings done "right" in Isabelle/HOL. -*) - -header {* Proving equalities in commutative rings *} - -theory Commutative_Ring -imports List Parity 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. *} - -primrec - Ipol :: "'a::{comm_ring_1} list \ 'a pol \ 'a" -where - "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 :: "'a::{comm_ring_1} list \ 'a polex \ 'a" -where - "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. *} - -definition - mkPinj :: "nat \ 'a pol \ 'a pol" where - "mkPinj x P = (case P of - Pc c \ Pc c | - Pinj y P \ Pinj (x + y) P | - PX p1 y p2 \ Pinj x P)" - -definition - mkPX :: "'a::{comm_ring} pol \ nat \ 'a pol \ 'a pol" where - "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 *} - -function - add :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) -where - "Pc a \ Pc b = Pc (a + b)" - | "Pc c \ Pinj i P = Pinj i (P \ Pc c)" - | "Pinj i P \ Pc c = Pinj i (P \ Pc c)" - | "Pc c \ PX P i Q = PX P i (Q \ Pc c)" - | "PX P i Q \ Pc c = PX P i (Q \ Pc c)" - | "Pinj x P \ Pinj y Q = - (if x = y then mkPinj x (P \ Q) - else (if x > y then mkPinj y (Pinj (x - y) P \ Q) - else mkPinj x (Pinj (y - x) Q \ P)))" - | "Pinj x P \ PX Q y R = - (if x = 0 then P \ PX Q y R - else (if x = 1 then PX Q y (R \ P) - else PX Q y (R \ Pinj (x - 1) P)))" - | "PX P x R \ Pinj y Q = - (if y = 0 then PX P x R \ Q - else (if y = 1 then PX P x (R \ Q) - else PX P x (R \ Pinj (y - 1) Q)))" - | "PX P1 x P2 \ PX Q1 y Q2 = - (if x = y then mkPX (P1 \ Q1) x (P2 \ Q2) - else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \ Q1) y (P2 \ Q2) - else mkPX (PX Q1 (y-x) (Pc 0) \ P1) x (P2 \ Q2)))" -by pat_completeness auto -termination by (relation "measure (\(x, y). size x + size y)") auto - -function - mul :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 70) -where - "Pc a \ Pc b = Pc (a * b)" - | "Pc c \ Pinj i P = - (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" - | "Pinj i P \ Pc c = - (if c = 0 then Pc 0 else mkPinj i (P \ Pc c))" - | "Pc c \ PX P i Q = - (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" - | "PX P i Q \ Pc c = - (if c = 0 then Pc 0 else mkPX (P \ Pc c) i (Q \ Pc c))" - | "Pinj x P \ Pinj y Q = - (if x = y then mkPinj x (P \ Q) else - (if x > y then mkPinj y (Pinj (x-y) P \ Q) - else mkPinj x (Pinj (y - x) Q \ P)))" - | "Pinj x P \ PX Q y R = - (if x = 0 then P \ PX Q y R else - (if x = 1 then mkPX (Pinj x P \ Q) y (R \ P) - else mkPX (Pinj x P \ Q) y (R \ Pinj (x - 1) P)))" - | "PX P x R \ Pinj y Q = - (if y = 0 then PX P x R \ Q else - (if y = 1 then mkPX (Pinj y Q \ P) x (R \ Q) - else mkPX (Pinj y Q \ P) x (R \ Pinj (y - 1) Q)))" - | "PX P1 x P2 \ PX Q1 y Q2 = - mkPX (P1 \ Q1) (x + y) (P2 \ Q2) \ - (mkPX (P1 \ mkPinj 1 Q2) x (Pc 0) \ - (mkPX (Q1 \ mkPinj 1 P2) y (Pc 0)))" -by pat_completeness auto -termination by (relation "measure (\(x, y). size x + size y)") - (auto simp add: mkPinj_def split: pol.split) - -text {* Negation*} -primrec - neg :: "'a::{comm_ring} pol \ 'a pol" -where - "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 *} -definition - sub :: "'a::{comm_ring} pol \ 'a pol \ 'a pol" (infixl "\" 65) -where - "sub P Q = P \ neg Q" - -text {* Square for Fast Exponentation *} -primrec - sqr :: "'a::{comm_ring_1} pol \ 'a pol" -where - "sqr (Pc c) = Pc (c * c)" - | "sqr (Pinj i P) = mkPinj i (sqr P)" - | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \ - mkPX (Pc (1 + 1) \ A \ mkPinj 1 B) x (Pc 0)" - -text {* Fast Exponentation *} -fun - pow :: "nat \ 'a::{comm_ring_1} pol \ 'a pol" -where - "pow 0 P = Pc 1" - | "pow n P = (if even n then pow (n div 2) (sqr P) - else P \ pow (n div 2) (sqr P))" - -lemma pow_if: - "pow n P = - (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) - else P \ pow (n div 2) (sqr P))" - by (cases n) simp_all - - -text {* Normalization of polynomial expressions *} - -primrec - norm :: "'a::{comm_ring_1} polex \ 'a pol" -where - "norm (Pol P) = P" - | "norm (Add P Q) = norm P \ norm Q" - | "norm (Sub P Q) = norm P \ norm Q" - | "norm (Mul P Q) = norm P \ norm Q" - | "norm (Pow P n) = pow n (norm P)" - | "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 algebra_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 algebra_simps) - -text {* Correctness theorems for the implemented operations *} - -text {* Negation *} -lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" - by (induct P arbitrary: l) auto - -text {* Addition *} -lemma add_ci: "Ipol l (P \ Q) = Ipol l P + Ipol l Q" -proof (induct P Q arbitrary: l 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 algebra_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 algebra_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: algebra_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 algebra_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 algebra_simps) - next - assume "x = y" - with 9 show ?case by (simp add: mkPX_ci algebra_simps) - qed -qed (auto simp add: algebra_simps) - -text {* Multiplication *} -lemma mul_ci: "Ipol l (P \ Q) = Ipol l P * Ipol l Q" - by (induct P Q arbitrary: l rule: mul.induct) - (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add) - -text {* Substraction *} -lemma sub_ci: "Ipol l (P \ Q) = Ipol l P - Ipol l Q" - by (simp add: add_ci neg_ci sub_def) - -text {* Square *} -lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" - by (induct P arbitrary: ls) - (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add) - -text {* Power *} -lemma even_pow:"even n \ pow n P = pow (n div 2) (sqr P)" - by (induct n) simp_all - -lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" -proof (induct n arbitrary: P rule: nat_less_induct) - case (1 k) - show ?case - proof (cases k) - case 0 - then show ?thesis by simp - next - case (Suc l) - show ?thesis - proof cases - assume "even l" - then have "Suc l div 2 = l div 2" - by (simp add: nat_number even_nat_plus_one_div_two) - moreover - from Suc have "l < k" by simp - with 1 have "\P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp - moreover - note Suc `even l` even_nat_plus_one_div_two - ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) - next - assume "odd l" - { - fix p - have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" - proof (cases l) - case 0 - with `odd l` show ?thesis by simp - next - case (Suc w) - with `odd l` have "even w" by simp - have two_times: "2 * (w div 2) = w" - by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) - have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" - by (simp add: power_Suc) - then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" - by (simp add: numerals) - with Suc show ?thesis - by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci - simp del: power_Suc) - qed - } with 1 Suc `odd l` show ?thesis by simp - qed - 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 "norm P1 = norm P2" - shows "Ipolex l P1 = Ipolex l P2" -proof - - from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp - then show ?thesis by (simp only: norm_ci) -qed - - -use "comm_ring.ML" -setup CommRing.setup - -end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Library/Library.thy Fri Oct 30 14:02:42 2009 +0100 @@ -11,7 +11,6 @@ Code_Char_chr Code_Integer Coinductive_List - Commutative_Ring Continuity ContNotDenum Countable diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Library/Numeral_Type.thy Fri Oct 30 14:02:42 2009 +0100 @@ -188,7 +188,7 @@ by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" -by (simp add: Abs_inverse IntDiv.pos_mod_conj [OF size0]) +by (simp add: Abs_inverse pos_mod_conj [OF size0]) lemma Rep_Abs_0: "Rep (Abs 0) = 0" by (simp add: Abs_inverse size0) diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Library/Word.thy Fri Oct 30 14:02:42 2009 +0100 @@ -5,7 +5,7 @@ header {* Binary Words *} theory Word -imports "~~/src/HOL/Main" +imports Main begin subsection {* Auxilary Lemmas *} @@ -561,35 +561,17 @@ shows "nat_to_bv (2 * bv_to_nat w + bitval x) = norm_unsigned (w @ [x])" proof (cases x) assume [simp]: "x = \" - show ?thesis - apply (simp add: nat_to_bv_non0) - apply safe - proof - - fix q - assume "Suc (2 * bv_to_nat w) = 2 * q" - hence orig: "(2 * bv_to_nat w + 1) mod 2 = 2 * q mod 2" (is "?lhs = ?rhs") - by simp - have "?lhs = (1 + 2 * bv_to_nat w) mod 2" - by (simp add: add_commute) - also have "... = 1" - by (subst mod_add_eq) simp - finally have eq1: "?lhs = 1" . - have "?rhs = 0" by simp - with orig and eq1 - show "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" - by simp - next - have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = - nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\]" - by (simp add: add_commute) - also have "... = nat_to_bv (bv_to_nat w) @ [\]" - by (subst div_add1_eq) simp - also have "... = norm_unsigned w @ [\]" - by (subst ass) (rule refl) - also have "... = norm_unsigned (w @ [\])" - by (cases "norm_unsigned w") simp_all - finally show "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" . - qed + have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = + nat_to_bv ((1 + 2 * bv_to_nat w) div 2) @ [\]" + by (simp add: add_commute) + also have "... = nat_to_bv (bv_to_nat w) @ [\]" + by (subst div_add1_eq) simp + also have "... = norm_unsigned w @ [\]" + by (subst ass) (rule refl) + also have "... = norm_unsigned (w @ [\])" + by (cases "norm_unsigned w") simp_all + finally have "nat_to_bv (Suc (2 * bv_to_nat w) div 2) @ [\] = norm_unsigned (w @ [\])" . + then show ?thesis by (simp add: nat_to_bv_non0) next assume [simp]: "x = \" show ?thesis diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Library/comm_ring.ML --- a/src/HOL/Library/comm_ring.ML Fri Oct 30 11:31:34 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,109 +0,0 @@ -(* Author: Amine Chaieb - -Tactic for solving equalities over commutative rings. -*) - -signature COMM_RING = -sig - val comm_ring_tac : Proof.context -> int -> tactic - val setup : theory -> theory -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 (@{const_name HOL.zero}, T); -fun cring_one T = Const (@{const_name HOL.one}, T); - -(* reification functions *) -(* add two polynom expressions *) -fun polT t = Type ("Commutative_Ring.pol", [t]); -fun polexT t = Type ("Commutative_Ring.polex", [t]); - -(* pol *) -fun pol_Pc t = Const ("Commutative_Ring.pol.Pc", t --> polT t); -fun pol_Pinj t = Const ("Commutative_Ring.pol.Pinj", HOLogic.natT --> polT t --> polT t); -fun pol_PX t = Const ("Commutative_Ring.pol.PX", polT t --> HOLogic.natT --> 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 --> HOLogic.natT --> polexT t); - -(* reification of polynoms : primitive cring expressions *) -fun reif_pol T vs (t as Free _) = - let - val one = @{term "1::nat"}; - val i = find_index (fn t' => t' = 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 $ HOLogic.mk_nat i - $ (pol_PX T $ (pol_Pc T $ cring_one T) - $ one $ (pol_Pc T $ cring_zero T)) - end - | reif_pol T vs t = pol_Pc T $ t; - -(* reification of polynom expressions *) -fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) = - polex_add T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) = - polex_sub T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) = - polex_mul T $ reif_polex T vs a $ reif_polex T vs b - | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) = - polex_neg T $ reif_polex T vs a - | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) = - polex_pow T $ reif_polex T vs a $ n - | reif_polex T vs t = polex_pol T $ reif_pol T vs t; - -(* reification of the equation *) -val cr_sort = @{sort "comm_ring_1"}; - -fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) = - if Sign.of_sort thy (T, cr_sort) then - let - val fs = OldTerm.term_frees eq; - val cvs = cterm_of thy (HOLogic.mk_list T fs); - val clhs = cterm_of thy (reif_polex T fs lhs); - val crhs = cterm_of thy (reif_polex T fs rhs); - val ca = ctyp_of thy T; - in (ca, cvs, clhs, crhs) end - else raise CRing ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort) - | reif_eq _ _ = 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 *) -val cring_simps = - [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add}, - @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]]; - -fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) => - let - val thy = ProofContext.theory_of ctxt; - val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) - addsimps cring_simps; - val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) - val norm_eq_th = - simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq}) - in - cut_rules_tac [norm_eq_th] i - THEN (simp_tac cring_ss i) - THEN (simp_tac cring_ss i) - end); - -val setup = - Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o comm_ring_tac)) - "reflective decision procedure for equalities over commutative rings" #> - Method.setup @{binding algebra} (Scan.succeed (SIMPLE_METHOD' o comm_ring_tac)) - "method for proving algebraic properties (same as comm_ring)"; - -end; diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Parity.thy --- a/src/HOL/Parity.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Parity.thy Fri Oct 30 14:02:42 2009 +0100 @@ -315,42 +315,6 @@ qed qed -subsection {* General Lemmas About Division *} - -(*FIXME move to Divides.thy*) - -lemma Suc_times_mod_eq: "1 Suc (k * m) mod k = 1" -apply (induct "m") -apply (simp_all add: mod_Suc) -done - -declare Suc_times_mod_eq [of "number_of w", standard, simp] - -lemma [simp]: "n div k \ (Suc n) div k" -by (simp add: div_le_mono) - -lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" -by arith - -lemma div_2_gt_zero [simp]: "(1::nat) < n ==> 0 < n div 2" -by arith - - (* Potential use of algebra : Equality modulo n*) -lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" -by (simp add: mult_ac add_ac) - -lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" -proof - - have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp - also have "... = Suc m mod n" by (rule mod_mult_self3) - finally show ?thesis . -qed - -lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" -apply (subst mod_Suc [of m]) -apply (subst mod_Suc [of "m mod n"], simp) -done - subsection {* More Even/Odd Results *} diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Tools/arith_data.ML --- a/src/HOL/Tools/arith_data.ML Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Tools/arith_data.ML Fri Oct 30 14:02:42 2009 +0100 @@ -1,7 +1,7 @@ (* Title: HOL/Tools/arith_data.ML Author: Markus Wenzel, Stefan Berghofer, and Tobias Nipkow -Common arithmetic proof auxiliary. +Common arithmetic proof auxiliary and legacy. *) signature ARITH_DATA = @@ -11,6 +11,11 @@ val add_tactic: string -> (bool -> Proof.context -> int -> tactic) -> theory -> theory val get_arith_facts: Proof.context -> thm list + val mk_number: typ -> int -> term + val mk_sum: typ -> term list -> term + val long_mk_sum: typ -> term list -> term + val dest_sum: term -> term list + val prove_conv_nohyps: tactic list -> Proof.context -> term * term -> thm option val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option val prove_conv2: tactic -> (simpset -> tactic) -> simpset -> term * term -> thm @@ -67,6 +72,36 @@ "various arithmetic decision procedures"; +(* some specialized syntactic operations *) + +fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; + +val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; + +fun mk_minus t = + let val T = Term.fastype_of t + in Const (@{const_name HOL.uminus}, T --> T) $ t end; + +(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) +fun mk_sum T [] = mk_number T 0 + | mk_sum T [t,u] = mk_plus (t, u) + | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum T [] = mk_number T 0 + | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else mk_minus t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + + (* various auxiliary and legacy *) fun prove_conv_nohyps tacs ctxt (t, u) = diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/Tools/numeral_simprocs.ML Fri Oct 30 14:02:42 2009 +0100 @@ -16,9 +16,6 @@ signature NUMERAL_SIMPROCS = sig - val mk_sum: typ -> term list -> term - val dest_sum: term -> term list - val assoc_fold_simproc: simproc val combine_numerals: simproc val cancel_numerals: simproc list @@ -32,39 +29,10 @@ structure Numeral_Simprocs : NUMERAL_SIMPROCS = struct -fun mk_number T n = HOLogic.number_of_const T $ HOLogic.mk_numeral n; - -fun find_first_numeral past (t::terms) = - ((snd (HOLogic.dest_number t), rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val mk_plus = HOLogic.mk_binop @{const_name HOL.plus}; - -fun mk_minus t = - let val T = Term.fastype_of t - in Const (@{const_name HOL.uminus}, T --> T) $ t end; - -(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) -fun mk_sum T [] = mk_number T 0 - | mk_sum T [t,u] = mk_plus (t, u) - | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum T [] = mk_number T 0 - | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); - -val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} Term.dummyT; - -(*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const (@{const_name HOL.plus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const (@{const_name HOL.minus}, _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (not pos, u, ts)) - | dest_summing (pos, t, ts) = - if pos then t::ts else mk_minus t :: ts; - -fun dest_sum t = dest_summing (true, t, []); +val mk_number = Arith_Data.mk_number; +val mk_sum = Arith_Data.mk_sum; +val long_mk_sum = Arith_Data.long_mk_sum; +val dest_sum = Arith_Data.dest_sum; val mk_diff = HOLogic.mk_binop @{const_name HOL.minus}; val dest_diff = HOLogic.dest_bin @{const_name HOL.minus} Term.dummyT; @@ -95,6 +63,11 @@ in dest_prod t @ dest_prod u end handle TERM _ => [t]; +fun find_first_numeral past (t::terms) = + ((snd (HOLogic.dest_number t), rev past @ terms) + handle TERM _ => find_first_numeral (t::past) terms) + | find_first_numeral past [] = raise TERM("find_first_numeral", []); + (*DON'T do the obvious simplifications; that would create special cases*) fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/ex/Codegenerator_Candidates.thy --- a/src/HOL/ex/Codegenerator_Candidates.thy Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/ex/Codegenerator_Candidates.thy Fri Oct 30 14:02:42 2009 +0100 @@ -9,7 +9,6 @@ AssocList Binomial Fset - Commutative_Ring Enum List_Prefix Nat_Infinity @@ -22,7 +21,7 @@ Tree While_Combinator Word - "~~/src/HOL/ex/Commutative_Ring_Complete" + "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete" "~~/src/HOL/ex/Records" begin diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/ex/Commutative_RingEx.thy --- a/src/HOL/ex/Commutative_RingEx.thy Fri Oct 30 11:31:34 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,50 +0,0 @@ -(* ID: $Id$ - Author: Bernhard Haeupler -*) - -header {* Some examples demonstrating the comm-ring method *} - -theory Commutative_RingEx -imports Commutative_Ring -begin - -lemma "4*(x::int)^5*y^3*x^2*3 + x*z + 3^5 = 12*x^7*y^3 + z*x + 243" -by comm_ring - -lemma "((x::int) + y)^2 = x^2 + y^2 + 2*x*y" -by comm_ring - -lemma "((x::int) + y)^3 = x^3 + y^3 + 3*x^2*y + 3*y^2*x" -by comm_ring - -lemma "((x::int) - y)^3 = x^3 + 3*x*y^2 + (-3)*y*x^2 - y^3" -by comm_ring - -lemma "((x::int) - y)^2 = x^2 + y^2 - 2*x*y" -by comm_ring - -lemma " ((a::int) + b + c)^2 = a^2 + b^2 + c^2 + 2*a*b + 2*b*c + 2*a*c" -by comm_ring - -lemma "((a::int) - b - c)^2 = a^2 + b^2 + c^2 - 2*a*b + 2*b*c - 2*a*c" -by comm_ring - -lemma "(a::int)*b + a*c = a*(b+c)" -by comm_ring - -lemma "(a::int)^2 - b^2 = (a - b) * (a + b)" -by comm_ring - -lemma "(a::int)^3 - b^3 = (a - b) * (a^2 + a*b + b^2)" -by comm_ring - -lemma "(a::int)^3 + b^3 = (a + b) * (a^2 - a*b + b^2)" -by comm_ring - -lemma "(a::int)^4 - b^4 = (a - b) * (a + b)*(a^2 + b^2)" -by comm_ring - -lemma "(a::int)^10 - b^10 = (a - b) * (a^9 + a^8*b + a^7*b^2 + a^6*b^3 + a^5*b^4 + a^4*b^5 + a^3*b^6 + a^2*b^7 + a*b^8 + b^9 )" -by comm_ring - -end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/ex/Commutative_Ring_Complete.thy --- a/src/HOL/ex/Commutative_Ring_Complete.thy Fri Oct 30 11:31:34 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,391 +0,0 @@ -(* Author: Bernhard Haeupler - -This theory is about of the relative completeness of method comm-ring -method. As long as the reified atomic polynomials of type 'a pol are -in normal form, the cring method is complete. -*) - -header {* Proof of the relative completeness of method comm-ring *} - -theory Commutative_Ring_Complete -imports Commutative_Ring -begin - -text {* Formalization of normal form *} -fun - isnorm :: "('a::{comm_ring}) pol \ bool" -where - "isnorm (Pc c) \ True" - | "isnorm (Pinj i (Pc c)) \ False" - | "isnorm (Pinj i (Pinj j Q)) \ False" - | "isnorm (Pinj 0 P) \ False" - | "isnorm (Pinj i (PX Q1 j Q2)) \ isnorm (PX Q1 j Q2)" - | "isnorm (PX P 0 Q) \ False" - | "isnorm (PX (Pc c) i Q) \ c \ 0 \ isnorm Q" - | "isnorm (PX (PX P1 j (Pc c)) i Q) \ c \ 0 \ isnorm (PX P1 j (Pc c)) \ isnorm Q" - | "isnorm (PX P i Q) \ isnorm P \ isnorm Q" - -(* Some helpful lemmas *) -lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False" -by(cases P, auto) - -lemma norm_PX_0_False:"isnorm (PX (Pc 0) i Q) = False" -by(cases i, auto) - -lemma norm_Pinj:"isnorm (Pinj i Q) \ isnorm Q" -by(cases i,simp add: norm_Pinj_0_False norm_PX_0_False,cases Q) auto - -lemma norm_PX2:"isnorm (PX P i Q) \ isnorm Q" -by(cases i, auto, cases P, auto, case_tac pol2, auto) - -lemma norm_PX1:"isnorm (PX P i Q) \ isnorm P" -by(cases i, auto, cases P, auto, case_tac pol2, auto) - -lemma mkPinj_cn:"\y~=0; isnorm Q\ \ isnorm (mkPinj y Q)" -apply(auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) -apply(case_tac nat, auto simp add: norm_Pinj_0_False) -by(case_tac pol, auto) (case_tac y, auto) - -lemma norm_PXtrans: - assumes A:"isnorm (PX P x Q)" and "isnorm Q2" - shows "isnorm (PX P x Q2)" -proof(cases P) - case (PX p1 y p2) from prems show ?thesis by(cases x, auto, cases p2, auto) -next - case Pc from prems show ?thesis by(cases x, auto) -next - case Pinj from prems show ?thesis by(cases x, auto) -qed - -lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)" -proof(cases P) - case (PX p1 y p2) - from prems show ?thesis by(cases x, auto, cases p2, auto) -next - case Pc - from prems show ?thesis by(cases x, auto) -next - case Pinj - from prems show ?thesis by(cases x, auto) -qed - -text {* mkPX conserves normalizedness (@{text "_cn"}) *} -lemma mkPX_cn: - assumes "x \ 0" and "isnorm P" and "isnorm Q" - shows "isnorm (mkPX P x Q)" -proof(cases P) - case (Pc c) - from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) -next - case (Pinj i Q) - from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) -next - case (PX P1 y P2) - from prems have Y0:"y>0" by(cases y, auto) - from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) - with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto) -qed - -text {* add conserves normalizedness *} -lemma add_cn:"isnorm P \ isnorm Q \ isnorm (P \ Q)" -proof(induct P Q rule: add.induct) - case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all) -next - case (3 i P2 c) thus ?case by (cases P2, simp_all, cases i, simp_all) -next - case (4 c P2 i Q2) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) -next - case (5 P2 i Q2 c) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) -next - case (6 x P2 y Q2) - from prems have Y0:"y>0" by (cases y, auto simp add: norm_Pinj_0_False) - from prems have X0:"x>0" by (cases x, auto simp add: norm_Pinj_0_False) - have "x < y \ x = y \ x > y" by arith - moreover - { assume "xy" hence "EX d. x=d+y" by arith - then obtain d where "x=d+y".. - moreover - note prems Y0 - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) - ultimately have ?case by (simp add: mkPinj_cn)} - ultimately show ?case by blast -next - case (7 x P2 Q2 y R) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" by simp - with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact - with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - ultimately show ?case by blast -next - case (8 Q2 y R x P2) - have "x = 0 \ x = 1 \ x > 1" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" by simp - with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - with prems NR have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (PX Q2 y R)" by simp fact - with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } - ultimately show ?case by blast -next - case (9 P1 x P2 Q1 y Q2) - from prems have Y0:"y>0" by(cases y, auto) - from prems have X0:"x>0" by(cases x, auto) - from prems have NP1:"isnorm P1" and NP2:"isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) - from prems have NQ1:"isnorm Q1" and NQ2:"isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) - have "y < x \ x = y \ x < y" by arith - moreover - {assume sm1:"y < x" hence "EX d. x=d+y" by arith - then obtain d where sm2:"x=d+y".. - note prems NQ1 NP1 NP2 NQ2 sm1 sm2 - moreover - have "isnorm (PX P1 d (Pc 0))" - proof(cases P1) - case (PX p1 y p2) - with prems show ?thesis by(cases d, simp,cases p2, auto) - next case Pc from prems show ?thesis by(cases d, auto) - next case Pinj from prems show ?thesis by(cases d, auto) - qed - ultimately have "isnorm (P2 \ Q2)" "isnorm (PX P1 (x - y) (Pc 0) \ Q1)" by auto - with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)} - moreover - {assume "x=y" - from prems NP1 NP2 NQ1 NQ2 have "isnorm (P2 \ Q2)" "isnorm (P1 \ Q1)" by auto - with Y0 prems have ?case by (simp add: mkPX_cn) } - moreover - {assume sm1:"x Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \ P1)" by auto - with X0 sm1 sm2 have ?case by (simp add: mkPX_cn)} - ultimately show ?case by blast -qed simp - -text {* mul concerves normalizedness *} -lemma mul_cn :"isnorm P \ isnorm Q \ isnorm (P \ Q)" -proof(induct P Q rule: mul.induct) - case (2 c i P2) thus ?case - by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) -next - case (3 i P2 c) thus ?case - by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) -next - case (4 c P2 i Q2) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case - by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) -next - case (5 P2 i Q2 c) - from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with prems show ?case - by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) -next - case (6 x P2 y Q2) - have "x < y \ x = y \ x > y" by arith - moreover - { assume "x0" by (cases x, auto simp add: norm_Pinj_0_False) - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) - ultimately have ?case by (simp add: mkPinj_cn)} - moreover - { assume "x=y" - moreover - from prems have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) - moreover - note prems - moreover - ultimately have ?case by (simp add: mkPinj_cn) } - moreover - { assume "x>y" hence "EX d. x=d+y" by arith - then obtain d where "x=d+y".. - moreover - note prems - moreover - from prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) - moreover - from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - moreover - with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) - ultimately have ?case by (simp add: mkPinj_cn) } - ultimately show ?case by blast -next - case (7 x P2 Q2 y R) - from prems have Y0:"y>0" by(cases y, auto) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) - with Y0 prems have ?case by (simp add: mkPX_cn)} - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) - moreover - from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - moreover - from prems have "isnorm (Pinj x P2)" by(cases P2, auto) - moreover - note prems - ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto - with Y0 X have ?case by (simp add: mkPX_cn)} - ultimately show ?case by blast -next - case (8 Q2 y R x P2) - from prems have Y0:"y>0" by(cases y, auto) - have "x=0 \ (x = 1) \ (x > 1)" by arith - moreover - { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} - moreover - { assume "x=1" - from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) - with prems have "isnorm (R \ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) - with Y0 prems have ?case by (simp add: mkPX_cn) } - moreover - { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith - then obtain d where X:"x=Suc (Suc d)" .. - from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) - moreover - from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) - moreover - from prems have "isnorm (Pinj x P2)" by(cases P2, auto) - moreover - note prems - ultimately have "isnorm (R \ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \ Q2)" by auto - with Y0 X have ?case by (simp add: mkPX_cn) } - ultimately show ?case by blast -next - case (9 P1 x P2 Q1 y Q2) - from prems have X0:"x>0" by(cases x, auto) - from prems have Y0:"y>0" by(cases y, auto) - note prems - moreover - from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - moreover - from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) - ultimately have "isnorm (P1 \ Q1)" "isnorm (P2 \ Q2)" - "isnorm (P1 \ mkPinj 1 Q2)" "isnorm (Q1 \ mkPinj 1 P2)" - by (auto simp add: mkPinj_cn) - with prems X0 Y0 have - "isnorm (mkPX (P1 \ Q1) (x + y) (P2 \ Q2))" - "isnorm (mkPX (P1 \ mkPinj (Suc 0) Q2) x (Pc 0))" - "isnorm (mkPX (Q1 \ mkPinj (Suc 0) P2) y (Pc 0))" - by (auto simp add: mkPX_cn) - thus ?case by (simp add: add_cn) -qed(simp) - -text {* neg conserves normalizedness *} -lemma neg_cn: "isnorm P \ isnorm (neg P)" -proof (induct P) - case (Pinj i P2) - from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2]) - with prems show ?case by(cases P2, auto, cases i, auto) -next - case (PX P1 x P2) - from prems have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - with prems show ?case - proof(cases P1) - case (PX p1 y p2) - with prems show ?thesis by(cases x, auto, cases p2, auto) - next - case Pinj - with prems show ?thesis by(cases x, auto) - qed(cases x, auto) -qed(simp) - -text {* sub conserves normalizedness *} -lemma sub_cn:"isnorm p \ isnorm q \ isnorm (p \ q)" -by (simp add: sub_def add_cn neg_cn) - -text {* sqr conserves normalizizedness *} -lemma sqr_cn:"isnorm P \ isnorm (sqr P)" -proof(induct P) - case (Pinj i Q) - from prems show ?case by(cases Q, auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn) -next - case (PX P1 x P2) - from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) - with prems have - "isnorm (mkPX (Pc (1 + 1) \ P1 \ mkPinj (Suc 0) P2) x (Pc 0))" - and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" - by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) - thus ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) -qed simp - -text {* pow conserves normalizedness *} -lemma pow_cn:"isnorm P \ isnorm (pow n P)" -proof (induct n arbitrary: P rule: nat_less_induct) - case (1 k) - show ?case - proof (cases "k=0") - case False - then have K2:"k div 2 < k" by (cases k, auto) - from prems have "isnorm (sqr P)" by (simp add: sqr_cn) - with prems K2 show ?thesis - by (simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn) - qed simp -qed - -end diff -r ee5b5ef5e970 -r 85adf83af7ce src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Fri Oct 30 11:31:34 2009 +0100 +++ b/src/HOL/ex/ROOT.ML Fri Oct 30 14:02:42 2009 +0100 @@ -45,8 +45,6 @@ "Groebner_Examples", "MT", "Unification", - "Commutative_RingEx", - "Commutative_Ring_Complete", "Primrec", "Tarski", "Adder", diff -r ee5b5ef5e970 -r 85adf83af7ce src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Fri Oct 30 11:31:34 2009 +0100 +++ b/src/Pure/Isar/expression.ML Fri Oct 30 14:02:42 2009 +0100 @@ -611,7 +611,7 @@ else raise Match); (* define one predicate including its intro rule and axioms - - bname: predicate name + - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) @@ -619,9 +619,9 @@ - thy: the theory *) -fun def_pred bname parms defs ts norm_ts thy = +fun def_pred binding parms defs ts norm_ts thy = let - val name = Sign.full_name thy bname; + val name = Sign.full_name thy binding; val (body, bodyT, body_eq) = atomize_spec thy norm_ts; val env = Term.add_free_names body []; @@ -639,9 +639,9 @@ val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], []) - |> Sign.declare_const ((bname, predT), NoSyn) |> snd + |> Sign.declare_const ((Binding.conceal binding, predT), NoSyn) |> snd |> PureThy.add_defs false - [((Binding.conceal (Binding.map_name Thm.def_name bname), + [((Binding.conceal (Binding.map_name Thm.def_name binding), Logic.mk_equals (head, body)), [])]; val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head; @@ -667,7 +667,7 @@ (* main predicate definition function *) -fun define_preds pname parms (((exts, exts'), (ints, ints')), defs) thy = +fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val defs' = map (cterm_of thy #> Assumption.assume #> Drule.abs_def) defs; @@ -675,13 +675,13 @@ if null exts then (NONE, NONE, [], thy) else let - val aname = if null ints then pname else Binding.suffix_name ("_" ^ axiomsN) pname; + val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy') = thy - |> def_pred aname parms defs' exts exts'; + |> def_pred abinding parms defs' exts exts'; val (_, thy'') = thy' - |> Sign.mandatory_path (Binding.name_of aname) + |> Sign.mandatory_path (Binding.name_of abinding) |> PureThy.note_thmss Thm.internalK [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.unfold_add])])] ||> Sign.restore_naming thy'; @@ -692,10 +692,10 @@ let val ((statement, intro, axioms), thy''') = thy'' - |> def_pred pname parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); + |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val (_, thy'''') = thy''' - |> Sign.mandatory_path (Binding.name_of pname) + |> Sign.mandatory_path (Binding.name_of binding) |> PureThy.note_thmss Thm.internalK [((Binding.conceal (Binding.name introN), []), [([intro], [Locale.intro_add])]), ((Binding.conceal (Binding.name axiomsN), []), @@ -723,9 +723,9 @@ | defines_to_notes _ e = e; fun gen_add_locale prep_decl - bname raw_predicate_bname raw_import raw_body thy = + binding raw_predicate_binding raw_import raw_body thy = let - val name = Sign.full_name thy bname; + val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); @@ -733,17 +733,17 @@ prep_decl raw_import I raw_body (ProofContext.init thy); val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; - val predicate_bname = - if Binding.is_empty raw_predicate_bname then bname - else raw_predicate_bname; + val predicate_binding = + if Binding.is_empty raw_predicate_binding then binding + else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = - define_preds predicate_bname parms text thy; + define_preds predicate_binding parms text thy; val extraTs = subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []); val _ = if null extraTs then () else warning ("Additional type variable(s) in locale specification " ^ - quote (Binding.str_of bname)); + quote (Binding.str_of binding)); val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; @@ -755,7 +755,7 @@ val notes = if is_some asm then - [(Thm.internalK, [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) bname), []), + [(Thm.internalK, [((Binding.conceal (Binding.suffix_name ("_" ^ axiomsN) binding), []), [([Assumption.assume (cterm_of thy' (the asm))], [(Attrib.internal o K) Locale.witness_add])])])] else []; @@ -772,7 +772,7 @@ val axioms = map Element.conclude_witness b_axioms; val loc_ctxt = thy' - |> Locale.register_locale bname (extraTs, params) + |> Locale.register_locale binding (extraTs, params) (asm, rev defs) (a_intro, b_intro) axioms ([], []) (rev notes) (rev deps') |> TheoryTarget.init (SOME name) |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes';