# HG changeset patch # User wenzelm # Date 1182451727 -7200 # Node ID 8f8835aac29909e711f68aff6d2ec08a94ab9f06 # Parent bc2563c37b1aef9d5d831989f1dffb87f62753bc moved Presburger setup back to Presburger.thy; diff -r bc2563c37b1a -r 8f8835aac299 src/HOL/Arith_Tools.thy --- a/src/HOL/Arith_Tools.thy Thu Jun 21 20:07:26 2007 +0200 +++ b/src/HOL/Arith_Tools.thy Thu Jun 21 20:48:47 2007 +0200 @@ -7,16 +7,12 @@ header {* Setup of arithmetic tools *} theory Arith_Tools -imports Groebner_Basis Dense_Linear_Order SetInterval +imports Groebner_Basis Dense_Linear_Order uses "~~/src/Provers/Arith/cancel_numeral_factor.ML" "~~/src/Provers/Arith/extract_common_term.ML" "int_factor_simprocs.ML" "nat_simprocs.ML" - "Tools/Presburger/cooper_data.ML" - "Tools/Presburger/generated_cooper.ML" - ("Tools/Presburger/cooper.ML") - ("Tools/Presburger/presburger.ML") begin subsection {* Simprocs for the Naturals *} @@ -941,681 +937,4 @@ end *} - -subsection {* Decision Procedure for Presburger Arithmetic *} - -setup CooperData.setup - -subsection{* The @{text "-\"} and @{text "+\"} Properties *} - -lemma minf: - "\\(z ::'a::linorder).\xz.\x - \ \z.\x Q x) = (P' x \ Q' x)" - "\\(z ::'a::linorder).\xz.\x - \ \z.\x Q x) = (P' x \ Q' x)" - "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" - "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" - "\(z ::'a::{linorder}).\x t) = False" - "\(z ::'a::{linorder}).\x t) = False" - "\z.\(x::'a::{linorder,plus,times})z.\(x::'a::{linorder,plus,times}) d dvd x + s) = (\ d dvd x + s)" - "\z.\x\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ - \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" - "\\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ - \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" - "\(z ::'a::{linorder}).\x>z.(x = t) = False" - "\(z ::'a::{linorder}).\x>z.(x \ t) = True" - "\(z ::'a::{linorder}).\x>z.(x < t) = False" - "\(z ::'a::{linorder}).\x>z.(x \ t) = False" - "\(z ::'a::{linorder}).\x>z.(x > t) = True" - "\(z ::'a::{linorder}).\x>z.(x \ t) = True" - "\z.\(x::'a::{linorder,plus,times})>z. (d dvd x + s) = (d dvd x + s)" - "\z.\(x::'a::{linorder,plus,times})>z. (\ d dvd x + s) = (\ d dvd x + s)" - "\z.\x>z. F = F" - by ((erule exE, erule exE,rule_tac x="max z za" in exI,simp)+,(rule_tac x="t" in exI,fastsimp)+) simp_all - -lemma inf_period: - "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ - \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" - "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ - \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" - "(d::'a::{comm_ring}) dvd D \ \x k. (d dvd x + t) = (d dvd (x - k*D) + t)" - "(d::'a::{comm_ring}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" - "\x k. F = F" -by simp_all - (clarsimp simp add: dvd_def, rule iffI, clarsimp,rule_tac x = "kb - ka*k" in exI, - simp add: ring_eq_simps, clarsimp,rule_tac x = "kb + ka*k" in exI,simp add: ring_eq_simps)+ - -section{* The A and B sets *} -lemma bset: - "\\x.(\j \ {1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; - \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ - \x.(\j\{1 .. D}. \b\B. x \ b + j) \ (P x \ Q x) \ (P(x - D) \ Q (x - D))" - "\\x.(\j\{1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; - \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ - \x.(\j\{1 .. D}. \b\B. x \ b + j)\ (P x \ Q x) \ (P(x - D) \ Q (x - D))" - "\D>0; t - 1\ B\ \ (\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" - "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" - "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" - "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" - "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t))" - "\D>0 ; t - 1 \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" - "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t))" - "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\ d dvd (x - D) + t))" - "\x.(\j\{1 .. D}. \b\B. x \ b + j) \ F \ F" -proof (blast, blast) - assume dp: "D > 0" and tB: "t - 1\ B" - show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" - apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t - 1"]) - using dp tB by simp_all -next - assume dp: "D > 0" and tB: "t \ B" - show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" - apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) - using dp tB by simp_all -next - assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" by arith -next - assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by arith -next - assume dp: "D > 0" and tB:"t \ B" - {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x > t" and ng: "\ (x - D) > t" - hence "x -t \ D" and "1 \ x - t" by simp+ - hence "\j \ {1 .. D}. x - t = j" by auto - hence "\j \ {1 .. D}. x = t + j" by (simp add: ring_eq_simps) - with nob tB have "False" by simp} - thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t)" by blast -next - assume dp: "D > 0" and tB:"t - 1\ B" - {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x \ t" and ng: "\ (x - D) \ t" - hence "x - (t - 1) \ D" and "1 \ x - (t - 1)" by simp+ - hence "\j \ {1 .. D}. x - (t - 1) = j" by auto - hence "\j \ {1 .. D}. x = (t - 1) + j" by (simp add: ring_eq_simps) - with nob tB have "False" by simp} - thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by blast -next - assume d: "d dvd D" - {fix x assume H: "d dvd x + t" with d have "d dvd (x - D) + t" - by (clarsimp simp add: dvd_def,rule_tac x= "ka - k" in exI,simp add: ring_eq_simps)} - thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t)" by simp -next - assume d: "d dvd D" - {fix x assume H: "\(d dvd x + t)" with d have "\d dvd (x - D) + t" - by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: ring_eq_simps)} - thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\d dvd (x - D) + t)" by auto -qed blast - -lemma aset: - "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; - \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ - \x.(\j\{1 .. D}. \b\A. x \ b - j) \ (P x \ Q x) \ (P(x + D) \ Q (x + D))" - "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; - \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ - \x.(\j\{1 .. D}. \b\A. x \ b - j)\ (P x \ Q x) \ (P(x + D) \ Q (x + D))" - "\D>0; t + 1\ A\ \ (\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" - "\D>0 ; t \ A\ \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" - "\D>0; t\ A\ \(\(x::int). (\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t))" - "\D>0; t + 1 \ A\ \ (\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" - "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" - "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" - "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t))" - "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\ d dvd (x + D) + t))" - "\x.(\j\{1 .. D}. \b\A. x \ b - j) \ F \ F" -proof (blast, blast) - assume dp: "D > 0" and tA: "t + 1 \ A" - show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" - apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t + 1"]) - using dp tA by simp_all -next - assume dp: "D > 0" and tA: "t \ A" - show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" - apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) - using dp tA by simp_all -next - assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" by arith -next - assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by arith -next - assume dp: "D > 0" and tA:"t \ A" - {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x < t" and ng: "\ (x + D) < t" - hence "t - x \ D" and "1 \ t - x" by simp+ - hence "\j \ {1 .. D}. t - x = j" by auto - hence "\j \ {1 .. D}. x = t - j" by (auto simp add: ring_eq_simps) - with nob tA have "False" by simp} - thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t)" by blast -next - assume dp: "D > 0" and tA:"t + 1\ A" - {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x \ t" and ng: "\ (x + D) \ t" - hence "(t + 1) - x \ D" and "1 \ (t + 1) - x" by (simp_all add: ring_eq_simps) - hence "\j \ {1 .. D}. (t + 1) - x = j" by auto - hence "\j \ {1 .. D}. x = (t + 1) - j" by (auto simp add: ring_eq_simps) - with nob tA have "False" by simp} - thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by blast -next - assume d: "d dvd D" - {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t" - by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: ring_eq_simps)} - thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t)" by simp -next - assume d: "d dvd D" - {fix x assume H: "\(d dvd x + t)" with d have "\d dvd (x + D) + t" - by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: ring_eq_simps)} - thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\d dvd (x + D) + t)" by auto -qed blast - -subsection{* Cooper's Theorem @{text "-\"} and @{text "+\"} Version *} - -subsubsection{* First some trivial facts about periodic sets or predicates *} -lemma periodic_finite_ex: - assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)" - shows "(EX x. P x) = (EX j : {1..d}. P j)" - (is "?LHS = ?RHS") -proof - assume ?LHS - then obtain x where P: "P x" .. - have "x mod d = x - (x div d)*d" by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) - hence Pmod: "P x = P(x mod d)" using modd by simp - show ?RHS - proof (cases) - assume "x mod d = 0" - hence "P 0" using P Pmod by simp - moreover have "P 0 = P(0 - (-1)*d)" using modd by blast - ultimately have "P d" by simp - moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff) - ultimately show ?RHS .. - next - assume not0: "x mod d \ 0" - have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) - moreover have "x mod d : {1..d}" - proof - - from dpos have "0 \ x mod d" by(rule pos_mod_sign) - moreover from dpos have "x mod d < d" by(rule pos_mod_bound) - ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) - qed - ultimately show ?RHS .. - qed -qed auto - -subsubsection{* The @{text "-\"} Version*} - -lemma decr_lemma: "0 < (d::int) \ x - (abs(x-z)+1) * d < z" -by(induct rule: int_gr_induct,simp_all add:int_distrib) - -lemma incr_lemma: "0 < (d::int) \ z < x + (abs(x-z)+1) * d" -by(induct rule: int_gr_induct, simp_all add:int_distrib) - -theorem int_induct[case_names base step1 step2]: - assumes - base: "P(k::int)" and step1: "\i. \k \ i; P i\ \ P(i+1)" and - step2: "\i. \k \ i; P i\ \ P(i - 1)" - shows "P i" -proof - - have "i \ k \ i\ k" by arith - thus ?thesis using prems int_ge_induct[where P="P" and k="k" and i="i"] int_le_induct[where P="P" and k="k" and i="i"] by blast -qed - -lemma decr_mult_lemma: - assumes dpos: "(0::int) < d" and minus: "\x. P x \ P(x - d)" and knneg: "0 <= k" - shows "ALL x. P x \ P(x - k*d)" -using knneg -proof (induct rule:int_ge_induct) - case base thus ?case by simp -next - case (step i) - {fix x - have "P x \ P (x - i * d)" using step.hyps by blast - also have "\ \ P(x - (i + 1) * d)" using minus[THEN spec, of "x - i * d"] - by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric]) - ultimately have "P x \ P(x - (i + 1) * d)" by blast} - thus ?case .. -qed - -lemma minusinfinity: - assumes dpos: "0 < d" and - P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and ePeqP1: "EX z::int. ALL x. x < z \ (P x = P1 x)" - shows "(EX x. P1 x) \ (EX x. P x)" -proof - assume eP1: "EX x. P1 x" - then obtain x where P1: "P1 x" .. - from ePeqP1 obtain z where P1eqP: "ALL x. x < z \ (P x = P1 x)" .. - let ?w = "x - (abs(x-z)+1) * d" - from dpos have w: "?w < z" by(rule decr_lemma) - have "P1 x = P1 ?w" using P1eqP1 by blast - also have "\ = P(?w)" using w P1eqP by blast - finally have "P ?w" using P1 by blast - thus "EX x. P x" .. -qed - -lemma cpmi: - assumes dp: "0 < D" and p1:"\z. \ x< z. P x = P' x" - and nb:"\x.(\ j\ {1..D}. \(b::int) \ B. x \ b+j) --> P (x) --> P (x - D)" - and pd: "\ x k. P' x = P' (x-k*D)" - shows "(\x. P x) = ((\ j\ {1..D} . P' j) | (\ j \ {1..D}.\ b\ B. P (b+j)))" - (is "?L = (?R1 \ ?R2)") -proof- - {assume "?R2" hence "?L" by blast} - moreover - {assume H:"?R1" hence "?L" using minusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} - moreover - { fix x - assume P: "P x" and H: "\ ?R2" - {fix y assume "\ (\j\{1..D}. \b\B. P (b + j))" and P: "P y" - hence "~(EX (j::int) : {1..D}. EX (b::int) : B. y = b+j)" by auto - with nb P have "P (y - D)" by auto } - hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D)" by blast - with H P have th: " \x. P x \ P (x - D)" by auto - from p1 obtain z where z: "ALL x. x < z --> (P x = P' x)" by blast - let ?y = "x - (\x - z\ + 1)*D" - have zp: "0 <= (\x - z\ + 1)" by arith - from dp have yz: "?y < z" using decr_lemma[OF dp] by simp - from z[rule_format, OF yz] decr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto - with periodic_finite_ex[OF dp pd] - have "?R1" by blast} - ultimately show ?thesis by blast -qed - -subsubsection {* The @{text "+\"} Version*} - -lemma plusinfinity: - assumes dpos: "(0::int) < d" and - P1eqP1: "\x k. P' x = P'(x - k*d)" and ePeqP1: "\ z. \ x>z. P x = P' x" - shows "(\ x. P' x) \ (\ x. P x)" -proof - assume eP1: "EX x. P' x" - then obtain x where P1: "P' x" .. - from ePeqP1 obtain z where P1eqP: "\x>z. P x = P' x" .. - let ?w' = "x + (abs(x-z)+1) * d" - let ?w = "x - (-(abs(x-z) + 1))*d" - have ww'[simp]: "?w = ?w'" by (simp add: ring_eq_simps) - from dpos have w: "?w > z" by(simp only: ww' incr_lemma) - hence "P' x = P' ?w" using P1eqP1 by blast - also have "\ = P(?w)" using w P1eqP by blast - finally have "P ?w" using P1 by blast - thus "EX x. P x" .. -qed - -lemma incr_mult_lemma: - assumes dpos: "(0::int) < d" and plus: "ALL x::int. P x \ P(x + d)" and knneg: "0 <= k" - shows "ALL x. P x \ P(x + k*d)" -using knneg -proof (induct rule:int_ge_induct) - case base thus ?case by simp -next - case (step i) - {fix x - have "P x \ P (x + i * d)" using step.hyps by blast - also have "\ \ P(x + (i + 1) * d)" using plus[THEN spec, of "x + i * d"] - by (simp add:int_distrib zadd_ac) - ultimately have "P x \ P(x + (i + 1) * d)" by blast} - thus ?case .. -qed - -lemma cppi: - assumes dp: "0 < D" and p1:"\z. \ x> z. P x = P' x" - and nb:"\x.(\ j\ {1..D}. \(b::int) \ A. x \ b - j) --> P (x) --> P (x + D)" - and pd: "\ x k. P' x= P' (x-k*D)" - shows "(\x. P x) = ((\ j\ {1..D} . P' j) | (\ j \ {1..D}.\ b\ A. P (b - j)))" (is "?L = (?R1 \ ?R2)") -proof- - {assume "?R2" hence "?L" by blast} - moreover - {assume H:"?R1" hence "?L" using plusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} - moreover - { fix x - assume P: "P x" and H: "\ ?R2" - {fix y assume "\ (\j\{1..D}. \b\A. P (b - j))" and P: "P y" - hence "~(EX (j::int) : {1..D}. EX (b::int) : A. y = b - j)" by auto - with nb P have "P (y + D)" by auto } - hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : A. P(b-j)) --> P (x) --> P (x + D)" by blast - with H P have th: " \x. P x \ P (x + D)" by auto - from p1 obtain z where z: "ALL x. x > z --> (P x = P' x)" by blast - let ?y = "x + (\x - z\ + 1)*D" - have zp: "0 <= (\x - z\ + 1)" by arith - from dp have yz: "?y > z" using incr_lemma[OF dp] by simp - from z[rule_format, OF yz] incr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto - with periodic_finite_ex[OF dp pd] - have "?R1" by blast} - ultimately show ?thesis by blast -qed - -lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})" -apply(simp add:atLeastAtMost_def atLeast_def atMost_def) -apply(fastsimp) -done - -theorem unity_coeff_ex: "(\(x::'a::{semiring_0}). P (l * x)) \ (\x. l dvd (x + 0) \ P x)" - apply (rule eq_reflection[symmetric]) - apply (rule iffI) - defer - apply (erule exE) - apply (rule_tac x = "l * x" in exI) - apply (simp add: dvd_def) - apply (rule_tac x="x" in exI, simp) - apply (erule exE) - apply (erule conjE) - apply (erule dvdE) - apply (rule_tac x = k in exI) - apply simp - done - -lemma zdvd_mono: assumes not0: "(k::int) \ 0" -shows "((m::int) dvd t) \ (k*m dvd k*t)" - using not0 by (simp add: dvd_def) - -lemma uminus_dvd_conv: "(d dvd (t::int)) \ (-d dvd t)" "(d dvd (t::int)) \ (d dvd -t)" - by simp_all -text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*} -lemma all_nat: "(\x::nat. P x) = (\x::int. 0 <= x \ P (nat x))" - by (simp split add: split_nat) - -lemma ex_nat: "(\x::nat. P x) = (\x::int. 0 <= x \ P (nat x))" - apply (auto split add: split_nat) - apply (rule_tac x="int x" in exI, simp) - apply (rule_tac x = "nat x" in exI,erule_tac x = "nat x" in allE, simp) - done - -lemma zdiff_int_split: "P (int (x - y)) = - ((y \ x \ P (int x - int y)) \ (x < y \ P 0))" - by (case_tac "y \ x", simp_all add: zdiff_int) - -lemma number_of1: "(0::int) <= number_of n \ (0::int) <= number_of (n BIT b)" by simp -lemma number_of2: "(0::int) <= Numeral0" by simp -lemma Suc_plus1: "Suc n = n + 1" by simp - -text {* - \medskip Specific instances of congruence rules, to prevent - simplifier from looping. *} - -theorem imp_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" by simp - -theorem conj_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" - by (simp cong: conj_cong) -lemma int_eq_number_of_eq: - "(((number_of v)::int) = (number_of w)) = iszero ((number_of (v + (uminus w)))::int)" - by simp - -lemma mod_eq0_dvd_iff[presburger]: "(m::nat) mod n = 0 \ n dvd m" -unfolding dvd_eq_mod_eq_0[symmetric] .. - -lemma zmod_eq0_zdvd_iff[presburger]: "(m::int) mod n = 0 \ n dvd m" -unfolding zdvd_iff_zmod_eq_0[symmetric] .. -declare mod_1[presburger] -declare mod_0[presburger] -declare zmod_1[presburger] -declare zmod_zero[presburger] -declare zmod_self[presburger] -declare mod_self[presburger] -declare DIVISION_BY_ZERO_MOD[presburger] -declare nat_mod_div_trivial[presburger] -declare div_mod_equality2[presburger] -declare div_mod_equality[presburger] -declare mod_div_equality2[presburger] -declare mod_div_equality[presburger] -declare mod_mult_self1[presburger] -declare mod_mult_self2[presburger] -declare zdiv_zmod_equality2[presburger] -declare zdiv_zmod_equality[presburger] -declare mod2_Suc_Suc[presburger] -lemma [presburger]: "(a::int) div 0 = 0" and [presburger]: "a mod 0 = a" -using IntDiv.DIVISION_BY_ZERO by blast+ - -use "Tools/Presburger/cooper.ML" -oracle linzqe_oracle ("term") = Coopereif.cooper_oracle - -use "Tools/Presburger/presburger.ML" - -setup {* - arith_tactic_add - (mk_arith_tactic "presburger" (fn i => fn st => - (warning "Trying Presburger arithmetic ..."; - Presburger.cooper_tac true [] [] ((ProofContext.init o theory_of_thm) st) i st))) - (* FIXME!!!!!!! get the right context!!*) -*} - -method_setup presburger = {* -let - fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () - fun simple_keyword k = Scan.lift (Args.$$$ k) >> K () - val addN = "add" - val delN = "del" - val elimN = "elim" - val any_keyword = keyword addN || keyword delN || simple_keyword elimN - val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; -in - fn src => Method.syntax - ((Scan.optional (simple_keyword elimN >> K false) true) -- - (Scan.optional (keyword addN |-- thms) []) -- - (Scan.optional (keyword delN |-- thms) [])) src - #> (fn (((elim, add_ths), del_ths),ctxt) => - Method.SIMPLE_METHOD' (Presburger.cooper_tac elim add_ths del_ths ctxt)) end -*} "Cooper's algorithm for Presburger arithmetic" - -lemma [presburger]: "m mod 2 = (1::nat) \ \ 2 dvd m " by presburger -lemma [presburger]: "m mod 2 = Suc 0 \ \ 2 dvd m " by presburger -lemma [presburger]: "m mod (Suc (Suc 0)) = (1::nat) \ \ 2 dvd m " by presburger -lemma [presburger]: "m mod (Suc (Suc 0)) = Suc 0 \ \ 2 dvd m " by presburger -lemma [presburger]: "m mod 2 = (1::int) \ \ 2 dvd m " by presburger - - -subsection {* Code generator setup *} - -text {* - Presburger arithmetic is convenient to prove some - of the following code lemmas on integer numerals: -*} - -lemma eq_Pls_Pls: - "Numeral.Pls = Numeral.Pls \ True" by presburger - -lemma eq_Pls_Min: - "Numeral.Pls = Numeral.Min \ False" - unfolding Pls_def Numeral.Min_def by presburger - -lemma eq_Pls_Bit0: - "Numeral.Pls = Numeral.Bit k bit.B0 \ Numeral.Pls = k" - unfolding Pls_def Bit_def bit.cases by presburger - -lemma eq_Pls_Bit1: - "Numeral.Pls = Numeral.Bit k bit.B1 \ False" - unfolding Pls_def Bit_def bit.cases by presburger - -lemma eq_Min_Pls: - "Numeral.Min = Numeral.Pls \ False" - unfolding Pls_def Numeral.Min_def by presburger - -lemma eq_Min_Min: - "Numeral.Min = Numeral.Min \ True" by presburger - -lemma eq_Min_Bit0: - "Numeral.Min = Numeral.Bit k bit.B0 \ False" - unfolding Numeral.Min_def Bit_def bit.cases by presburger - -lemma eq_Min_Bit1: - "Numeral.Min = Numeral.Bit k bit.B1 \ Numeral.Min = k" - unfolding Numeral.Min_def Bit_def bit.cases by presburger - -lemma eq_Bit0_Pls: - "Numeral.Bit k bit.B0 = Numeral.Pls \ Numeral.Pls = k" - unfolding Pls_def Bit_def bit.cases by presburger - -lemma eq_Bit1_Pls: - "Numeral.Bit k bit.B1 = Numeral.Pls \ False" - unfolding Pls_def Bit_def bit.cases by presburger - -lemma eq_Bit0_Min: - "Numeral.Bit k bit.B0 = Numeral.Min \ False" - unfolding Numeral.Min_def Bit_def bit.cases by presburger - -lemma eq_Bit1_Min: - "(Numeral.Bit k bit.B1) = Numeral.Min \ Numeral.Min = k" - unfolding Numeral.Min_def Bit_def bit.cases by presburger - -lemma eq_Bit_Bit: - "Numeral.Bit k1 v1 = Numeral.Bit k2 v2 \ - v1 = v2 \ k1 = k2" - unfolding Bit_def - apply (cases v1) - apply (cases v2) - apply auto - apply presburger - apply (cases v2) - apply auto - apply presburger - apply (cases v2) - apply auto - done - -lemma eq_number_of: - "(number_of k \ int) = number_of l \ k = l" - unfolding number_of_is_id .. - - -lemma less_eq_Pls_Pls: - "Numeral.Pls \ Numeral.Pls \ True" by rule+ - -lemma less_eq_Pls_Min: - "Numeral.Pls \ Numeral.Min \ False" - unfolding Pls_def Numeral.Min_def by presburger - -lemma less_eq_Pls_Bit: - "Numeral.Pls \ Numeral.Bit k v \ Numeral.Pls \ k" - unfolding Pls_def Bit_def by (cases v) auto - -lemma less_eq_Min_Pls: - "Numeral.Min \ Numeral.Pls \ True" - unfolding Pls_def Numeral.Min_def by presburger - -lemma less_eq_Min_Min: - "Numeral.Min \ Numeral.Min \ True" by rule+ - -lemma less_eq_Min_Bit0: - "Numeral.Min \ Numeral.Bit k bit.B0 \ Numeral.Min < k" - unfolding Numeral.Min_def Bit_def by auto - -lemma less_eq_Min_Bit1: - "Numeral.Min \ Numeral.Bit k bit.B1 \ Numeral.Min \ k" - unfolding Numeral.Min_def Bit_def by auto - -lemma less_eq_Bit0_Pls: - "Numeral.Bit k bit.B0 \ Numeral.Pls \ k \ Numeral.Pls" - unfolding Pls_def Bit_def by simp - -lemma less_eq_Bit1_Pls: - "Numeral.Bit k bit.B1 \ Numeral.Pls \ k < Numeral.Pls" - unfolding Pls_def Bit_def by auto - -lemma less_eq_Bit_Min: - "Numeral.Bit k v \ Numeral.Min \ k \ Numeral.Min" - unfolding Numeral.Min_def Bit_def by (cases v) auto - -lemma less_eq_Bit0_Bit: - "Numeral.Bit k1 bit.B0 \ Numeral.Bit k2 v \ k1 \ k2" - unfolding Bit_def bit.cases by (cases v) auto - -lemma less_eq_Bit_Bit1: - "Numeral.Bit k1 v \ Numeral.Bit k2 bit.B1 \ k1 \ k2" - unfolding Bit_def bit.cases by (cases v) auto - -lemma less_eq_Bit1_Bit0: - "Numeral.Bit k1 bit.B1 \ Numeral.Bit k2 bit.B0 \ k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_eq_number_of: - "(number_of k \ int) \ number_of l \ k \ l" - unfolding number_of_is_id .. - - -lemma less_Pls_Pls: - "Numeral.Pls < Numeral.Pls \ False" by simp - -lemma less_Pls_Min: - "Numeral.Pls < Numeral.Min \ False" - unfolding Pls_def Numeral.Min_def by presburger - -lemma less_Pls_Bit0: - "Numeral.Pls < Numeral.Bit k bit.B0 \ Numeral.Pls < k" - unfolding Pls_def Bit_def by auto - -lemma less_Pls_Bit1: - "Numeral.Pls < Numeral.Bit k bit.B1 \ Numeral.Pls \ k" - unfolding Pls_def Bit_def by auto - -lemma less_Min_Pls: - "Numeral.Min < Numeral.Pls \ True" - unfolding Pls_def Numeral.Min_def by presburger - -lemma less_Min_Min: - "Numeral.Min < Numeral.Min \ False" by simp - -lemma less_Min_Bit: - "Numeral.Min < Numeral.Bit k v \ Numeral.Min < k" - unfolding Numeral.Min_def Bit_def by (auto split: bit.split) - -lemma less_Bit_Pls: - "Numeral.Bit k v < Numeral.Pls \ k < Numeral.Pls" - unfolding Pls_def Bit_def by (auto split: bit.split) - -lemma less_Bit0_Min: - "Numeral.Bit k bit.B0 < Numeral.Min \ k \ Numeral.Min" - unfolding Numeral.Min_def Bit_def by auto - -lemma less_Bit1_Min: - "Numeral.Bit k bit.B1 < Numeral.Min \ k < Numeral.Min" - unfolding Numeral.Min_def Bit_def by auto - -lemma less_Bit_Bit0: - "Numeral.Bit k1 v < Numeral.Bit k2 bit.B0 \ k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_Bit1_Bit: - "Numeral.Bit k1 bit.B1 < Numeral.Bit k2 v \ k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_Bit0_Bit1: - "Numeral.Bit k1 bit.B0 < Numeral.Bit k2 bit.B1 \ k1 \ k2" - unfolding Bit_def bit.cases by arith - -lemma less_number_of: - "(number_of k \ int) < number_of l \ k < l" - unfolding number_of_is_id .. - -lemmas pred_succ_numeral_code [code func] = - arith_simps(5-12) - -lemmas plus_numeral_code [code func] = - arith_simps(13-17) - arith_simps(26-27) - arith_extra_simps(1) [where 'a = int] - -lemmas minus_numeral_code [code func] = - arith_simps(18-21) - arith_extra_simps(2) [where 'a = int] - arith_extra_simps(5) [where 'a = int] - -lemmas times_numeral_code [code func] = - arith_simps(22-25) - arith_extra_simps(4) [where 'a = int] - -lemmas eq_numeral_code [code func] = - eq_Pls_Pls eq_Pls_Min eq_Pls_Bit0 eq_Pls_Bit1 - eq_Min_Pls eq_Min_Min eq_Min_Bit0 eq_Min_Bit1 - eq_Bit0_Pls eq_Bit1_Pls eq_Bit0_Min eq_Bit1_Min eq_Bit_Bit - eq_number_of - -lemmas less_eq_numeral_code [code func] = less_eq_Pls_Pls less_eq_Pls_Min less_eq_Pls_Bit - less_eq_Min_Pls less_eq_Min_Min less_eq_Min_Bit0 less_eq_Min_Bit1 - less_eq_Bit0_Pls less_eq_Bit1_Pls less_eq_Bit_Min less_eq_Bit0_Bit less_eq_Bit_Bit1 less_eq_Bit1_Bit0 - less_eq_number_of - -lemmas less_numeral_code [code func] = less_Pls_Pls less_Pls_Min less_Pls_Bit0 - less_Pls_Bit1 less_Min_Pls less_Min_Min less_Min_Bit less_Bit_Pls - less_Bit0_Min less_Bit1_Min less_Bit_Bit0 less_Bit1_Bit less_Bit0_Bit1 - less_number_of - -end diff -r bc2563c37b1a -r 8f8835aac299 src/HOL/PreList.thy --- a/src/HOL/PreList.thy Thu Jun 21 20:07:26 2007 +0200 +++ b/src/HOL/PreList.thy Thu Jun 21 20:48:47 2007 +0200 @@ -7,7 +7,7 @@ header {* A Basis for Building the Theory of Lists *} theory PreList -imports Wellfounded_Relations Arith_Tools Relation_Power SAT +imports Wellfounded_Relations Presburger Relation_Power SAT FunDef Recdef Extraction begin diff -r bc2563c37b1a -r 8f8835aac299 src/HOL/Presburger.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Presburger.thy Thu Jun 21 20:48:47 2007 +0200 @@ -0,0 +1,691 @@ +(* Title: HOL/Presburger.thy + ID: $Id$ + Author: Amine Chaieb, TU Muenchen +*) + +theory Presburger +imports Arith_Tools SetInterval +uses + "Tools/Qelim/cooper_data.ML" + "Tools/Qelim/generated_cooper.ML" + ("Tools/Qelim/cooper.ML") + ("Tools/Qelim/presburger.ML") +begin + +subsection {* Decision Procedure for Presburger Arithmetic *} + +setup CooperData.setup + +subsection{* The @{text "-\"} and @{text "+\"} Properties *} + +lemma minf: + "\\(z ::'a::linorder).\xz.\x + \ \z.\x Q x) = (P' x \ Q' x)" + "\\(z ::'a::linorder).\xz.\x + \ \z.\x Q x) = (P' x \ Q' x)" + "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" + "\(z ::'a::{linorder}).\x(z ::'a::{linorder}).\x t) = True" + "\(z ::'a::{linorder}).\x t) = False" + "\(z ::'a::{linorder}).\x t) = False" + "\z.\(x::'a::{linorder,plus,times})z.\(x::'a::{linorder,plus,times}) d dvd x + s) = (\ d dvd x + s)" + "\z.\x\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ + \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" + "\\(z ::'a::linorder).\x>z. P x = P' x; \z.\x>z. Q x = Q' x\ + \ \z.\x>z. (P x \ Q x) = (P' x \ Q' x)" + "\(z ::'a::{linorder}).\x>z.(x = t) = False" + "\(z ::'a::{linorder}).\x>z.(x \ t) = True" + "\(z ::'a::{linorder}).\x>z.(x < t) = False" + "\(z ::'a::{linorder}).\x>z.(x \ t) = False" + "\(z ::'a::{linorder}).\x>z.(x > t) = True" + "\(z ::'a::{linorder}).\x>z.(x \ t) = True" + "\z.\(x::'a::{linorder,plus,times})>z. (d dvd x + s) = (d dvd x + s)" + "\z.\(x::'a::{linorder,plus,times})>z. (\ d dvd x + s) = (\ d dvd x + s)" + "\z.\x>z. F = F" + by ((erule exE, erule exE,rule_tac x="max z za" in exI,simp)+,(rule_tac x="t" in exI,fastsimp)+) simp_all + +lemma inf_period: + "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ + \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" + "\\x k. P x = P (x - k*D); \x k. Q x = Q (x - k*D)\ + \ \x k. (P x \ Q x) = (P (x - k*D) \ Q (x - k*D))" + "(d::'a::{comm_ring}) dvd D \ \x k. (d dvd x + t) = (d dvd (x - k*D) + t)" + "(d::'a::{comm_ring}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" + "\x k. F = F" +by simp_all + (clarsimp simp add: dvd_def, rule iffI, clarsimp,rule_tac x = "kb - ka*k" in exI, + simp add: ring_eq_simps, clarsimp,rule_tac x = "kb + ka*k" in exI,simp add: ring_eq_simps)+ + +section{* The A and B sets *} +lemma bset: + "\\x.(\j \ {1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; + \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ + \x.(\j\{1 .. D}. \b\B. x \ b + j) \ (P x \ Q x) \ (P(x - D) \ Q (x - D))" + "\\x.(\j\{1 .. D}. \b\B. x \ b + j)\ P x \ P(x - D) ; + \x.(\j\{1 .. D}. \b\B. x \ b + j)\ Q x \ Q(x - D)\ \ + \x.(\j\{1 .. D}. \b\B. x \ b + j)\ (P x \ Q x) \ (P(x - D) \ Q (x - D))" + "\D>0; t - 1\ B\ \ (\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" + "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" + "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" + "D>0 \ (\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" + "\D>0 ; t \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t))" + "\D>0 ; t - 1 \ B\ \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" + "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t))" + "d dvd D \(\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\ d dvd (x - D) + t))" + "\x.(\j\{1 .. D}. \b\B. x \ b + j) \ F \ F" +proof (blast, blast) + assume dp: "D > 0" and tB: "t - 1\ B" + show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x = t) \ (x - D = t))" + apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t - 1"]) + using dp tB by simp_all +next + assume dp: "D > 0" and tB: "t \ B" + show "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t))" + apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) + using dp tB by simp_all +next + assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x < t) \ (x - D < t))" by arith +next + assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by arith +next + assume dp: "D > 0" and tB:"t \ B" + {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x > t" and ng: "\ (x - D) > t" + hence "x -t \ D" and "1 \ x - t" by simp+ + hence "\j \ {1 .. D}. x - t = j" by auto + hence "\j \ {1 .. D}. x = t + j" by (simp add: ring_eq_simps) + with nob tB have "False" by simp} + thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x > t) \ (x - D > t)" by blast +next + assume dp: "D > 0" and tB:"t - 1\ B" + {fix x assume nob: "\j\{1 .. D}. \b\B. x \ b + j" and g: "x \ t" and ng: "\ (x - D) \ t" + hence "x - (t - 1) \ D" and "1 \ x - (t - 1)" by simp+ + hence "\j \ {1 .. D}. x - (t - 1) = j" by auto + hence "\j \ {1 .. D}. x = (t - 1) + j" by (simp add: ring_eq_simps) + with nob tB have "False" by simp} + thus "\x.(\j\{1 .. D}. \b\B. x \ b + j)\ (x \ t) \ (x - D \ t)" by blast +next + assume d: "d dvd D" + {fix x assume H: "d dvd x + t" with d have "d dvd (x - D) + t" + by (clarsimp simp add: dvd_def,rule_tac x= "ka - k" in exI,simp add: ring_eq_simps)} + thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (d dvd x+t) \ (d dvd (x - D) + t)" by simp +next + assume d: "d dvd D" + {fix x assume H: "\(d dvd x + t)" with d have "\d dvd (x - D) + t" + by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: ring_eq_simps)} + thus "\(x::int).(\j\{1 .. D}. \b\B. x \ b + j)\ (\d dvd x+t) \ (\d dvd (x - D) + t)" by auto +qed blast + +lemma aset: + "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; + \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ + \x.(\j\{1 .. D}. \b\A. x \ b - j) \ (P x \ Q x) \ (P(x + D) \ Q (x + D))" + "\\x.(\j\{1 .. D}. \b\A. x \ b - j)\ P x \ P(x + D) ; + \x.(\j\{1 .. D}. \b\A. x \ b - j)\ Q x \ Q(x + D)\ \ + \x.(\j\{1 .. D}. \b\A. x \ b - j)\ (P x \ Q x) \ (P(x + D) \ Q (x + D))" + "\D>0; t + 1\ A\ \ (\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" + "\D>0 ; t \ A\ \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" + "\D>0; t\ A\ \(\(x::int). (\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t))" + "\D>0; t + 1 \ A\ \ (\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" + "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" + "D>0 \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" + "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t))" + "d dvd D \(\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\ d dvd (x + D) + t))" + "\x.(\j\{1 .. D}. \b\A. x \ b - j) \ F \ F" +proof (blast, blast) + assume dp: "D > 0" and tA: "t + 1 \ A" + show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x = t) \ (x + D = t))" + apply (rule allI, rule impI,erule ballE[where x="1"],erule ballE[where x="t + 1"]) + using dp tA by simp_all +next + assume dp: "D > 0" and tA: "t \ A" + show "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t))" + apply (rule allI, rule impI,erule ballE[where x="D"],erule ballE[where x="t"]) + using dp tA by simp_all +next + assume dp: "D > 0" thus "(\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x > t) \ (x + D > t))" by arith +next + assume dp: "D > 0" thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by arith +next + assume dp: "D > 0" and tA:"t \ A" + {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x < t" and ng: "\ (x + D) < t" + hence "t - x \ D" and "1 \ t - x" by simp+ + hence "\j \ {1 .. D}. t - x = j" by auto + hence "\j \ {1 .. D}. x = t - j" by (auto simp add: ring_eq_simps) + with nob tA have "False" by simp} + thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x < t) \ (x + D < t)" by blast +next + assume dp: "D > 0" and tA:"t + 1\ A" + {fix x assume nob: "\j\{1 .. D}. \b\A. x \ b - j" and g: "x \ t" and ng: "\ (x + D) \ t" + hence "(t + 1) - x \ D" and "1 \ (t + 1) - x" by (simp_all add: ring_eq_simps) + hence "\j \ {1 .. D}. (t + 1) - x = j" by auto + hence "\j \ {1 .. D}. x = (t + 1) - j" by (auto simp add: ring_eq_simps) + with nob tA have "False" by simp} + thus "\x.(\j\{1 .. D}. \b\A. x \ b - j)\ (x \ t) \ (x + D \ t)" by blast +next + assume d: "d dvd D" + {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t" + by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: ring_eq_simps)} + thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (d dvd x+t) \ (d dvd (x + D) + t)" by simp +next + assume d: "d dvd D" + {fix x assume H: "\(d dvd x + t)" with d have "\d dvd (x + D) + t" + by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: ring_eq_simps)} + thus "\(x::int).(\j\{1 .. D}. \b\A. x \ b - j)\ (\d dvd x+t) \ (\d dvd (x + D) + t)" by auto +qed blast + +subsection{* Cooper's Theorem @{text "-\"} and @{text "+\"} Version *} + +subsubsection{* First some trivial facts about periodic sets or predicates *} +lemma periodic_finite_ex: + assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)" + shows "(EX x. P x) = (EX j : {1..d}. P j)" + (is "?LHS = ?RHS") +proof + assume ?LHS + then obtain x where P: "P x" .. + have "x mod d = x - (x div d)*d" by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) + hence Pmod: "P x = P(x mod d)" using modd by simp + show ?RHS + proof (cases) + assume "x mod d = 0" + hence "P 0" using P Pmod by simp + moreover have "P 0 = P(0 - (-1)*d)" using modd by blast + ultimately have "P d" by simp + moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff) + ultimately show ?RHS .. + next + assume not0: "x mod d \ 0" + have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) + moreover have "x mod d : {1..d}" + proof - + from dpos have "0 \ x mod d" by(rule pos_mod_sign) + moreover from dpos have "x mod d < d" by(rule pos_mod_bound) + ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) + qed + ultimately show ?RHS .. + qed +qed auto + +subsubsection{* The @{text "-\"} Version*} + +lemma decr_lemma: "0 < (d::int) \ x - (abs(x-z)+1) * d < z" +by(induct rule: int_gr_induct,simp_all add:int_distrib) + +lemma incr_lemma: "0 < (d::int) \ z < x + (abs(x-z)+1) * d" +by(induct rule: int_gr_induct, simp_all add:int_distrib) + +theorem int_induct[case_names base step1 step2]: + assumes + base: "P(k::int)" and step1: "\i. \k \ i; P i\ \ P(i+1)" and + step2: "\i. \k \ i; P i\ \ P(i - 1)" + shows "P i" +proof - + have "i \ k \ i\ k" by arith + thus ?thesis using prems int_ge_induct[where P="P" and k="k" and i="i"] int_le_induct[where P="P" and k="k" and i="i"] by blast +qed + +lemma decr_mult_lemma: + assumes dpos: "(0::int) < d" and minus: "\x. P x \ P(x - d)" and knneg: "0 <= k" + shows "ALL x. P x \ P(x - k*d)" +using knneg +proof (induct rule:int_ge_induct) + case base thus ?case by simp +next + case (step i) + {fix x + have "P x \ P (x - i * d)" using step.hyps by blast + also have "\ \ P(x - (i + 1) * d)" using minus[THEN spec, of "x - i * d"] + by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric]) + ultimately have "P x \ P(x - (i + 1) * d)" by blast} + thus ?case .. +qed + +lemma minusinfinity: + assumes dpos: "0 < d" and + P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and ePeqP1: "EX z::int. ALL x. x < z \ (P x = P1 x)" + shows "(EX x. P1 x) \ (EX x. P x)" +proof + assume eP1: "EX x. P1 x" + then obtain x where P1: "P1 x" .. + from ePeqP1 obtain z where P1eqP: "ALL x. x < z \ (P x = P1 x)" .. + let ?w = "x - (abs(x-z)+1) * d" + from dpos have w: "?w < z" by(rule decr_lemma) + have "P1 x = P1 ?w" using P1eqP1 by blast + also have "\ = P(?w)" using w P1eqP by blast + finally have "P ?w" using P1 by blast + thus "EX x. P x" .. +qed + +lemma cpmi: + assumes dp: "0 < D" and p1:"\z. \ x< z. P x = P' x" + and nb:"\x.(\ j\ {1..D}. \(b::int) \ B. x \ b+j) --> P (x) --> P (x - D)" + and pd: "\ x k. P' x = P' (x-k*D)" + shows "(\x. P x) = ((\ j\ {1..D} . P' j) | (\ j \ {1..D}.\ b\ B. P (b+j)))" + (is "?L = (?R1 \ ?R2)") +proof- + {assume "?R2" hence "?L" by blast} + moreover + {assume H:"?R1" hence "?L" using minusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} + moreover + { fix x + assume P: "P x" and H: "\ ?R2" + {fix y assume "\ (\j\{1..D}. \b\B. P (b + j))" and P: "P y" + hence "~(EX (j::int) : {1..D}. EX (b::int) : B. y = b+j)" by auto + with nb P have "P (y - D)" by auto } + hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D)" by blast + with H P have th: " \x. P x \ P (x - D)" by auto + from p1 obtain z where z: "ALL x. x < z --> (P x = P' x)" by blast + let ?y = "x - (\x - z\ + 1)*D" + have zp: "0 <= (\x - z\ + 1)" by arith + from dp have yz: "?y < z" using decr_lemma[OF dp] by simp + from z[rule_format, OF yz] decr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto + with periodic_finite_ex[OF dp pd] + have "?R1" by blast} + ultimately show ?thesis by blast +qed + +subsubsection {* The @{text "+\"} Version*} + +lemma plusinfinity: + assumes dpos: "(0::int) < d" and + P1eqP1: "\x k. P' x = P'(x - k*d)" and ePeqP1: "\ z. \ x>z. P x = P' x" + shows "(\ x. P' x) \ (\ x. P x)" +proof + assume eP1: "EX x. P' x" + then obtain x where P1: "P' x" .. + from ePeqP1 obtain z where P1eqP: "\x>z. P x = P' x" .. + let ?w' = "x + (abs(x-z)+1) * d" + let ?w = "x - (-(abs(x-z) + 1))*d" + have ww'[simp]: "?w = ?w'" by (simp add: ring_eq_simps) + from dpos have w: "?w > z" by(simp only: ww' incr_lemma) + hence "P' x = P' ?w" using P1eqP1 by blast + also have "\ = P(?w)" using w P1eqP by blast + finally have "P ?w" using P1 by blast + thus "EX x. P x" .. +qed + +lemma incr_mult_lemma: + assumes dpos: "(0::int) < d" and plus: "ALL x::int. P x \ P(x + d)" and knneg: "0 <= k" + shows "ALL x. P x \ P(x + k*d)" +using knneg +proof (induct rule:int_ge_induct) + case base thus ?case by simp +next + case (step i) + {fix x + have "P x \ P (x + i * d)" using step.hyps by blast + also have "\ \ P(x + (i + 1) * d)" using plus[THEN spec, of "x + i * d"] + by (simp add:int_distrib zadd_ac) + ultimately have "P x \ P(x + (i + 1) * d)" by blast} + thus ?case .. +qed + +lemma cppi: + assumes dp: "0 < D" and p1:"\z. \ x> z. P x = P' x" + and nb:"\x.(\ j\ {1..D}. \(b::int) \ A. x \ b - j) --> P (x) --> P (x + D)" + and pd: "\ x k. P' x= P' (x-k*D)" + shows "(\x. P x) = ((\ j\ {1..D} . P' j) | (\ j \ {1..D}.\ b\ A. P (b - j)))" (is "?L = (?R1 \ ?R2)") +proof- + {assume "?R2" hence "?L" by blast} + moreover + {assume H:"?R1" hence "?L" using plusinfinity[OF dp pd p1] periodic_finite_ex[OF dp pd] by simp} + moreover + { fix x + assume P: "P x" and H: "\ ?R2" + {fix y assume "\ (\j\{1..D}. \b\A. P (b - j))" and P: "P y" + hence "~(EX (j::int) : {1..D}. EX (b::int) : A. y = b - j)" by auto + with nb P have "P (y + D)" by auto } + hence "ALL x.~(EX (j::int) : {1..D}. EX (b::int) : A. P(b-j)) --> P (x) --> P (x + D)" by blast + with H P have th: " \x. P x \ P (x + D)" by auto + from p1 obtain z where z: "ALL x. x > z --> (P x = P' x)" by blast + let ?y = "x + (\x - z\ + 1)*D" + have zp: "0 <= (\x - z\ + 1)" by arith + from dp have yz: "?y > z" using incr_lemma[OF dp] by simp + from z[rule_format, OF yz] incr_mult_lemma[OF dp th zp, rule_format, OF P] have th2: " P' ?y" by auto + with periodic_finite_ex[OF dp pd] + have "?R1" by blast} + ultimately show ?thesis by blast +qed + +lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})" +apply(simp add:atLeastAtMost_def atLeast_def atMost_def) +apply(fastsimp) +done + +theorem unity_coeff_ex: "(\(x::'a::{semiring_0}). P (l * x)) \ (\x. l dvd (x + 0) \ P x)" + apply (rule eq_reflection[symmetric]) + apply (rule iffI) + defer + apply (erule exE) + apply (rule_tac x = "l * x" in exI) + apply (simp add: dvd_def) + apply (rule_tac x="x" in exI, simp) + apply (erule exE) + apply (erule conjE) + apply (erule dvdE) + apply (rule_tac x = k in exI) + apply simp + done + +lemma zdvd_mono: assumes not0: "(k::int) \ 0" +shows "((m::int) dvd t) \ (k*m dvd k*t)" + using not0 by (simp add: dvd_def) + +lemma uminus_dvd_conv: "(d dvd (t::int)) \ (-d dvd t)" "(d dvd (t::int)) \ (d dvd -t)" + by simp_all +text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*} +lemma all_nat: "(\x::nat. P x) = (\x::int. 0 <= x \ P (nat x))" + by (simp split add: split_nat) + +lemma ex_nat: "(\x::nat. P x) = (\x::int. 0 <= x \ P (nat x))" + apply (auto split add: split_nat) + apply (rule_tac x="int x" in exI, simp) + apply (rule_tac x = "nat x" in exI,erule_tac x = "nat x" in allE, simp) + done + +lemma zdiff_int_split: "P (int (x - y)) = + ((y \ x \ P (int x - int y)) \ (x < y \ P 0))" + by (case_tac "y \ x", simp_all add: zdiff_int) + +lemma number_of1: "(0::int) <= number_of n \ (0::int) <= number_of (n BIT b)" by simp +lemma number_of2: "(0::int) <= Numeral0" by simp +lemma Suc_plus1: "Suc n = n + 1" by simp + +text {* + \medskip Specific instances of congruence rules, to prevent + simplifier from looping. *} + +theorem imp_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" by simp + +theorem conj_le_cong: "(0 <= x \ P = P') \ (0 <= (x::int) \ P) = (0 <= x \ P')" + by (simp cong: conj_cong) +lemma int_eq_number_of_eq: + "(((number_of v)::int) = (number_of w)) = iszero ((number_of (v + (uminus w)))::int)" + by simp + +lemma mod_eq0_dvd_iff[presburger]: "(m::nat) mod n = 0 \ n dvd m" +unfolding dvd_eq_mod_eq_0[symmetric] .. + +lemma zmod_eq0_zdvd_iff[presburger]: "(m::int) mod n = 0 \ n dvd m" +unfolding zdvd_iff_zmod_eq_0[symmetric] .. +declare mod_1[presburger] +declare mod_0[presburger] +declare zmod_1[presburger] +declare zmod_zero[presburger] +declare zmod_self[presburger] +declare mod_self[presburger] +declare DIVISION_BY_ZERO_MOD[presburger] +declare nat_mod_div_trivial[presburger] +declare div_mod_equality2[presburger] +declare div_mod_equality[presburger] +declare mod_div_equality2[presburger] +declare mod_div_equality[presburger] +declare mod_mult_self1[presburger] +declare mod_mult_self2[presburger] +declare zdiv_zmod_equality2[presburger] +declare zdiv_zmod_equality[presburger] +declare mod2_Suc_Suc[presburger] +lemma [presburger]: "(a::int) div 0 = 0" and [presburger]: "a mod 0 = a" +using IntDiv.DIVISION_BY_ZERO by blast+ + +use "Tools/Qelim/cooper.ML" +oracle linzqe_oracle ("term") = Coopereif.cooper_oracle + +use "Tools/Qelim/presburger.ML" + +setup {* + arith_tactic_add + (mk_arith_tactic "presburger" (fn i => fn st => + (warning "Trying Presburger arithmetic ..."; + Presburger.cooper_tac true [] [] ((ProofContext.init o theory_of_thm) st) i st))) + (* FIXME!!!!!!! get the right context!!*) +*} + +method_setup presburger = {* +let + fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () + fun simple_keyword k = Scan.lift (Args.$$$ k) >> K () + val addN = "add" + val delN = "del" + val elimN = "elim" + val any_keyword = keyword addN || keyword delN || simple_keyword elimN + val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; +in + fn src => Method.syntax + ((Scan.optional (simple_keyword elimN >> K false) true) -- + (Scan.optional (keyword addN |-- thms) []) -- + (Scan.optional (keyword delN |-- thms) [])) src + #> (fn (((elim, add_ths), del_ths),ctxt) => + Method.SIMPLE_METHOD' (Presburger.cooper_tac elim add_ths del_ths ctxt)) +end +*} "Cooper's algorithm for Presburger arithmetic" + +lemma [presburger]: "m mod 2 = (1::nat) \ \ 2 dvd m " by presburger +lemma [presburger]: "m mod 2 = Suc 0 \ \ 2 dvd m " by presburger +lemma [presburger]: "m mod (Suc (Suc 0)) = (1::nat) \ \ 2 dvd m " by presburger +lemma [presburger]: "m mod (Suc (Suc 0)) = Suc 0 \ \ 2 dvd m " by presburger +lemma [presburger]: "m mod 2 = (1::int) \ \ 2 dvd m " by presburger + + +subsection {* Code generator setup *} + +text {* + Presburger arithmetic is convenient to prove some + of the following code lemmas on integer numerals: +*} + +lemma eq_Pls_Pls: + "Numeral.Pls = Numeral.Pls \ True" by presburger + +lemma eq_Pls_Min: + "Numeral.Pls = Numeral.Min \ False" + unfolding Pls_def Numeral.Min_def by presburger + +lemma eq_Pls_Bit0: + "Numeral.Pls = Numeral.Bit k bit.B0 \ Numeral.Pls = k" + unfolding Pls_def Bit_def bit.cases by presburger + +lemma eq_Pls_Bit1: + "Numeral.Pls = Numeral.Bit k bit.B1 \ False" + unfolding Pls_def Bit_def bit.cases by presburger + +lemma eq_Min_Pls: + "Numeral.Min = Numeral.Pls \ False" + unfolding Pls_def Numeral.Min_def by presburger + +lemma eq_Min_Min: + "Numeral.Min = Numeral.Min \ True" by presburger + +lemma eq_Min_Bit0: + "Numeral.Min = Numeral.Bit k bit.B0 \ False" + unfolding Numeral.Min_def Bit_def bit.cases by presburger + +lemma eq_Min_Bit1: + "Numeral.Min = Numeral.Bit k bit.B1 \ Numeral.Min = k" + unfolding Numeral.Min_def Bit_def bit.cases by presburger + +lemma eq_Bit0_Pls: + "Numeral.Bit k bit.B0 = Numeral.Pls \ Numeral.Pls = k" + unfolding Pls_def Bit_def bit.cases by presburger + +lemma eq_Bit1_Pls: + "Numeral.Bit k bit.B1 = Numeral.Pls \ False" + unfolding Pls_def Bit_def bit.cases by presburger + +lemma eq_Bit0_Min: + "Numeral.Bit k bit.B0 = Numeral.Min \ False" + unfolding Numeral.Min_def Bit_def bit.cases by presburger + +lemma eq_Bit1_Min: + "(Numeral.Bit k bit.B1) = Numeral.Min \ Numeral.Min = k" + unfolding Numeral.Min_def Bit_def bit.cases by presburger + +lemma eq_Bit_Bit: + "Numeral.Bit k1 v1 = Numeral.Bit k2 v2 \ + v1 = v2 \ k1 = k2" + unfolding Bit_def + apply (cases v1) + apply (cases v2) + apply auto + apply presburger + apply (cases v2) + apply auto + apply presburger + apply (cases v2) + apply auto + done + +lemma eq_number_of: + "(number_of k \ int) = number_of l \ k = l" + unfolding number_of_is_id .. + + +lemma less_eq_Pls_Pls: + "Numeral.Pls \ Numeral.Pls \ True" by rule+ + +lemma less_eq_Pls_Min: + "Numeral.Pls \ Numeral.Min \ False" + unfolding Pls_def Numeral.Min_def by presburger + +lemma less_eq_Pls_Bit: + "Numeral.Pls \ Numeral.Bit k v \ Numeral.Pls \ k" + unfolding Pls_def Bit_def by (cases v) auto + +lemma less_eq_Min_Pls: + "Numeral.Min \ Numeral.Pls \ True" + unfolding Pls_def Numeral.Min_def by presburger + +lemma less_eq_Min_Min: + "Numeral.Min \ Numeral.Min \ True" by rule+ + +lemma less_eq_Min_Bit0: + "Numeral.Min \ Numeral.Bit k bit.B0 \ Numeral.Min < k" + unfolding Numeral.Min_def Bit_def by auto + +lemma less_eq_Min_Bit1: + "Numeral.Min \ Numeral.Bit k bit.B1 \ Numeral.Min \ k" + unfolding Numeral.Min_def Bit_def by auto + +lemma less_eq_Bit0_Pls: + "Numeral.Bit k bit.B0 \ Numeral.Pls \ k \ Numeral.Pls" + unfolding Pls_def Bit_def by simp + +lemma less_eq_Bit1_Pls: + "Numeral.Bit k bit.B1 \ Numeral.Pls \ k < Numeral.Pls" + unfolding Pls_def Bit_def by auto + +lemma less_eq_Bit_Min: + "Numeral.Bit k v \ Numeral.Min \ k \ Numeral.Min" + unfolding Numeral.Min_def Bit_def by (cases v) auto + +lemma less_eq_Bit0_Bit: + "Numeral.Bit k1 bit.B0 \ Numeral.Bit k2 v \ k1 \ k2" + unfolding Bit_def bit.cases by (cases v) auto + +lemma less_eq_Bit_Bit1: + "Numeral.Bit k1 v \ Numeral.Bit k2 bit.B1 \ k1 \ k2" + unfolding Bit_def bit.cases by (cases v) auto + +lemma less_eq_Bit1_Bit0: + "Numeral.Bit k1 bit.B1 \ Numeral.Bit k2 bit.B0 \ k1 < k2" + unfolding Bit_def by (auto split: bit.split) + +lemma less_eq_number_of: + "(number_of k \ int) \ number_of l \ k \ l" + unfolding number_of_is_id .. + + +lemma less_Pls_Pls: + "Numeral.Pls < Numeral.Pls \ False" by simp + +lemma less_Pls_Min: + "Numeral.Pls < Numeral.Min \ False" + unfolding Pls_def Numeral.Min_def by presburger + +lemma less_Pls_Bit0: + "Numeral.Pls < Numeral.Bit k bit.B0 \ Numeral.Pls < k" + unfolding Pls_def Bit_def by auto + +lemma less_Pls_Bit1: + "Numeral.Pls < Numeral.Bit k bit.B1 \ Numeral.Pls \ k" + unfolding Pls_def Bit_def by auto + +lemma less_Min_Pls: + "Numeral.Min < Numeral.Pls \ True" + unfolding Pls_def Numeral.Min_def by presburger + +lemma less_Min_Min: + "Numeral.Min < Numeral.Min \ False" by simp + +lemma less_Min_Bit: + "Numeral.Min < Numeral.Bit k v \ Numeral.Min < k" + unfolding Numeral.Min_def Bit_def by (auto split: bit.split) + +lemma less_Bit_Pls: + "Numeral.Bit k v < Numeral.Pls \ k < Numeral.Pls" + unfolding Pls_def Bit_def by (auto split: bit.split) + +lemma less_Bit0_Min: + "Numeral.Bit k bit.B0 < Numeral.Min \ k \ Numeral.Min" + unfolding Numeral.Min_def Bit_def by auto + +lemma less_Bit1_Min: + "Numeral.Bit k bit.B1 < Numeral.Min \ k < Numeral.Min" + unfolding Numeral.Min_def Bit_def by auto + +lemma less_Bit_Bit0: + "Numeral.Bit k1 v < Numeral.Bit k2 bit.B0 \ k1 < k2" + unfolding Bit_def by (auto split: bit.split) + +lemma less_Bit1_Bit: + "Numeral.Bit k1 bit.B1 < Numeral.Bit k2 v \ k1 < k2" + unfolding Bit_def by (auto split: bit.split) + +lemma less_Bit0_Bit1: + "Numeral.Bit k1 bit.B0 < Numeral.Bit k2 bit.B1 \ k1 \ k2" + unfolding Bit_def bit.cases by arith + +lemma less_number_of: + "(number_of k \ int) < number_of l \ k < l" + unfolding number_of_is_id .. + +lemmas pred_succ_numeral_code [code func] = + arith_simps(5-12) + +lemmas plus_numeral_code [code func] = + arith_simps(13-17) + arith_simps(26-27) + arith_extra_simps(1) [where 'a = int] + +lemmas minus_numeral_code [code func] = + arith_simps(18-21) + arith_extra_simps(2) [where 'a = int] + arith_extra_simps(5) [where 'a = int] + +lemmas times_numeral_code [code func] = + arith_simps(22-25) + arith_extra_simps(4) [where 'a = int] + +lemmas eq_numeral_code [code func] = + eq_Pls_Pls eq_Pls_Min eq_Pls_Bit0 eq_Pls_Bit1 + eq_Min_Pls eq_Min_Min eq_Min_Bit0 eq_Min_Bit1 + eq_Bit0_Pls eq_Bit1_Pls eq_Bit0_Min eq_Bit1_Min eq_Bit_Bit + eq_number_of + +lemmas less_eq_numeral_code [code func] = less_eq_Pls_Pls less_eq_Pls_Min less_eq_Pls_Bit + less_eq_Min_Pls less_eq_Min_Min less_eq_Min_Bit0 less_eq_Min_Bit1 + less_eq_Bit0_Pls less_eq_Bit1_Pls less_eq_Bit_Min less_eq_Bit0_Bit less_eq_Bit_Bit1 less_eq_Bit1_Bit0 + less_eq_number_of + +lemmas less_numeral_code [code func] = less_Pls_Pls less_Pls_Min less_Pls_Bit0 + less_Pls_Bit1 less_Min_Pls less_Min_Min less_Min_Bit less_Bit_Pls + less_Bit0_Min less_Bit1_Min less_Bit_Bit0 less_Bit1_Bit less_Bit0_Bit1 + less_number_of + +end