diff --git a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy --- a/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy +++ b/src/HOL/Decision_Procs/Commutative_Ring_Complete.thy @@ -1,570 +1,515 @@ (* Author: Bernhard Haeupler This theory is about of the relative completeness of the ring method. As long as the reified atomic polynomials of type pol are in normal form, the ring method is complete. *) section \Proof of the relative completeness of method comm-ring\ theory Commutative_Ring_Complete imports Commutative_Ring begin text \Formalization of normal form\ fun isnorm :: "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" subsection \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) + using isnorm.elims(2) by fastforce lemma norm_PX2: "isnorm (PX P i Q) \ isnorm Q" - apply (cases i) - apply auto - apply (cases P) - apply auto - subgoal for \ pol2 by (cases pol2) auto - done + using isnorm.elims(1) by auto -lemma norm_PX1: "isnorm (PX P i Q) \ isnorm P" - apply (cases i) - apply auto - apply (cases P) - apply auto - subgoal for \ pol2 by (cases pol2) auto - done +lemma norm_PX1: assumes "isnorm (PX P i Q)" shows "isnorm P" +proof (cases P) + case (Pc x1) + then show ?thesis + by auto +qed (use assms isnorm.elims(1) in 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 (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False) - apply (rename_tac pol a, case_tac pol, auto) - apply (case_tac y, auto) - done +lemma mkPinj_cn: + assumes "y \ 0" and "isnorm Q" + shows "isnorm (mkPinj y Q)" +proof (cases Q) + case Pc + with assms show ?thesis + by (simp add: mkPinj_def) +next + case Pinj + with assms show ?thesis + using isnorm.elims(2) isnorm.simps(5) mkPinj_def by fastforce +next + case PX + with assms show ?thesis + using isnorm.elims(1) mkPinj_def by auto +qed lemma norm_PXtrans: - assumes A: "isnorm (PX P x Q)" + assumes "isnorm (PX P x Q)" and "isnorm Q2" + shows "isnorm (PX P x Q2)" + using assms isnorm.elims(3) by fastforce + + +lemma norm_PXtrans2: + assumes "isnorm (PX P x Q)" and "isnorm Q2" - shows "isnorm (PX P x Q2)" + shows "isnorm (PX P (Suc (n + x)) Q2)" proof (cases P) case (PX p1 y p2) with assms show ?thesis - apply (cases x) - apply auto - apply (cases p2) - apply auto - done + using isnorm.elims(2) by fastforce next case Pc with assms show ?thesis by (cases x) auto next case Pinj with assms show ?thesis by (cases x) auto qed -lemma norm_PXtrans2: - assumes "isnorm (PX P x Q)" - and "isnorm Q2" - shows "isnorm (PX P (Suc (n + x)) Q2)" -proof (cases P) - case (PX p1 y p2) - with assms show ?thesis - apply (cases x) - apply auto - apply (cases p2) - apply auto - done -next - case Pc - with assms show ?thesis - by (cases x) auto -next - case Pinj - with assms show ?thesis - by (cases x) auto -qed - -text \\mkPX\ conserves normalizedness (\_cn\)\ +text \\mkPX\ preserves the normal property (\_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) with assms show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) next case (Pinj i Q) with assms show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) next case (PX P1 y P2) with assms have Y0: "y > 0" by (cases y) auto from assms PX have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) from assms PX Y0 show ?thesis - apply (cases x) - apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _]) - apply (cases P2) - apply auto - done + proof (cases P2) + case (Pc x1) + then show ?thesis + using assms gr0_conv_Suc PX by (auto simp add: mkPX_def norm_PXtrans2) + next + case (Pinj x21 x22) + with assms gr0_conv_Suc PX show ?thesis + by (auto simp: mkPX_def) + next + case (PX x31 x32 x33) + with assms gr0_conv_Suc \P = PX P1 y P2\ + show ?thesis + using assms PX by (auto simp add: mkPX_def norm_PXtrans2) + qed qed -text \\add\ conserves normalizedness\ +text \\add\ preserves the normal property\ lemma add_cn: "isnorm P \ isnorm Q \ isnorm (P \+\ Q)" proof (induct P Q rule: add.induct) case 1 then show ?case by simp next case (2 c i P2) then show ?case - apply (cases P2) - apply simp_all - apply (cases i) - apply simp_all - done + using isnorm.elims(2) by fastforce next case (3 i P2 c) then show ?case - apply (cases P2) - apply simp_all - apply (cases i) - apply simp_all - done + using isnorm.elims(2) by fastforce next case (4 c P2 i Q2) then have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) with 4 show ?case - apply (cases i) - apply simp - apply (cases P2) - apply auto - subgoal for \ pol2 by (cases pol2) auto - done + using isnorm.elims(2) by fastforce next case (5 P2 i Q2 c) then have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) with 5 show ?case - apply (cases i) - apply simp - apply (cases P2) - apply auto - subgoal for \ pol2 by (cases pol2) auto - done + using isnorm.elims(2) by fastforce next case (6 x P2 y Q2) then have Y0: "y > 0" by (cases y) (auto simp add: norm_Pinj_0_False) with 6 have X0: "x > 0" by (cases x) (auto simp add: norm_Pinj_0_False) consider "x < y" | "x = y" | "x > y" by arith then show ?case proof cases case xy: 1 then obtain d where y: "y = d + x" by atomize_elim arith from 6 have *: "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) from 6 xy y have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) with 6 X0 y * show ?thesis by (simp add: mkPinj_cn) next case xy: 2 from 6 have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) with xy 6 Y0 show ?thesis by (simp add: mkPinj_cn) next case xy: 3 then obtain d where x: "x = d + y" by atomize_elim arith from 6 have *: "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) from 6 xy x have "isnorm (Pinj d P2)" by (cases d) (simp, cases P2, auto) with xy 6 Y0 * x show ?thesis by (simp add: mkPinj_cn) qed next case (7 x P2 Q2 y R) consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" by atomize_elim arith then show ?case proof cases case 1 with 7 show ?thesis by (auto simp add: norm_Pinj_0_False) next case x: 2 from 7 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 7 x have "isnorm (R \+\ P2)" by simp with 7 x show ?thesis by (simp add: norm_PXtrans[of Q2 y _]) next case x: 3 with 7 have NR: "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 7 x have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto with 7 x NR have "isnorm (R \+\ Pinj (x - 1) P2)" by simp with \isnorm (PX Q2 y R)\ x show ?thesis by (simp add: norm_PXtrans[of Q2 y _]) qed next case (8 Q2 y R x P2) consider "x = 0" | "x = 1" | "x > 1" by arith then show ?case proof cases case 1 with 8 show ?thesis by (auto simp add: norm_Pinj_0_False) next case x: 2 with 8 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 8 x have "isnorm (R \+\ P2)" by simp with 8 x show ?thesis by (simp add: norm_PXtrans[of Q2 y _]) next case x: 3 then obtain d where x: "x = Suc (Suc d)" by atomize_elim arith with 8 have NR: "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 8 x have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto with 8 x NR have "isnorm (R \+\ Pinj (x - 1) P2)" by simp with \isnorm (PX Q2 y R)\ x show ?thesis by (simp add: norm_PXtrans[of Q2 y _]) qed next case (9 P1 x P2 Q1 y Q2) then have Y0: "y > 0" by (cases y) auto with 9 have X0: "x > 0" by (cases x) auto with 9 have NP1: "isnorm P1" and NP2: "isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) with 9 have NQ1: "isnorm Q1" and NQ2: "isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) consider "y < x" | "x = y" | "x < y" by arith then show ?case proof cases case xy: 1 then obtain d where x: "x = d + y" by atomize_elim arith have "isnorm (PX P1 d (Pc 0))" proof (cases P1) case (PX p1 y p2) with 9 xy x show ?thesis by (cases d) (simp, cases p2, auto) next case Pc with 9 xy x show ?thesis by (cases d) auto next case Pinj with 9 xy x show ?thesis by (cases d) auto qed with 9 NQ1 NP1 NP2 NQ2 xy x have "isnorm (P2 \+\ Q2)" "isnorm (PX P1 (x - y) (Pc 0) \+\ Q1)" by auto with Y0 xy x show ?thesis by (simp add: mkPX_cn) next case xy: 2 with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \+\ Q2)" "isnorm (P1 \+\ Q1)" by auto with xy Y0 show ?thesis by (simp add: mkPX_cn) next case xy: 3 then obtain d where y: "y = d + x" by atomize_elim arith have "isnorm (PX Q1 d (Pc 0))" proof (cases Q1) case (PX p1 y p2) with 9 xy y show ?thesis by (cases d) (simp, cases p2, auto) next case Pc with 9 xy y show ?thesis by (cases d) auto next case Pinj with 9 xy y show ?thesis by (cases d) auto qed with 9 NQ1 NP1 NP2 NQ2 xy y have "isnorm (P2 \+\ Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \+\ P1)" by auto with X0 xy y show ?thesis by (simp add: mkPX_cn) qed qed text \\mul\ concerves normalizedness\ lemma mul_cn: "isnorm P \ isnorm Q \ isnorm (P \*\ Q)" proof (induct P Q rule: mul.induct) case 1 then show ?case by simp next case (2 c i P2) then show ?case - apply (cases P2) - apply simp_all - apply (cases i) - apply (simp_all add: mkPinj_cn) - done + by (metis mkPinj_cn mul.simps(2) norm_Pinj norm_Pinj_0_False) next case (3 i P2 c) then show ?case - apply (cases P2) - apply simp_all - apply (cases i) - apply (simp_all add: mkPinj_cn) - done + by (metis mkPinj_cn mul.simps(3) norm_Pinj norm_Pinj_0_False) next case (4 c P2 i Q2) - then have "isnorm P2" "isnorm Q2" - by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with 4 show ?case - apply (cases "c = 0") - apply simp_all - apply (cases "i = 0") - apply (simp_all add: mkPX_cn) - done + then show ?case + by (metis isnorm.simps(6) mkPX_cn mul.simps(4) norm_PX1 norm_PX2) next case (5 P2 i Q2 c) - then have "isnorm P2" "isnorm Q2" - by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) - with 5 show ?case - apply (cases "c = 0") - apply simp_all - apply (cases "i = 0") - apply (simp_all add: mkPX_cn) - done + then show ?case + by (metis isnorm.simps(6) mkPX_cn mul.simps(5) norm_PX1 norm_PX2) next case (6 x P2 y Q2) consider "x < y" | "x = y" | "x > y" by arith then show ?case proof cases case xy: 1 then obtain d where y: "y = d + x" by atomize_elim arith from 6 have *: "x > 0" by (cases x) (auto simp add: norm_Pinj_0_False) from 6 have **: "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) from 6 xy y have "isnorm (Pinj d Q2)" - apply (cases d) - apply simp - apply (cases Q2) - apply auto - done + apply simp + by (smt (verit, ccfv_SIG) "**"(2) Suc_pred isnorm.elims(1) isnorm.simps(2) isnorm.simps(3) isnorm.simps(5)) with 6 * ** y show ?thesis by (simp add: mkPinj_cn) next case xy: 2 from 6 have *: "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) from 6 have **: "y > 0" by (cases y) (auto simp add: norm_Pinj_0_False) with 6 xy * ** show ?thesis by (simp add: mkPinj_cn) next case xy: 3 then obtain d where x: "x = d + y" by atomize_elim arith from 6 have *: "y > 0" by (cases y) (auto simp add: norm_Pinj_0_False) from 6 have **: "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) - from 6 xy x have "isnorm (Pinj d P2)" - apply (cases d) + with 6 xy x have "isnorm (Pinj d P2)" apply simp - apply (cases P2) - apply auto - done + by (smt (verit, ccfv_SIG) Suc_pred isnorm.elims(1) isnorm.simps(2) isnorm.simps(3) isnorm.simps(5)) with 6 xy * ** x show ?thesis by (simp add: mkPinj_cn) qed next case (7 x P2 Q2 y R) then have Y0: "y > 0" by (cases y) auto consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" by atomize_elim arith then show ?case proof cases case 1 with 7 show ?thesis by (auto simp add: norm_Pinj_0_False) next case x: 2 from 7 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 7 x have "isnorm (R \*\ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) with 7 x Y0 show ?thesis by (simp add: mkPX_cn) next case x: 3 from 7 have *: "isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) from 7 x have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto with 7 x * have "isnorm (R \*\ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \*\ Q2)" by auto with Y0 x show ?thesis by (simp add: mkPX_cn) qed next case (8 Q2 y R x P2) then have Y0: "y > 0" by (cases y) auto consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" by atomize_elim arith then show ?case proof cases case 1 with 8 show ?thesis by (auto simp add: norm_Pinj_0_False) next case x: 2 from 8 have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) with 8 x have "isnorm (R \*\ P2)" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) with 8 x Y0 show ?thesis by (simp add: mkPX_cn) next case x: 3 from 8 have *: "isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) from 8 x have "isnorm (Pinj (x - 1) P2)" by (cases P2) auto with 8 x * have "isnorm (R \*\ Pinj (x - 1) P2)" "isnorm (Pinj x P2 \*\ Q2)" by auto with Y0 x show ?thesis by (simp add: mkPX_cn) qed next case (9 P1 x P2 Q1 y Q2) from 9 have X0: "x > 0" by (cases x) auto from 9 have Y0: "y > 0" by (cases y) auto from 9 have *: "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) from 9 have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) with 9 * have "isnorm (P1 \*\ Q1)" "isnorm (P2 \*\ Q2)" "isnorm (P1 \*\ mkPinj 1 Q2)" "isnorm (Q1 \*\ mkPinj 1 P2)" by (auto simp add: mkPinj_cn) with 9 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) then show ?case by (simp add: add_cn) qed -text \neg conserves normalizedness\ +text \neg preserves the normal property\ lemma neg_cn: "isnorm P \ isnorm (neg P)" proof (induct P) case Pc then show ?case by simp next case (Pinj i P2) then have "isnorm P2" by (simp add: norm_Pinj[of i P2]) with Pinj show ?case by (cases P2) (auto, cases i, auto) next case (PX P1 x P2) note PX1 = this from PX have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) with PX show ?case proof (cases P1) case (PX p1 y p2) with PX1 show ?thesis by (cases x) (auto, cases p2, auto) next case Pinj with PX1 show ?thesis by (cases x) auto qed (cases x, auto) qed -text \sub conserves normalizedness\ +text \sub preserves the normal property\ 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 Pc then show ?case by simp next case (Pinj i Q) then show ?case - apply (cases Q) - apply (auto simp add: mkPX_cn mkPinj_cn) - apply (cases i) - apply (auto simp add: mkPX_cn mkPinj_cn) - done + by (metis Commutative_Ring.sqr.simps(2) mkPinj_cn norm_Pinj norm_Pinj_0_False) next case (PX P1 x P2) then have "x + x \ 0" "isnorm P2" "isnorm P1" by (cases x) (use PX in \auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]\) with PX 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) then show ?case by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) qed -text \\pow\ conserves normalizedness\ +text \\pow\ preserves the normal property\ lemma pow_cn: "isnorm P \ isnorm (pow n P)" proof (induct n arbitrary: P rule: less_induct) case (less k) show ?case proof (cases "k = 0") case True then show ?thesis by simp next case False then have K2: "k div 2 < k" by (cases k) auto from less have "isnorm (sqr P)" by (simp add: sqr_cn) with less False K2 show ?thesis by (cases "even k") (auto simp add: mul_cn elim: evenE oddE) qed qed end diff --git a/src/HOL/Decision_Procs/MIR.thy b/src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy +++ b/src/HOL/Decision_Procs/MIR.thy @@ -1,5721 +1,5709 @@ (* Title: HOL/Decision_Procs/MIR.thy Author: Amine Chaieb *) theory MIR imports Complex_Main Dense_Linear_Order DP_Library "HOL-Library.Code_Target_Numeral" "HOL-Library.Old_Recdef" begin section \Prelude\ abbreviation (input) UNION :: "'a set \ ('a \ 'b set) \ 'b set" where "UNION A f \ \ (f ` A)" \ \legacy\ section \Quantifier elimination for \\ (0, 1, +, floor, <)\\ declare of_int_floor_cancel [simp del] lemma myle: fixes a b :: "'a::{ordered_ab_group_add}" shows "(a \ b) = (0 \ b - a)" by (metis add_0_left add_le_cancel_right diff_add_cancel) lemma myless: fixes a b :: "'a::{ordered_ab_group_add}" shows "(a < b) = (0 < b - a)" by (metis le_iff_diff_le_0 less_le_not_le myle) (* Periodicity of dvd *) lemmas dvd_period = zdvd_period (* The Divisibility relation between reals *) definition rdvd:: "real \ real \ bool" (infixl "rdvd" 50) where "x rdvd y \ (\k::int. y = x * real_of_int k)" lemma int_rdvd_real: "real_of_int (i::int) rdvd x = (i dvd \x\ \ real_of_int \x\ = x)" (is "?l = ?r") proof assume "?l" hence th: "\ k. x=real_of_int (i*k)" by (simp add: rdvd_def) hence th': "real_of_int \x\ = x" by (auto simp del: of_int_mult) with th have "\ k. real_of_int \x\ = real_of_int (i*k)" by simp hence "\k. \x\ = i*k" by presburger thus ?r using th' by (simp add: dvd_def) next assume "?r" hence "(i::int) dvd \x::real\" .. hence "\k. real_of_int \x\ = real_of_int (i*k)" by (metis (no_types) dvd_def) thus ?l using \?r\ by (simp add: rdvd_def) qed lemma int_rdvd_iff: "(real_of_int (i::int) rdvd real_of_int t) = (i dvd t)" by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only: of_int_mult[symmetric]) lemma rdvd_abs1: "(\real_of_int d\ rdvd t) = (real_of_int (d ::int) rdvd t)" proof assume d: "real_of_int d rdvd t" from d int_rdvd_real have d2: "d dvd \t\" and ti: "real_of_int \t\ = t" by auto from iffD2[OF abs_dvd_iff] d2 have "\d\ dvd \t\" by blast with ti int_rdvd_real[symmetric] have "real_of_int \d\ rdvd t" by blast thus "\real_of_int d\ rdvd t" by simp next assume "\real_of_int d\ rdvd t" hence "real_of_int \d\ rdvd t" by simp with int_rdvd_real[where i="\d\" and x="t"] have d2: "\d\ dvd \t\" and ti: "real_of_int \t\ = t" by auto from iffD1[OF abs_dvd_iff] d2 have "d dvd \t\" by blast with ti int_rdvd_real[symmetric] show "real_of_int d rdvd t" by blast qed lemma rdvd_minus: "(real_of_int (d::int) rdvd t) = (real_of_int d rdvd -t)" - apply (auto simp add: rdvd_def) - apply (rule_tac x="-k" in exI, simp) - apply (rule_tac x="-k" in exI, simp) - done + by (metis equation_minus_iff mult_minus_right of_int_minus rdvd_def) lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)" by (auto simp add: rdvd_def) lemma rdvd_mult: assumes knz: "k\0" shows "(real_of_int (n::int) * real_of_int (k::int) rdvd x * real_of_int k) = (real_of_int n rdvd x)" using knz by (simp add: rdvd_def) (*********************************************************************************) (**** SHADOW SYNTAX AND SEMANTICS ****) (*********************************************************************************) datatype (plugins del: size) num = C int | Bound nat | CN nat int num | Neg num | Add num num | Sub num num | Mul int num | Floor num | CF int num num instantiation num :: size begin primrec size_num :: "num \ nat" where "size_num (C c) = 1" | "size_num (Bound n) = 1" | "size_num (Neg a) = 1 + size_num a" | "size_num (Add a b) = 1 + size_num a + size_num b" | "size_num (Sub a b) = 3 + size_num a + size_num b" | "size_num (CN n c a) = 4 + size_num a " | "size_num (CF c a b) = 4 + size_num a + size_num b" | "size_num (Mul c a) = 1 + size_num a" | "size_num (Floor a) = 1 + size_num a" instance .. end (* Semantics of numeral terms (num) *) primrec Inum :: "real list \ num \ real" where "Inum bs (C c) = (real_of_int c)" | "Inum bs (Bound n) = bs!n" | "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (Inum bs a)" | "Inum bs (Neg a) = -(Inum bs a)" | "Inum bs (Add a b) = Inum bs a + Inum bs b" | "Inum bs (Sub a b) = Inum bs a - Inum bs b" | "Inum bs (Mul c a) = (real_of_int c) * Inum bs a" | "Inum bs (Floor a) = real_of_int \Inum bs a\" | "Inum bs (CF c a b) = real_of_int c * real_of_int \Inum bs a\ + Inum bs b" definition "isint t bs \ real_of_int \Inum bs t\ = Inum bs t" lemma isint_iff: "isint n bs = (real_of_int \Inum bs n\ = Inum bs n)" by (simp add: isint_def) lemma isint_Floor: "isint (Floor n) bs" by (simp add: isint_iff) lemma isint_Mul: "isint e bs \ isint (Mul c e) bs" proof- let ?e = "Inum bs e" assume be: "isint e bs" hence efe:"real_of_int \?e\ = ?e" by (simp add: isint_iff) have "real_of_int \Inum bs (Mul c e)\ = real_of_int \real_of_int (c * \?e\)\" using efe by simp also have "\ = real_of_int (c* \?e\)" by (metis floor_of_int) also have "\ = real_of_int c * ?e" using efe by simp finally show ?thesis using isint_iff by simp qed lemma isint_neg: "isint e bs \ isint (Neg e) bs" proof- let ?I = "\ t. Inum bs t" assume ie: "isint e bs" hence th: "real_of_int \?I e\ = ?I e" by (simp add: isint_def) have "real_of_int \?I (Neg e)\ = real_of_int \- (real_of_int \?I e\)\" by (simp add: th) also have "\ = - real_of_int \?I e\" by simp finally show "isint (Neg e) bs" by (simp add: isint_def th) qed lemma isint_sub: assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" proof- let ?I = "\ t. Inum bs t" from ie have th: "real_of_int \?I e\ = ?I e" by (simp add: isint_def) have "real_of_int \?I (Sub (C c) e)\ = real_of_int \real_of_int (c - \?I e\)\" by (simp add: th) also have "\ = real_of_int (c - \?I e\)" by simp finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) qed lemma isint_add: assumes ai: "isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs" proof- let ?a = "Inum bs a" let ?b = "Inum bs b" from ai bi isint_iff have "real_of_int \?a + ?b\ = real_of_int \real_of_int \?a\ + real_of_int \?b\\" by simp also have "\ = real_of_int \?a\ + real_of_int \?b\" by simp also have "\ = ?a + ?b" using ai bi isint_iff by simp finally show "isint (Add a b) bs" by (simp add: isint_iff) qed lemma isint_c: "isint (C j) bs" by (simp add: isint_iff) (* FORMULAE *) datatype (plugins del: size) fm = T | F | Lt num | Le num | Gt num | Ge num | Eq num | NEq num | Dvd int num | NDvd int num | Not fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm instantiation fm :: size begin primrec size_fm :: "fm \ nat" where "size_fm (Not p) = 1 + size_fm p" | "size_fm (And p q) = 1 + size_fm p + size_fm q" | "size_fm (Or p q) = 1 + size_fm p + size_fm q" | "size_fm (Imp p q) = 3 + size_fm p + size_fm q" | "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)" | "size_fm (E p) = 1 + size_fm p" | "size_fm (A p) = 4 + size_fm p" | "size_fm (Dvd i t) = 2" | "size_fm (NDvd i t) = 2" | "size_fm T = 1" | "size_fm F = 1" | "size_fm (Lt _) = 1" | "size_fm (Le _) = 1" | "size_fm (Gt _) = 1" | "size_fm (Ge _) = 1" | "size_fm (Eq _) = 1" | "size_fm (NEq _) = 1" instance .. end lemma size_fm_pos [simp]: "size p > 0" for p :: fm by (induct p) simp_all (* Semantics of formulae (fm) *) primrec Ifm ::"real list \ fm \ bool" where "Ifm bs T \ True" | "Ifm bs F \ False" | "Ifm bs (Lt a) \ Inum bs a < 0" | "Ifm bs (Gt a) \ Inum bs a > 0" | "Ifm bs (Le a) \ Inum bs a \ 0" | "Ifm bs (Ge a) \ Inum bs a \ 0" | "Ifm bs (Eq a) \ Inum bs a = 0" | "Ifm bs (NEq a) \ Inum bs a \ 0" | "Ifm bs (Dvd i b) \ real_of_int i rdvd Inum bs b" | "Ifm bs (NDvd i b) \ \ (real_of_int i rdvd Inum bs b)" | "Ifm bs (Not p) \ \ (Ifm bs p)" | "Ifm bs (And p q) \ Ifm bs p \ Ifm bs q" | "Ifm bs (Or p q) \ Ifm bs p \ Ifm bs q" | "Ifm bs (Imp p q) \ (Ifm bs p \ Ifm bs q)" | "Ifm bs (Iff p q) \ (Ifm bs p \ Ifm bs q)" | "Ifm bs (E p) \ (\x. Ifm (x # bs) p)" | "Ifm bs (A p) \ (\x. Ifm (x # bs) p)" fun prep :: "fm \ fm" where "prep (E T) = T" | "prep (E F) = F" | "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" | "prep (E (Imp p q)) = Or (prep (E (Not p))) (prep (E q))" | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (Not p) (Not q))))" | "prep (E (Not (And p q))) = Or (prep (E (Not p))) (prep (E(Not q)))" | "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))" | "prep (E (Not (Iff p q))) = Or (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))" | "prep (E p) = E (prep p)" | "prep (A (And p q)) = And (prep (A p)) (prep (A q))" | "prep (A p) = prep (Not (E (Not p)))" | "prep (Not (Not p)) = prep p" | "prep (Not (And p q)) = Or (prep (Not p)) (prep (Not q))" | "prep (Not (A p)) = prep (E (Not p))" | "prep (Not (Or p q)) = And (prep (Not p)) (prep (Not q))" | "prep (Not (Imp p q)) = And (prep p) (prep (Not q))" | "prep (Not (Iff p q)) = Or (prep (And p (Not q))) (prep (And (Not p) q))" | "prep (Not p) = Not (prep p)" | "prep (Or p q) = Or (prep p) (prep q)" | "prep (And p q) = And (prep p) (prep q)" | "prep (Imp p q) = prep (Or (Not p) q)" | "prep (Iff p q) = Or (prep (And p q)) (prep (And (Not p) (Not q)))" | "prep p = p" lemma prep: "\ bs. Ifm bs (prep p) = Ifm bs p" by (induct p rule: prep.induct) auto (* Quantifier freeness *) fun qfree:: "fm \ bool" where "qfree (E p) = False" | "qfree (A p) = False" | "qfree (Not p) = qfree p" | "qfree (And p q) = (qfree p \ qfree q)" | "qfree (Or p q) = (qfree p \ qfree q)" | "qfree (Imp p q) = (qfree p \ qfree q)" | "qfree (Iff p q) = (qfree p \ qfree q)" | "qfree p = True" (* Boundedness and substitution *) primrec numbound0 :: "num \ bool" (* a num is INDEPENDENT of Bound 0 *) where "numbound0 (C c) = True" | "numbound0 (Bound n) = (n>0)" | "numbound0 (CN n i a) = (n > 0 \ numbound0 a)" | "numbound0 (Neg a) = numbound0 a" | "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" | "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" | "numbound0 (Mul i a) = numbound0 a" | "numbound0 (Floor a) = numbound0 a" | "numbound0 (CF c a b) = (numbound0 a \ numbound0 b)" lemma numbound0_I: assumes nb: "numbound0 a" shows "Inum (b#bs) a = Inum (b'#bs) a" using nb by (induct a) auto lemma numbound0_gen: assumes nb: "numbound0 t" and ti: "isint t (x#bs)" shows "\ y. isint t (y#bs)" using nb ti proof(clarify) fix y from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def] show "isint t (y#bs)" by (simp add: isint_def) qed primrec bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) where "bound0 T = True" | "bound0 F = True" | "bound0 (Lt a) = numbound0 a" | "bound0 (Le a) = numbound0 a" | "bound0 (Gt a) = numbound0 a" | "bound0 (Ge a) = numbound0 a" | "bound0 (Eq a) = numbound0 a" | "bound0 (NEq a) = numbound0 a" | "bound0 (Dvd i a) = numbound0 a" | "bound0 (NDvd i a) = numbound0 a" | "bound0 (Not p) = bound0 p" | "bound0 (And p q) = (bound0 p \ bound0 q)" | "bound0 (Or p q) = (bound0 p \ bound0 q)" | "bound0 (Imp p q) = ((bound0 p) \ (bound0 q))" | "bound0 (Iff p q) = (bound0 p \ bound0 q)" | "bound0 (E p) = False" | "bound0 (A p) = False" lemma bound0_I: assumes bp: "bound0 p" shows "Ifm (b#bs) p = Ifm (b'#bs) p" using bp numbound0_I [where b="b" and bs="bs" and b'="b'"] by (induct p) auto primrec numsubst0:: "num \ num \ num" (* substitute a num into a num for Bound 0 *) where "numsubst0 t (C c) = (C c)" | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))" | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)" | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" | "numsubst0 t (Floor a) = Floor (numsubst0 t a)" lemma numsubst0_I: shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" by (induct t) simp_all primrec subst0:: "num \ fm \ fm" (* substitue a num into a formula for Bound 0 *) where "subst0 t T = T" | "subst0 t F = F" | "subst0 t (Lt a) = Lt (numsubst0 t a)" | "subst0 t (Le a) = Le (numsubst0 t a)" | "subst0 t (Gt a) = Gt (numsubst0 t a)" | "subst0 t (Ge a) = Ge (numsubst0 t a)" | "subst0 t (Eq a) = Eq (numsubst0 t a)" | "subst0 t (NEq a) = NEq (numsubst0 t a)" | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" | "subst0 t (Not p) = Not (subst0 t p)" | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" lemma subst0_I: assumes qfp: "qfree p" shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p" using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] by (induct p) simp_all fun decrnum:: "num \ num" where "decrnum (Bound n) = Bound (n - 1)" | "decrnum (Neg a) = Neg (decrnum a)" | "decrnum (Add a b) = Add (decrnum a) (decrnum b)" | "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" | "decrnum (Mul c a) = Mul c (decrnum a)" | "decrnum (Floor a) = Floor (decrnum a)" | "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" | "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)" | "decrnum a = a" fun decr :: "fm \ fm" where "decr (Lt a) = Lt (decrnum a)" | "decr (Le a) = Le (decrnum a)" | "decr (Gt a) = Gt (decrnum a)" | "decr (Ge a) = Ge (decrnum a)" | "decr (Eq a) = Eq (decrnum a)" | "decr (NEq a) = NEq (decrnum a)" | "decr (Dvd i a) = Dvd i (decrnum a)" | "decr (NDvd i a) = NDvd i (decrnum a)" | "decr (Not p) = Not (decr p)" | "decr (And p q) = And (decr p) (decr q)" | "decr (Or p q) = Or (decr p) (decr q)" | "decr (Imp p q) = Imp (decr p) (decr q)" | "decr (Iff p q) = Iff (decr p) (decr q)" | "decr p = p" lemma decrnum: assumes nb: "numbound0 t" shows "Inum (x#bs) t = Inum bs (decrnum t)" using nb by (induct t rule: decrnum.induct) simp_all lemma decr: assumes nb: "bound0 p" shows "Ifm (x#bs) p = Ifm bs (decr p)" using nb by (induct p rule: decr.induct) (simp_all add: decrnum) lemma decr_qf: "bound0 p \ qfree (decr p)" by (induct p) simp_all fun isatom :: "fm \ bool" (* test for atomicity *) where "isatom T = True" | "isatom F = True" | "isatom (Lt a) = True" | "isatom (Le a) = True" | "isatom (Gt a) = True" | "isatom (Ge a) = True" | "isatom (Eq a) = True" | "isatom (NEq a) = True" | "isatom (Dvd i b) = True" | "isatom (NDvd i b) = True" | "isatom p = False" lemma numsubst0_numbound0: assumes nb: "numbound0 t" shows "numbound0 (numsubst0 t a)" using nb by (induct a) auto lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" shows "bound0 (subst0 t p)" using qf numsubst0_numbound0[OF nb] by (induct p) auto lemma bound0_qf: "bound0 p \ qfree p" by (induct p) simp_all definition djf:: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q = (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or fp q))" definition evaldjf:: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps = foldr (djf f) ps F" lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) (cases "f p", simp_all add: Let_def djf_def) lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" by (induct ps) (simp_all add: evaldjf_def djf_Or) lemma evaldjf_bound0: - assumes nb: "\ x\ set xs. bound0 (f x)" + assumes "\x \ set xs. bound0 (f x)" shows "bound0 (evaldjf f xs)" - using nb - apply (induct xs) - apply (auto simp add: evaldjf_def djf_def Let_def) - apply (case_tac "f a") - apply auto - done + using assms + by (induct xs) (auto simp add: evaldjf_def djf_def Let_def split: fm.split) + lemma evaldjf_qf: - assumes nb: "\ x\ set xs. qfree (f x)" + assumes "\x \ set xs. qfree (f x)" shows "qfree (evaldjf f xs)" - using nb - apply (induct xs) - apply (auto simp add: evaldjf_def djf_def Let_def) - apply (case_tac "f a") - apply auto - done + using assms + by (induct xs) (auto simp add: evaldjf_def djf_def Let_def split: fm.split) + fun disjuncts :: "fm \ fm list" where "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" | "disjuncts F = []" | "disjuncts p = [p]" fun conjuncts :: "fm \ fm list" where "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)" | "conjuncts T = []" | "conjuncts p = [p]" lemma conjuncts: "(\ q\ set (conjuncts p). Ifm bs q) = Ifm bs p" by (induct p rule: conjuncts.induct) auto lemma disjuncts_qf: "qfree p \ \ q\ set (disjuncts p). qfree q" proof - assume qf: "qfree p" hence "list_all qfree (disjuncts p)" by (induct p rule: disjuncts.induct, auto) thus ?thesis by (simp only: list_all_iff) qed lemma conjuncts_qf: "qfree p \ \ q\ set (conjuncts p). qfree q" proof- assume qf: "qfree p" hence "list_all qfree (conjuncts p)" by (induct p rule: conjuncts.induct, auto) thus ?thesis by (simp only: list_all_iff) qed definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. f (Or p q) = Or (f p) (f q)" and fF: "f F = F" shows "Ifm bs (DJ f p) = Ifm bs (f p)" proof - have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" by (simp add: DJ_def evaldjf_ex) also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) finally show ?thesis . qed lemma DJ_qf: assumes fqf: "\ p. qfree p \ qfree (f p)" shows "\p. qfree p \ qfree (DJ f p) " proof(clarify) fix p assume qf: "qfree p" have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp qed lemma DJ_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" proof(clarify) fix p::fm and bs assume qf: "qfree p" from qe have qth: "\ p. qfree p \ qfree (qe p)" by blast from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto have "Ifm bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm bs (qe q))" by (simp add: DJ_def evaldjf_ex) also have "\ = (\ q \ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto also have "\ = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) finally show "qfree (DJ qe p) \ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast qed (* Simplification *) (* Algebraic simplifications for nums *) fun bnds:: "num \ nat list" where "bnds (Bound n) = [n]" | "bnds (CN n c a) = n#(bnds a)" | "bnds (Neg a) = bnds a" | "bnds (Add a b) = (bnds a)@(bnds b)" | "bnds (Sub a b) = (bnds a)@(bnds b)" | "bnds (Mul i a) = bnds a" | "bnds (Floor a) = bnds a" | "bnds (CF c a b) = (bnds a)@(bnds b)" | "bnds a = []" fun lex_ns:: "nat list \ nat list \ bool" where "lex_ns [] ms = True" | "lex_ns ns [] = False" | "lex_ns (n#ns) (m#ms) = (n ((n = m) \ lex_ns ns ms)) " definition lex_bnd :: "num \ num \ bool" where "lex_bnd t s \ lex_ns (bnds t) (bnds s)" fun maxcoeff:: "num \ int" where "maxcoeff (C i) = \i\" | "maxcoeff (CN n c t) = max \c\ (maxcoeff t)" | "maxcoeff (CF c t s) = max \c\ (maxcoeff s)" | "maxcoeff t = 1" lemma maxcoeff_pos: "maxcoeff t \ 0" by (induct t rule: maxcoeff.induct) auto fun numgcdh:: "num \ int \ int" where "numgcdh (C i) = (\g. gcd i g)" | "numgcdh (CN n c t) = (\g. gcd c (numgcdh t g))" | "numgcdh (CF c s t) = (\g. gcd c (numgcdh t g))" | "numgcdh t = (\g. 1)" definition numgcd :: "num \ int" where "numgcd t = numgcdh t (maxcoeff t)" fun reducecoeffh:: "num \ int \ num" where "reducecoeffh (C i) = (\ g. C (i div g))" | "reducecoeffh (CN n c t) = (\ g. CN n (c div g) (reducecoeffh t g))" | "reducecoeffh (CF c s t) = (\ g. CF (c div g) s (reducecoeffh t g))" | "reducecoeffh t = (\g. t)" definition reducecoeff :: "num \ num" where "reducecoeff t = (let g = numgcd t in if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" fun dvdnumcoeff:: "num \ int \ bool" where "dvdnumcoeff (C i) = (\ g. g dvd i)" | "dvdnumcoeff (CN n c t) = (\ g. g dvd c \ (dvdnumcoeff t g))" | "dvdnumcoeff (CF c s t) = (\ g. g dvd c \ (dvdnumcoeff t g))" | "dvdnumcoeff t = (\g. False)" lemma dvdnumcoeff_trans: assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" shows "dvdnumcoeff t g" using dgt' gdg by (induct t rule: dvdnumcoeff.induct) (simp_all add: gdg dvd_trans[OF gdg]) declare dvd_trans [trans add] lemma numgcd0: assumes g0: "numgcd t = 0" shows "Inum bs t = 0" proof- have "\x. numgcdh t x= 0 \ Inum bs t = 0" by (induct t rule: numgcdh.induct, auto) thus ?thesis using g0[simplified numgcd_def] by blast qed lemma numgcdh_pos: assumes gp: "g \ 0" shows "numgcdh t g \ 0" using gp by (induct t rule: numgcdh.induct) auto lemma numgcd_pos: "numgcd t \0" by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) lemma reducecoeffh: assumes gt: "dvdnumcoeff t g" and gp: "g > 0" shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs t" using gt proof(induct t rule: reducecoeffh.induct) case (1 i) hence gd: "g dvd i" by simp from assms 1 show ?case by (simp add: real_of_int_div[OF gd]) next case (2 n c t) hence gd: "g dvd c" by simp from assms 2 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) next case (3 c s t) hence gd: "g dvd c" by simp from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) qed (auto simp add: numgcd_def gp) fun ismaxcoeff:: "num \ int \ bool" where "ismaxcoeff (C i) = (\ x. \i\ \ x)" | "ismaxcoeff (CN n c t) = (\x. \c\ \ x \ (ismaxcoeff t x))" | "ismaxcoeff (CF c s t) = (\x. \c\ \ x \ (ismaxcoeff t x))" | "ismaxcoeff t = (\x. True)" lemma ismaxcoeff_mono: "ismaxcoeff t c \ c \ c' \ ismaxcoeff t c'" by (induct t rule: ismaxcoeff.induct) auto lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" proof (induct t rule: maxcoeff.induct) case (2 n c t) hence H:"ismaxcoeff t (maxcoeff t)" . have thh: "maxcoeff t \ max \c\ (maxcoeff t)" by simp from ismaxcoeff_mono[OF H thh] show ?case by simp next case (3 c t s) hence H1:"ismaxcoeff s (maxcoeff s)" by auto have thh1: "maxcoeff s \ max \c\ (maxcoeff s)" by (simp add: max_def) from ismaxcoeff_mono[OF H1 thh1] show ?case by simp qed simp_all lemma zgcd_gt1: "\i\ > 1 \ \j\ > 1 \ \i\ = 0 \ \j\ > 1 \ \i\ > 1 \ \j\ = 0" if "gcd i j > 1" for i j :: int proof - have "\k\ \ 1 \ k = - 1 \ k = 0 \ k = 1" for k :: int by auto with that show ?thesis by (auto simp add: not_less) qed lemma numgcdh0:"numgcdh t m = 0 \ m =0" by (induct t rule: numgcdh.induct) auto lemma dvdnumcoeff_aux: assumes "ismaxcoeff t m" and mp:"m \ 0" and "numgcdh t m > 1" shows "dvdnumcoeff t (numgcdh t m)" using assms proof(induct t rule: numgcdh.induct) case (2 n c t) let ?g = "numgcdh t m" from 2 have th:"gcd c ?g > 1" by simp from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] have "(\c\ > 1 \ ?g > 1) \ (\c\ = 0 \ ?g > 1) \ (\c\ > 1 \ ?g = 0)" by simp moreover {assume "\c\ > 1" and gp: "?g > 1" with 2 have th: "dvdnumcoeff t ?g" by simp have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp } moreover {assume "\c\ = 0 \ ?g > 1" with 2 have th: "dvdnumcoeff t ?g" by simp have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp hence ?case by simp } moreover {assume "\c\ > 1" and g0:"?g = 0" from numgcdh0[OF g0] have "m=0". with 2 g0 have ?case by simp } ultimately show ?case by blast next case (3 c s t) let ?g = "numgcdh t m" from 3 have th:"gcd c ?g > 1" by simp from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] have "(\c\ > 1 \ ?g > 1) \ (\c\ = 0 \ ?g > 1) \ (\c\ > 1 \ ?g = 0)" by simp moreover {assume "\c\ > 1" and gp: "?g > 1" with 3 have th: "dvdnumcoeff t ?g" by simp have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp } moreover {assume "\c\ = 0 \ ?g > 1" with 3 have th: "dvdnumcoeff t ?g" by simp have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp hence ?case by simp } moreover {assume "\c\ > 1" and g0:"?g = 0" from numgcdh0[OF g0] have "m=0". with 3 g0 have ?case by simp } ultimately show ?case by blast qed auto lemma dvdnumcoeff_aux2: assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" using assms proof (simp add: numgcd_def) let ?mc = "maxcoeff t" let ?g = "numgcdh t ?mc" have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) have th2: "?mc \ 0" by (rule maxcoeff_pos) assume H: "numgcdh t ?mc > 1" from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . qed lemma reducecoeff: "real_of_int (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" proof- let ?g = "numgcd t" have "?g \ 0" by (simp add: numgcd_pos) hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} moreover { assume g1:"?g > 1" from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis by (simp add: reducecoeff_def Let_def)} ultimately show ?thesis by blast qed lemma reducecoeffh_numbound0: "numbound0 t \ numbound0 (reducecoeffh t g)" by (induct t rule: reducecoeffh.induct) auto lemma reducecoeff_numbound0: "numbound0 t \ numbound0 (reducecoeff t)" using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) consts numadd:: "num \ num \ num" recdef numadd "measure (\(t, s). size t + size s)" "numadd (CN n1 c1 r1,CN n2 c2 r2) = (if n1=n2 then (let c = c1 + c2 in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) else if n1 \ n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2)) else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" "numadd (CF c1 t1 r1,CF c2 t2 r2) = (if t1 = t2 then (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s)) else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2)) else CF c2 t2 (numadd(CF c1 t1 r1,r2)))" "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))" "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))" "numadd (C b1, C b2) = C (b1+b2)" "numadd (a,b) = Add a b" lemma numadd [simp]: "Inum bs (numadd (t, s)) = Inum bs (Add t s)" by (induct t s rule: numadd.induct) (simp_all add: Let_def algebra_simps add_eq_0_iff) lemma numadd_nb [simp]: "numbound0 t \ numbound0 s \ numbound0 (numadd (t, s))" by (induct t s rule: numadd.induct) (simp_all add: Let_def) fun nummul:: "num \ int \ num" where "nummul (C j) = (\ i. C (i*j))" | "nummul (CN n c t) = (\ i. CN n (c*i) (nummul t i))" | "nummul (CF c t s) = (\ i. CF (c*i) t (nummul s i))" | "nummul (Mul c t) = (\ i. nummul t (i*c))" | "nummul t = (\ i. Mul i t)" lemma nummul[simp]: "\ i. Inum bs (nummul t i) = Inum bs (Mul i t)" by (induct t rule: nummul.induct) (auto simp add: algebra_simps) lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" by (induct t rule: nummul.induct) auto definition numneg :: "num \ num" where "numneg t \ nummul t (- 1)" definition numsub :: "num \ num \ num" where "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" using numneg_def nummul by simp lemma numneg_nb[simp]: "numbound0 t \ numbound0 (numneg t)" using numneg_def by simp lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" using numsub_def by simp lemma numsub_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numsub t s)" using numsub_def by simp lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs" proof- have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor) have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def) also have "\" by (simp add: isint_add cti si) finally show ?thesis . qed fun split_int:: "num \ num \ num" where "split_int (C c) = (C 0, C c)" | "split_int (CN n c b) = (let (bv,bi) = split_int b in (CN n c bv, bi))" | "split_int (CF c a b) = (let (bv,bi) = split_int b in (bv, CF c a bi))" | "split_int a = (a,C 0)" lemma split_int: "\tv ti. split_int t = (tv,ti) \ (Inum bs (Add tv ti) = Inum bs t) \ isint ti bs" proof (induct t rule: split_int.induct) case (2 c n b tv ti) let ?bv = "fst (split_int b)" let ?bi = "snd (split_int b)" have "split_int b = (?bv,?bi)" by simp with 2(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ from 2(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def) from 2(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def) next case (3 c a b tv ti) let ?bv = "fst (split_int b)" let ?bi = "snd (split_int b)" have "split_int b = (?bv,?bi)" by simp with 3(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ from 3(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def) from 3(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF) qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def algebra_simps) lemma split_int_nb: "numbound0 t \ numbound0 (fst (split_int t)) \ numbound0 (snd (split_int t)) " by (induct t rule: split_int.induct) (auto simp add: Let_def split_def) definition numfloor:: "num \ num" where "numfloor t = (let (tv,ti) = split_int t in (case tv of C i \ numadd (tv,ti) | _ \ numadd(CF 1 tv (C 0),ti)))" lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)") proof- let ?tv = "fst (split_int t)" let ?ti = "snd (split_int t)" have tvti:"split_int t = (?tv,?ti)" by simp {assume H: "\ v. ?tv \ C v" hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" by (cases ?tv) (auto simp add: numfloor_def Let_def split_def) from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ hence "?N (Floor t) = real_of_int \?N (Add ?tv ?ti)\" by simp also have "\ = real_of_int (\?N ?tv\ + \?N ?ti\)" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) finally have ?thesis using th1 by simp} moreover {fix v assume H:"?tv = C v" from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ hence "?N (Floor t) = real_of_int \?N (Add ?tv ?ti)\" by simp also have "\ = real_of_int (\?N ?tv\ + \?N ?ti\)" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) finally have ?thesis by (simp add: H numfloor_def Let_def split_def) } ultimately show ?thesis by auto qed lemma numfloor_nb[simp]: "numbound0 t \ numbound0 (numfloor t)" using split_int_nb[where t="t"] by (cases "fst (split_int t)") (auto simp add: numfloor_def Let_def split_def) fun simpnum:: "num \ num" where "simpnum (C j) = C j" | "simpnum (Bound n) = CN n 1 (C 0)" | "simpnum (Neg t) = numneg (simpnum t)" | "simpnum (Add t s) = numadd (simpnum t,simpnum s)" | "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" | "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" | "simpnum (Floor t) = numfloor (simpnum t)" | "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))" | "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)" lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" by (induct t rule: simpnum.induct) auto lemma simpnum_numbound0[simp]: "numbound0 t \ numbound0 (simpnum t)" by (induct t rule: simpnum.induct) auto fun nozerocoeff:: "num \ bool" where "nozerocoeff (C c) = True" | "nozerocoeff (CN n c t) = (c\0 \ nozerocoeff t)" | "nozerocoeff (CF c s t) = (c \ 0 \ nozerocoeff t)" | "nozerocoeff (Mul c t) = (c\0 \ nozerocoeff t)" | "nozerocoeff t = True" lemma numadd_nz : "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numadd (a,b))" by (induct a b rule: numadd.induct) (auto simp add: Let_def) lemma nummul_nz : "\ i. i\0 \ nozerocoeff a \ nozerocoeff (nummul a i)" by (induct a rule: nummul.induct) (auto simp add: Let_def numadd_nz) lemma numneg_nz : "nozerocoeff a \ nozerocoeff (numneg a)" by (simp add: numneg_def nummul_nz) lemma numsub_nz: "nozerocoeff a \ nozerocoeff b \ nozerocoeff (numsub a b)" by (simp add: numsub_def numneg_nz numadd_nz) lemma split_int_nz: "nozerocoeff t \ nozerocoeff (fst (split_int t)) \ nozerocoeff (snd (split_int t))" by (induct t rule: split_int.induct) (auto simp add: Let_def split_def) lemma numfloor_nz: "nozerocoeff t \ nozerocoeff (numfloor t)" by (simp add: numfloor_def Let_def split_def) (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz) lemma simpnum_nz: "nozerocoeff (simpnum t)" by (induct t rule: simpnum.induct) (auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz) lemma maxcoeff_nz: "nozerocoeff t \ maxcoeff t = 0 \ t = C 0" proof (induct t rule: maxcoeff.induct) case (2 n c t) hence cnz: "c \0" and mx: "max \c\ (maxcoeff t) = 0" by simp+ have "max \c\ (maxcoeff t) \ \c\" by simp with cnz have "max \c\ (maxcoeff t) > 0" by arith with 2 show ?case by simp next case (3 c s t) hence cnz: "c \0" and mx: "max \c\ (maxcoeff t) = 0" by simp+ have "max \c\ (maxcoeff t) \ \c\" by simp with cnz have "max \c\ (maxcoeff t) > 0" by arith with 3 show ?case by simp qed auto lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" proof- from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) from numgcdh0[OF th] have th:"maxcoeff t = 0" . from maxcoeff_nz[OF nz th] show ?thesis . qed definition simp_num_pair :: "(num \ int) \ num \ int" where "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else (let t' = simpnum t ; g = numgcd t' in if g > 1 then (let g' = gcd n g in if g' = 1 then (t',n) else (reducecoeffh t' g', n div g')) else (t',n))))" lemma simp_num_pair_ci: shows "((\ (t,n). Inum bs t / real_of_int n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real_of_int n) (t,n))" (is "?lhs = ?rhs") proof- let ?t' = "simpnum t" let ?g = "numgcd ?t'" let ?g' = "gcd n ?g" {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} moreover { assume nnz: "n \ 0" {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} moreover {assume g1:"?g>1" hence g0: "?g > 0" by simp from g1 nnz have gp0: "?g' \ 0" by simp hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith hence "?g'= 1 \ ?g' > 1" by arith moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} moreover {assume g'1:"?g'>1" from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. let ?tt = "reducecoeffh ?t' ?g'" let ?t = "Inum bs ?tt" have gpdg: "?g' dvd ?g" by simp have gpdd: "?g' dvd n" by simp have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] have th2:"real_of_int ?g' * ?t = Inum bs ?t'" by simp from nnz g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')" by (simp add: simp_num_pair_def Let_def) also have "\ = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))" by simp also have "\ = (Inum bs ?t' / real_of_int n)" using real_of_int_div[OF gpdd] th2 gp0 by simp finally have "?lhs = Inum bs t / real_of_int n" by simp then have ?thesis using nnz g1 g'1 by (simp add: simp_num_pair_def) } ultimately have ?thesis by auto } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" shows "numbound0 t' \ n' >0" proof- let ?t' = "simpnum t" let ?g = "numgcd ?t'" let ?g' = "gcd n ?g" { assume nz: "n = 0" hence ?thesis using assms by (simp add: Let_def simp_num_pair_def) } moreover { assume nnz: "n \ 0" {assume "\ ?g > 1" hence ?thesis using assms by (auto simp add: Let_def simp_num_pair_def) } moreover {assume g1:"?g>1" hence g0: "?g > 0" by simp from g1 nnz have gp0: "?g' \ 0" by simp hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith hence "?g'= 1 \ ?g' > 1" by arith moreover {assume "?g'=1" hence ?thesis using assms g1 g0 by (auto simp add: Let_def simp_num_pair_def) } moreover {assume g'1:"?g'>1" have gpdg: "?g' dvd ?g" by simp have gpdd: "?g' dvd n" by simp have gpdgp: "?g' dvd ?g'" by simp from zdvd_imp_le[OF gpdd np] have g'n: "?g' \ n" . from zdiv_mono1[OF g'n g'p, simplified div_self[OF gp0]] have "n div ?g' >0" by simp hence ?thesis using assms g1 g'1 by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)} ultimately have ?thesis by auto } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed fun not:: "fm \ fm" where "not (Not p) = p" | "not T = F" | "not F = T" | "not (Lt t) = Ge t" | "not (Le t) = Gt t" | "not (Gt t) = Le t" | "not (Ge t) = Lt t" | "not (Eq t) = NEq t" | "not (NEq t) = Eq t" | "not (Dvd i t) = NDvd i t" | "not (NDvd i t) = Dvd i t" | "not (And p q) = Or (not p) (not q)" | "not (Or p q) = And (not p) (not q)" | "not p = Not p" lemma not[simp]: "Ifm bs (not p) = Ifm bs (Not p)" by (induct p) auto lemma not_qf[simp]: "qfree p \ qfree (not p)" by (induct p) auto lemma not_nb[simp]: "bound0 p \ bound0 (not p)" by (induct p) auto definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" by (cases "p=F \ q=F", simp_all add: conj_def) (cases p, simp_all) lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" using conj_def by auto lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" using conj_def by auto definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" using disj_def by auto lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" using disj_def by auto definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def) lemma imp_qf[simp]: "\qfree p ; qfree q\ \ qfree (imp p q)" using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp) (cases "not p= q", auto) lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" by (unfold iff_def,cases "p=q", auto) fun check_int:: "num \ bool" where "check_int (C i) = True" | "check_int (Floor t) = True" | "check_int (Mul i t) = check_int t" | "check_int (Add t s) = (check_int t \ check_int s)" | "check_int (Neg t) = check_int t" | "check_int (CF c t s) = check_int s" | "check_int t = False" lemma check_int: "check_int t \ isint t bs" by (induct t) (auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF) lemma rdvd_left1_int: "real_of_int \t\ = t \ 1 rdvd t" by (simp add: rdvd_def,rule_tac x="\t\" in exI) simp lemma rdvd_reduce: assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" shows "real_of_int (d::int) rdvd real_of_int (c::int)*t = (real_of_int (d div g) rdvd real_of_int (c div g)*t)" proof assume d: "real_of_int d rdvd real_of_int c * t" from d rdvd_def obtain k where k_def: "real_of_int c * t = real_of_int d* real_of_int (k::int)" by auto from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto from k_def kd_def kc_def have "real_of_int g * real_of_int kc * t = real_of_int g * real_of_int kd * real_of_int k" by simp hence "real_of_int kc * t = real_of_int kd * real_of_int k" using gp by simp hence th:"real_of_int kd rdvd real_of_int kc * t" using rdvd_def by blast from kd_def gp have th':"kd = d div g" by simp from kc_def gp have "kc = c div g" by simp with th th' show "real_of_int (d div g) rdvd real_of_int (c div g) * t" by simp next assume d: "real_of_int (d div g) rdvd real_of_int (c div g) * t" from gp have gnz: "g \ 0" by simp thus "real_of_int d rdvd real_of_int c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real_of_int (c div g) * t"] real_of_int_div[OF gd] real_of_int_div[OF gc] by simp qed definition simpdvd :: "int \ num \ (int \ num)" where "simpdvd d t \ (let g = numgcd t in if g > 1 then (let g' = gcd d g in if g' = 1 then (d, t) else (d div g',reducecoeffh t g')) else (d, t))" lemma simpdvd: assumes tnz: "nozerocoeff t" and dnz: "d \ 0" shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)" proof- let ?g = "numgcd t" let ?g' = "gcd d ?g" {assume "\ ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)} moreover {assume g1:"?g>1" hence g0: "?g > 0" by simp from g1 dnz have gp0: "?g' \ 0" by simp hence g'p: "?g' > 0" using gcd_ge_0_int[where x="d" and y="numgcd t"] by arith hence "?g'= 1 \ ?g' > 1" by arith moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)} moreover {assume g'1:"?g'>1" from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" .. let ?tt = "reducecoeffh t ?g'" let ?t = "Inum bs ?tt" have gpdg: "?g' dvd ?g" by simp have gpdd: "?g' dvd d" by simp have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] have th2:"real_of_int ?g' * ?t = Inum bs t" by simp from assms g1 g0 g'1 have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" by (simp add: simpdvd_def Let_def) also have "\ = (real_of_int d rdvd (Inum bs t))" using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]] th2[symmetric] by simp finally have ?thesis by simp } ultimately have ?thesis by auto } ultimately show ?thesis by blast qed fun simpfm :: "fm \ fm" where "simpfm (And p q) = conj (simpfm p) (simpfm q)" | "simpfm (Or p q) = disj (simpfm p) (simpfm q)" | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" | "simpfm (Not p) = not (simpfm p)" | "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F | _ \ Lt (reducecoeff a'))" | "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le (reducecoeff a'))" | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt (reducecoeff a'))" | "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Ge (reducecoeff a'))" | "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \ if (v = 0) then T else F | _ \ Eq (reducecoeff a'))" | "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ NEq (reducecoeff a'))" | "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) else if (\i\ = 1) \ check_int a then T else let a' = simpnum a in case a' of C v \ if (i dvd v) then T else F | _ \ (let (d,t) = simpdvd i a' in Dvd d t))" | "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) else if (\i\ = 1) \ check_int a then F else let a' = simpnum a in case a' of C v \ if (\(i dvd v)) then T else F | _ \ (let (d,t) = simpdvd i a' in NDvd d t))" | "simpfm p = p" lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p" proof(induct p rule: simpfm.induct) case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a < 0 = (real_of_int ?g * ?r < real_of_int ?g * 0)" by simp also have "\ = (?r < 0)" using gp by (simp only: mult_less_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} ultimately show ?case by blast next case (7 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ real_of_int ?g * 0)" by simp also have "\ = (?r \ 0)" using gp by (simp only: mult_le_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} ultimately show ?case by blast next case (8 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a > 0 = (real_of_int ?g * ?r > real_of_int ?g * 0)" by simp also have "\ = (?r > 0)" using gp by (simp only: mult_less_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} ultimately show ?case by blast next case (9 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ real_of_int ?g * 0)" by simp also have "\ = (?r \ 0)" using gp by (simp only: mult_le_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} ultimately show ?case by blast next case (10 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a = 0 = (real_of_int ?g * ?r = 0)" by simp also have "\ = (?r = 0)" using gp by simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} ultimately show ?case by blast next case (11 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume H:"\ (\ v. ?sa = C v)" let ?g = "numgcd ?sa" let ?rsa = "reducecoeff ?sa" let ?r = "Inum bs ?rsa" have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) hence gp: "real_of_int ?g > 0" by simp have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ 0)" by simp also have "\ = (?r \ 0)" using gp by simp finally have ?case using H by (cases "?sa") (simp_all add: Let_def) } ultimately show ?case by blast next case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 \ (\i\ = 1 \ check_int a) \ (i\0 \ ((\i\ \ 1) \ (\ check_int a)))" by auto {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)} moreover {assume ai1: "\i\ = 1" and ai: "check_int a" hence "i=1 \ i= - 1" by arith moreover {assume i1: "i = 1" from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] have ?case using i1 ai by simp } moreover {assume i1: "i = - 1" from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] rdvd_abs1[where d="- 1" and t="Inum bs a"] have ?case using i1 ai by simp } ultimately have ?case by blast} moreover {assume inz: "i\0" and cond: "(\i\ \ 1) \ (\ check_int a)" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "\i\ = 1", auto simp add: int_rdvd_iff) } moreover {assume H:"\ (\ v. ?sa = C v)" hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def) from simpnum_nz have nz:"nozerocoeff ?sa" by simp from simpdvd [OF nz inz] th have ?case using sa by simp} ultimately have ?case by blast} ultimately show ?case by blast next case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 \ (\i\ = 1 \ check_int a) \ (i\0 \ ((\i\ \ 1) \ (\ check_int a)))" by auto {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)} moreover {assume ai1: "\i\ = 1" and ai: "check_int a" hence "i=1 \ i= - 1" by arith moreover {assume i1: "i = 1" from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] have ?case using i1 ai by simp } moreover {assume i1: "i = - 1" from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] rdvd_abs1[where d="- 1" and t="Inum bs a"] have ?case using i1 ai by simp } ultimately have ?case by blast} moreover {assume inz: "i\0" and cond: "(\i\ \ 1) \ (\ check_int a)" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "\i\ = 1", auto simp add: int_rdvd_iff) } moreover {assume H:"\ (\ v. ?sa = C v)" hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def) from simpnum_nz have nz:"nozerocoeff ?sa" by simp from simpdvd [OF nz inz] th have ?case using sa by simp} ultimately have ?case by blast} ultimately show ?case by blast qed (induct p rule: simpfm.induct, simp_all) lemma simpdvd_numbound0: "numbound0 t \ numbound0 (snd (simpdvd d t))" by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0) lemma simpfm_bound0[simp]: "bound0 p \ bound0 (simpfm p)" proof(induct p rule: simpfm.induct) case (6 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (7 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (8 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (9 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (10 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (11 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) next case (12 i a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) next case (13 i a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) qed(auto simp add: disj_def imp_def iff_def conj_def) lemma simpfm_qf[simp]: "qfree p \ qfree (simpfm p)" by (induct p rule: simpfm.induct, auto simp add: Let_def) (case_tac "simpnum a",auto simp add: split_def Let_def)+ (* Generic quantifier elimination *) definition list_conj :: "fm list \ fm" where "list_conj ps \ foldr conj ps T" lemma list_conj: "Ifm bs (list_conj ps) = (\p\ set ps. Ifm bs p)" by (induct ps, auto simp add: list_conj_def) lemma list_conj_qf: " \p\ set ps. qfree p \ qfree (list_conj ps)" by (induct ps, auto simp add: list_conj_def) lemma list_conj_nb: " \p\ set ps. bound0 p \ bound0 (list_conj ps)" by (induct ps, auto simp add: list_conj_def) definition CJNB :: "(fm \ fm) \ fm \ fm" where "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs in conj (decr (list_conj yes)) (f (list_conj no)))" lemma CJNB_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" shows "\ bs p. qfree p \ qfree (CJNB qe p) \ (Ifm bs ((CJNB qe p)) = Ifm bs (E p))" proof(clarify) fix bs p assume qfp: "qfree p" let ?cjs = "conjuncts p" let ?yes = "fst (List.partition bound0 ?cjs)" let ?no = "snd (List.partition bound0 ?cjs)" let ?cno = "list_conj ?no" let ?cyes = "list_conj ?yes" have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf) from conjuncts_qf[OF qfp] partition_set[OF part] have " \q\ set ?no. qfree q" by auto hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) with qe have cno_qf:"qfree (qe ?cno )" and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+ from cno_qf yes_qf have qf: "qfree (CJNB qe p)" by (simp add: CJNB_def Let_def split_def) {fix bs from conjuncts have "Ifm bs p = (\q\ set ?cjs. Ifm bs q)" by blast also have "\ = ((\q\ set ?yes. Ifm bs q) \ (\q\ set ?no. Ifm bs q))" using partition_set[OF part] by auto finally have "Ifm bs p = ((Ifm bs ?cyes) \ (Ifm bs ?cno))" using list_conj by simp} hence "Ifm bs (E p) = (\x. (Ifm (x#bs) ?cyes) \ (Ifm (x#bs) ?cno))" by simp also fix y have "\ = (\x. (Ifm (y#bs) ?cyes) \ (Ifm (x#bs) ?cno))" using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast also have "\ = (Ifm bs (decr ?cyes) \ Ifm bs (E ?cno))" by (auto simp add: decr[OF yes_nb] simp del: partition_filter_conv) also have "\ = (Ifm bs (conj (decr ?cyes) (qe ?cno)))" using qe[rule_format, OF no_qf] by auto finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" by (simp add: Let_def CJNB_def split_def) with qf show "qfree (CJNB qe p) \ Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast qed fun qelim :: "fm \ (fm \ fm) \ fm" where "qelim (E p) = (\ qe. DJ (CJNB qe) (qelim p qe))" | "qelim (A p) = (\ qe. not (qe ((qelim (Not p) qe))))" | "qelim (Not p) = (\ qe. not (qelim p qe))" | "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" | "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" | "qelim (Imp p q) = (\ qe. disj (qelim (Not p) qe) (qelim q qe))" | "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" | "qelim p = (\ y. simpfm p)" lemma qelim_ci: assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] by (induct p rule: qelim.induct) (auto simp del: simpfm.simps) text \The \\\ Part\ text\Linearity for fm where Bound 0 ranges over \\\\ fun zsplit0 :: "num \ int \ num" (* splits the bounded from the unbounded part*) where "zsplit0 (C c) = (0,C c)" | "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" | "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)" | "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)" | "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" | "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; (ib,b') = zsplit0 b in (ia+ib, Add a' b'))" | "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; (ib,b') = zsplit0 b in (ia-ib, Sub a' b'))" | "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" | "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))" lemma zsplit0_I: shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real_of_int (x::int)) #bs) (CN 0 n a) = Inum (real_of_int x #bs) t) \ numbound0 a" (is "\ n a. ?S t = (n,a) \ (?I x (CN 0 n a) = ?I x t) \ ?N a") proof(induct t rule: zsplit0.induct) case (1 c n a) thus ?case by auto next case (2 m n a) thus ?case by (cases "m=0") auto next case (3 n i a n a') thus ?case by auto next case (4 c a b n a') thus ?case by auto next case (5 t n a) let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \ n=-?nt" using 5 by (simp add: Let_def split_def) from abj 5 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from th2[simplified] th[simplified] show ?case by simp next case (6 s t n a) let ?ns = "fst (zsplit0 s)" let ?as = "snd (zsplit0 s)" let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abjs: "zsplit0 s = (?ns,?as)" by simp moreover have abjt: "zsplit0 t = (?nt,?at)" by simp ultimately have th: "a=Add ?as ?at \ n=?ns + ?nt" using 6 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast from 6 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: distrib_right) next case (7 s t n a) let ?ns = "fst (zsplit0 s)" let ?as = "snd (zsplit0 s)" let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abjs: "zsplit0 s = (?ns,?as)" by simp moreover have abjt: "zsplit0 t = (?nt,?at)" by simp ultimately have th: "a=Sub ?as ?at \ n=?ns - ?nt" using 7 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast from 7 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: left_diff_distrib) next case (8 i t n a) let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \ n=i*?nt" using 8 by (simp add: Let_def split_def) from abj 8 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast hence "?I x (Mul i t) = (real_of_int i) * ?I x (CN 0 ?nt ?at)" by simp also have "\ = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: distrib_left) finally show ?case using th th2 by simp next case (9 t n a) let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \ n=?nt" using 9 by (simp add: Let_def split_def) from abj 9 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast hence na: "?N a" using th by simp have th': "(real_of_int ?nt)*(real_of_int x) = real_of_int (?nt * x)" by simp have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp also have "\ = real_of_int \real_of_int ?nt * real_of_int x + ?I x ?at\" by simp also have "\ = real_of_int \?I x ?at + real_of_int (?nt * x)\" by (simp add: ac_simps) also have "\ = real_of_int (\?I x ?at\ + (?nt * x))" by (simp add: of_int_mult[symmetric] del: of_int_mult) also have "\ = real_of_int (?nt)*(real_of_int x) + real_of_int \?I x ?at\" by (simp add: ac_simps) finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp with na show ?case by simp qed fun iszlfm :: "fm \ real list \ bool" (* Linearity test for fm *) where "iszlfm (And p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" | "iszlfm (Or p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" | "iszlfm (Eq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (NEq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (Lt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (Le (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (Gt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (Ge (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" | "iszlfm (Dvd i (CN 0 c e)) = (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" | "iszlfm (NDvd i (CN 0 c e))= (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" | "iszlfm p = (\ bs. isatom p \ (bound0 p))" lemma zlin_qfree: "iszlfm p bs \ qfree p" by (induct p rule: iszlfm.induct) auto lemma iszlfm_gen: assumes lp: "iszlfm p (x#bs)" shows "\ y. iszlfm p (y#bs)" proof fix y show "iszlfm p (y#bs)" using lp by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"]) qed lemma conj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (conj p q) bs" using conj_def by (cases p,auto) lemma disj_zl[simp]: "iszlfm p bs \ iszlfm q bs \ iszlfm (disj p q) bs" using disj_def by (cases p,auto) fun zlfm :: "fm \ fm" (* Linearity transformation for fm *) where "zlfm (And p q) = conj (zlfm p) (zlfm q)" | "zlfm (Or p q) = disj (zlfm p) (zlfm q)" | "zlfm (Imp p q) = disj (zlfm (Not p)) (zlfm q)" | "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (Not p)) (zlfm (Not q)))" | "zlfm (Lt a) = (let (c,r) = zsplit0 a in if c=0 then Lt r else if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" | "zlfm (Le a) = (let (c,r) = zsplit0 a in if c=0 then Le r else if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r))) else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" | "zlfm (Gt a) = (let (c,r) = zsplit0 a in if c=0 then Gt r else if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" | "zlfm (Ge a) = (let (c,r) = zsplit0 a in if c=0 then Ge r else if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" | "zlfm (Eq a) = (let (c,r) = zsplit0 a in if c=0 then Eq r else if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r))) else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))" | "zlfm (NEq a) = (let (c,r) = zsplit0 a in if c=0 then NEq r else if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r))) else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))" | "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) else (let (c,r) = zsplit0 a in if c=0 then Dvd \i\ r else if c>0 then And (Eq (Sub (Floor r) r)) (Dvd \i\ (CN 0 c (Floor r))) else And (Eq (Sub (Floor r) r)) (Dvd \i\ (CN 0 (-c) (Neg (Floor r))))))" | "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) else (let (c,r) = zsplit0 a in if c=0 then NDvd \i\ r else if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd \i\ (CN 0 c (Floor r))) else Or (NEq (Sub (Floor r) r)) (NDvd \i\ (CN 0 (-c) (Neg (Floor r))))))" | "zlfm (Not (And p q)) = disj (zlfm (Not p)) (zlfm (Not q))" | "zlfm (Not (Or p q)) = conj (zlfm (Not p)) (zlfm (Not q))" | "zlfm (Not (Imp p q)) = conj (zlfm p) (zlfm (Not q))" | "zlfm (Not (Iff p q)) = disj (conj(zlfm p) (zlfm(Not q))) (conj (zlfm(Not p)) (zlfm q))" | "zlfm (Not (Not p)) = zlfm p" | "zlfm (Not T) = F" | "zlfm (Not F) = T" | "zlfm (Not (Lt a)) = zlfm (Ge a)" | "zlfm (Not (Le a)) = zlfm (Gt a)" | "zlfm (Not (Gt a)) = zlfm (Le a)" | "zlfm (Not (Ge a)) = zlfm (Lt a)" | "zlfm (Not (Eq a)) = zlfm (NEq a)" | "zlfm (Not (NEq a)) = zlfm (Eq a)" | "zlfm (Not (Dvd i a)) = zlfm (NDvd i a)" | "zlfm (Not (NDvd i a)) = zlfm (Dvd i a)" | "zlfm p = p" lemma split_int_less_real: "(real_of_int (a::int) < b) = (a < \b\ \ (a = \b\ \ real_of_int \b\ < b))" proof( auto) assume alb: "real_of_int a < b" and agb: "\ a < \b\" from agb have "\b\ \ a" by simp hence th: "b < real_of_int a + 1" by (simp only: floor_le_iff) from floor_eq[OF alb th] show "a = \b\" by simp next assume alb: "a < \b\" hence "real_of_int a < real_of_int \b\" by simp moreover have "real_of_int \b\ \ b" by simp ultimately show "real_of_int a < b" by arith qed lemma split_int_less_real': "(real_of_int (a::int) + b < 0) = (real_of_int a - real_of_int \- b\ < 0 \ (real_of_int a - real_of_int \- b\ = 0 \ real_of_int \- b\ + b < 0))" proof- have "(real_of_int a + b <0) = (real_of_int a < -b)" by arith with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith qed lemma split_int_gt_real': "(real_of_int (a::int) + b > 0) = (real_of_int a + real_of_int \b\ > 0 \ (real_of_int a + real_of_int \b\ = 0 \ real_of_int \b\ - b < 0))" proof- have th: "(real_of_int a + b >0) = (real_of_int (-a) + (-b)< 0)" by arith show ?thesis by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) (auto simp add: algebra_simps) qed lemma split_int_le_real: "(real_of_int (a::int) \ b) = (a \ \b\ \ (a = \b\ \ real_of_int \b\ < b))" proof( auto) assume alb: "real_of_int a \ b" and agb: "\ a \ \b\" from alb have "\real_of_int a\ \ \b\" by (simp only: floor_mono) hence "a \ \b\" by simp with agb show "False" by simp next assume alb: "a \ \b\" hence "real_of_int a \ real_of_int \b\" by (simp only: floor_mono) also have "\\ b" by simp finally show "real_of_int a \ b" . qed lemma split_int_le_real': "(real_of_int (a::int) + b \ 0) = (real_of_int a - real_of_int \- b\ \ 0 \ (real_of_int a - real_of_int \- b\ = 0 \ real_of_int \- b\ + b < 0))" proof- have "(real_of_int a + b \0) = (real_of_int a \ -b)" by arith with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith qed lemma split_int_ge_real': "(real_of_int (a::int) + b \ 0) = (real_of_int a + real_of_int \b\ \ 0 \ (real_of_int a + real_of_int \b\ = 0 \ real_of_int \b\ - b < 0))" proof- have th: "(real_of_int a + b \0) = (real_of_int (-a) + (-b) \ 0)" by arith show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) (simp add: algebra_simps ,arith) qed lemma split_int_eq_real: "(real_of_int (a::int) = b) = ( a = \b\ \ b = real_of_int \b\)" (is "?l = ?r") by auto lemma split_int_eq_real': "(real_of_int (a::int) + b = 0) = ( a - \- b\ = 0 \ real_of_int \- b\ + b = 0)" (is "?l = ?r") proof- have "?l = (real_of_int a = -b)" by arith with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith qed lemma zlfm_I: assumes qfp: "qfree p" shows "(Ifm (real_of_int i #bs) (zlfm p) = Ifm (real_of_int i# bs) p) \ iszlfm (zlfm p) (real_of_int (i::int) #bs)" (is "(?I (?l p) = ?I p) \ ?L (?l p)") using qfp proof(induct p rule: zlfm.induct) case (5 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def,rename_tac nat a b,case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def) + also have "\ = (?I (?l (Lt a)))" + unfolding split_int_less_real'[where a="?c*i" and b="?N ?r"] + by (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast next case (6 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat",simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast next case (7 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast next case (8 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast next case (9 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult,arith) finally have ?case using l by simp} ultimately show ?case by blast next case (10 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult,arith) finally have ?case using l by simp} ultimately show ?case by blast next case (11 j a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) hence ?case using 11 j by (simp del: zlfm.simps add: rdvd_left_0_eq)} moreover {assume "?c=0" and "j\0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) also have "\ = (real_of_int \j\ rdvd real_of_int (?c*i) + (?N ?r))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp also have "\ = (\j\ dvd \(?N ?r) + real_of_int (?c*i)\ \ (real_of_int \(?N ?r) + real_of_int (?c*i)\ = (real_of_int (?c*i) + (?N ?r))))" by(simp only: int_rdvd_real[where i="\j\" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) also have "\ = (real_of_int \j\ rdvd real_of_int (?c*i) + (?N ?r))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp also have "\ = (\j\ dvd \(?N ?r) + real_of_int (?c*i)\ \ (real_of_int \(?N ?r) + real_of_int (?c*i)\ = (real_of_int (?c*i) + (?N ?r))))" by(simp only: int_rdvd_real[where i="\j\" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (Dvd j a)))" using cn cnz jnz using rdvd_minus [where d="\j\" and t="real_of_int (?c*i + \?N ?r\)", simplified, symmetric] by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast next case (12 j a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) hence ?case using 12 j by (simp del: zlfm.simps add: rdvd_left_0_eq)} moreover {assume "?c=0" and "j\0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) also have "\ = (\ (real_of_int \j\ rdvd real_of_int (?c*i) + (?N ?r)))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp also have "\ = (\ (\j\ dvd \(?N ?r) + real_of_int (?c*i)\ \ (real_of_int \(?N ?r) + real_of_int (?c*i)\ = (real_of_int (?c*i) + (?N ?r)))))" by(simp only: int_rdvd_real[where i="\j\" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) also have "\ = (\ (real_of_int \j\ rdvd real_of_int (?c*i) + (?N ?r)))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp also have "\ = (\ (\j\ dvd \(?N ?r) + real_of_int (?c*i)\ \ (real_of_int \(?N ?r) + real_of_int (?c*i)\ = (real_of_int (?c*i) + (?N ?r)))))" by(simp only: int_rdvd_real[where i="\j\" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (NDvd j a)))" using cn cnz jnz using rdvd_minus [where d="\j\" and t="real_of_int (?c*i + \?N ?r\)", simplified, symmetric] by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast qed auto text\plusinf : Virtual substitution of \+\\ minusinf: Virtual substitution of \-\\ \\\ Compute lcm \d| Dvd d c*x+t \ p\ \d_\\ checks if a given l divides all the ds above\ fun minusinf:: "fm \ fm" where "minusinf (And p q) = conj (minusinf p) (minusinf q)" | "minusinf (Or p q) = disj (minusinf p) (minusinf q)" | "minusinf (Eq (CN 0 c e)) = F" | "minusinf (NEq (CN 0 c e)) = T" | "minusinf (Lt (CN 0 c e)) = T" | "minusinf (Le (CN 0 c e)) = T" | "minusinf (Gt (CN 0 c e)) = F" | "minusinf (Ge (CN 0 c e)) = F" | "minusinf p = p" lemma minusinf_qfree: "qfree p \ qfree (minusinf p)" by (induct p rule: minusinf.induct, auto) fun plusinf:: "fm \ fm" where "plusinf (And p q) = conj (plusinf p) (plusinf q)" | "plusinf (Or p q) = disj (plusinf p) (plusinf q)" | "plusinf (Eq (CN 0 c e)) = F" | "plusinf (NEq (CN 0 c e)) = T" | "plusinf (Lt (CN 0 c e)) = F" | "plusinf (Le (CN 0 c e)) = F" | "plusinf (Gt (CN 0 c e)) = T" | "plusinf (Ge (CN 0 c e)) = T" | "plusinf p = p" fun \ :: "fm \ int" where "\ (And p q) = lcm (\ p) (\ q)" | "\ (Or p q) = lcm (\ p) (\ q)" | "\ (Dvd i (CN 0 c e)) = i" | "\ (NDvd i (CN 0 c e)) = i" | "\ p = 1" fun d_\ :: "fm \ int \ bool" where "d_\ (And p q) = (\ d. d_\ p d \ d_\ q d)" | "d_\ (Or p q) = (\ d. d_\ p d \ d_\ q d)" | "d_\ (Dvd i (CN 0 c e)) = (\ d. i dvd d)" | "d_\ (NDvd i (CN 0 c e)) = (\ d. i dvd d)" | "d_\ p = (\ d. True)" lemma delta_mono: assumes lin: "iszlfm p bs" and d: "d dvd d'" and ad: "d_\ p d" shows "d_\ p d'" using lin ad d proof(induct p rule: iszlfm.induct) case (9 i c e) thus ?case using d by (simp add: dvd_trans[of "i" "d" "d'"]) next case (10 i c e) thus ?case using d by (simp add: dvd_trans[of "i" "d" "d'"]) qed simp_all lemma \ : assumes lin:"iszlfm p bs" shows "d_\ p (\ p) \ \ p >0" using lin proof (induct p rule: iszlfm.induct) case (1 p q) let ?d = "\ (And p q)" from 1 lcm_pos_int have dp: "?d >0" by simp have d1: "\ p dvd \ (And p q)" using 1 by simp hence th: "d_\ p ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast have "\ q dvd \ (And p q)" using 1 by simp hence th': "d_\ q ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast from th th' dp show ?case by simp next case (2 p q) let ?d = "\ (And p q)" from 2 lcm_pos_int have dp: "?d >0" by simp have "\ p dvd \ (And p q)" using 2 by simp hence th: "d_\ p ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast have "\ q dvd \ (And p q)" using 2 by simp hence th': "d_\ q ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast from th th' dp show ?case by simp qed simp_all lemma minusinf_inf: assumes linp: "iszlfm p (a # bs)" shows "\ (z::int). \ x < z. Ifm ((real_of_int x)#bs) (minusinf p) = Ifm ((real_of_int x)#bs) p" (is "?P p" is "\ (z::int). \ x < z. ?I x (?M p) = ?I x p") using linp proof (induct p rule: minusinf.induct) case (1 f g) then have "?P f" by simp then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast with 1 have "?P g" by simp then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast let ?z = "min z1 z2" from z1_def z2_def have "\ x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp thus ?case by blast next case (2 f g) then have "?P f" by simp then obtain z1 where z1_def: "\ x < z1. ?I x (?M f) = ?I x f" by blast with 2 have "?P g" by simp then obtain z2 where z2_def: "\ x < z2. ?I x (?M g) = ?I x g" by blast let ?z = "min z1 z2" from z1_def z2_def have "\ x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp thus ?case by blast next case (3 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 3 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next case (4 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 4 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next case (5 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 5 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (6 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 6 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (7 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 7 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) thus "\ (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e>0)" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (8 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 8 have nbe: "numbound0 e" by simp fix y have "\ x < \- (Inum (y#bs) e) / (real_of_int c)\. ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) thus "\ real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast qed simp_all lemma minusinf_repeats: assumes d: "d_\ p d" and linp: "iszlfm p (a # bs)" shows "Ifm ((real_of_int(x - k*d))#bs) (minusinf p) = Ifm (real_of_int x #bs) (minusinf p)" using linp d proof(induct p rule: iszlfm.induct) case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast show ?case proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) assume "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next assume "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)" by (simp add: di_def) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" using rdvd_def by simp qed next case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast show ?case proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) assume "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next assume "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)" by (simp add: di_def) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" using rdvd_def by simp qed qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int(x - k*d)" and b'="real_of_int x"] simp del: of_int_mult of_int_diff) lemma minusinf_ex: assumes lin: "iszlfm p (real_of_int (a::int) #bs)" and exmi: "\ (x::int). Ifm (real_of_int x#bs) (minusinf p)" (is "\ x. ?P1 x") shows "\ (x::int). Ifm (real_of_int x#bs) p" (is "\ x. ?P x") proof- let ?d = "\ p" from \ [OF lin] have dpos: "?d >0" by simp from \ [OF lin] have alld: "d_\ p ?d" by simp from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P1 x = ?P1 (x - (k * ?d))" by simp from minusinf_inf[OF lin] have th2:"\ z. \ x. x (?P x = ?P1 x)" by blast from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast qed lemma minusinf_bex: assumes lin: "iszlfm p (real_of_int (a::int) #bs)" shows "(\ (x::int). Ifm (real_of_int x#bs) (minusinf p)) = (\ (x::int)\ {1..\ p}. Ifm (real_of_int x#bs) (minusinf p))" (is "(\ x. ?P x) = _") proof- let ?d = "\ p" from \ [OF lin] have dpos: "?d >0" by simp from \ [OF lin] have alld: "d_\ p ?d" by simp from minusinf_repeats[OF alld lin] have th1:"\ x k. ?P x = ?P (x - (k * ?d))" by simp from periodic_finite_ex[OF dpos th1] show ?thesis by blast qed lemma dvd1_eq1: "x > 0 \ is_unit x \ x = 1" for x :: int by simp fun a_\ :: "fm \ int \ fm" (* adjusts the coefficients of a formula *) where "a_\ (And p q) = (\ k. And (a_\ p k) (a_\ q k))" | "a_\ (Or p q) = (\ k. Or (a_\ p k) (a_\ q k))" | "a_\ (Eq (CN 0 c e)) = (\ k. Eq (CN 0 1 (Mul (k div c) e)))" | "a_\ (NEq (CN 0 c e)) = (\ k. NEq (CN 0 1 (Mul (k div c) e)))" | "a_\ (Lt (CN 0 c e)) = (\ k. Lt (CN 0 1 (Mul (k div c) e)))" | "a_\ (Le (CN 0 c e)) = (\ k. Le (CN 0 1 (Mul (k div c) e)))" | "a_\ (Gt (CN 0 c e)) = (\ k. Gt (CN 0 1 (Mul (k div c) e)))" | "a_\ (Ge (CN 0 c e)) = (\ k. Ge (CN 0 1 (Mul (k div c) e)))" | "a_\ (Dvd i (CN 0 c e)) =(\ k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | "a_\ (NDvd i (CN 0 c e))=(\ k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | "a_\ p = (\ k. p)" fun d_\ :: "fm \ int \ bool" (* tests if all coeffs c of c divide a given l*) where "d_\ (And p q) = (\ k. (d_\ p k) \ (d_\ q k))" | "d_\ (Or p q) = (\ k. (d_\ p k) \ (d_\ q k))" | "d_\ (Eq (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (NEq (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (Lt (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (Le (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (Gt (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (Ge (CN 0 c e)) = (\ k. c dvd k)" | "d_\ (Dvd i (CN 0 c e)) =(\ k. c dvd k)" | "d_\ (NDvd i (CN 0 c e))=(\ k. c dvd k)" | "d_\ p = (\ k. True)" fun \ :: "fm \ int" (* computes the lcm of all coefficients of x*) where "\ (And p q) = lcm (\ p) (\ q)" | "\ (Or p q) = lcm (\ p) (\ q)" | "\ (Eq (CN 0 c e)) = c" | "\ (NEq (CN 0 c e)) = c" | "\ (Lt (CN 0 c e)) = c" | "\ (Le (CN 0 c e)) = c" | "\ (Gt (CN 0 c e)) = c" | "\ (Ge (CN 0 c e)) = c" | "\ (Dvd i (CN 0 c e)) = c" | "\ (NDvd i (CN 0 c e))= c" | "\ p = 1" fun \ :: "fm \ num list" where "\ (And p q) = (\ p @ \ q)" | "\ (Or p q) = (\ p @ \ q)" | "\ (Eq (CN 0 c e)) = [Sub (C (- 1)) e]" | "\ (NEq (CN 0 c e)) = [Neg e]" | "\ (Lt (CN 0 c e)) = []" | "\ (Le (CN 0 c e)) = []" | "\ (Gt (CN 0 c e)) = [Neg e]" | "\ (Ge (CN 0 c e)) = [Sub (C (- 1)) e]" | "\ p = []" fun \ :: "fm \ num list" where "\ (And p q) = (\ p @ \ q)" | "\ (Or p q) = (\ p @ \ q)" | "\ (Eq (CN 0 c e)) = [Add (C (- 1)) e]" | "\ (NEq (CN 0 c e)) = [e]" | "\ (Lt (CN 0 c e)) = [e]" | "\ (Le (CN 0 c e)) = [Add (C (- 1)) e]" | "\ (Gt (CN 0 c e)) = []" | "\ (Ge (CN 0 c e)) = []" | "\ p = []" fun mirror :: "fm \ fm" where "mirror (And p q) = And (mirror p) (mirror q)" | "mirror (Or p q) = Or (mirror p) (mirror q)" | "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" | "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" | "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" | "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" | "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" | "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" | "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" | "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" | "mirror p = p" lemma mirror_\_\: assumes lp: "iszlfm p (a#bs)" shows "(Inum (real_of_int (i::int)#bs)) ` set (\ p) = (Inum (real_of_int i#bs)) ` set (\ (mirror p))" using lp by (induct p rule: mirror.induct) auto lemma mirror: assumes lp: "iszlfm p (a#bs)" shows "Ifm (real_of_int (x::int)#bs) (mirror p) = Ifm (real_of_int (- x)#bs) p" using lp proof(induct p rule: iszlfm.induct) case (9 j c e) have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) = (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int x # bs) e))" by (simp only: rdvd_minus[symmetric]) from 9 th show ?case by (simp add: algebra_simps numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int x"]) next case (10 j c e) have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) = (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int x # bs) e))" by (simp only: rdvd_minus[symmetric]) from 10 th show ?case by (simp add: algebra_simps numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int x"]) qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int x" and b'="- real_of_int x"]) lemma mirror_l: "iszlfm p (a#bs) \ iszlfm (mirror p) (a#bs)" by (induct p rule: mirror.induct) (auto simp add: isint_neg) lemma mirror_d_\: "iszlfm p (a#bs) \ d_\ p 1 \ iszlfm (mirror p) (a#bs) \ d_\ (mirror p) 1" by (induct p rule: mirror.induct) (auto simp add: isint_neg) lemma mirror_\: "iszlfm p (a#bs) \ \ (mirror p) = \ p" by (induct p rule: mirror.induct) auto lemma mirror_ex: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "(\ (x::int). Ifm (real_of_int x#bs) (mirror p)) = (\ (x::int). Ifm (real_of_int x#bs) p)" (is "(\ x. ?I x ?mp) = (\ x. ?I x p)") proof(auto) fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast thus "\ x. ?I x p" by blast next fix x assume "?I x p" hence "?I (- x) ?mp" using mirror[OF lp, where x="- x", symmetric] by auto thus "\ x. ?I x ?mp" by blast qed lemma \_numbound0: assumes lp: "iszlfm p bs" shows "\ b\ set (\ p). numbound0 b" using lp by (induct p rule: \.induct,auto) lemma d_\_mono: assumes linp: "iszlfm p (a #bs)" and dr: "d_\ p l" and d: "l dvd l'" shows "d_\ p l'" using dr linp dvd_trans[of _ "l" "l'", simplified d] by (induct p rule: iszlfm.induct) simp_all lemma \_l: assumes lp: "iszlfm p (a#bs)" shows "\ b\ set (\ p). numbound0 b \ isint b (a#bs)" using lp by(induct p rule: \.induct, auto simp add: isint_add isint_c) lemma \: assumes linp: "iszlfm p (a #bs)" shows "\ p > 0 \ d_\ p (\ p)" using linp proof(induct p rule: iszlfm.induct) case (1 p q) then have dl1: "\ p dvd lcm (\ p) (\ q)" by simp from 1 have dl2: "\ q dvd lcm (\ p) (\ q)" by simp from 1 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] dl1 dl2 show ?case by (auto simp add: lcm_pos_int) next case (2 p q) then have dl1: "\ p dvd lcm (\ p) (\ q)" by simp from 2 have dl2: "\ q dvd lcm (\ p) (\ q)" by simp from 2 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] dl1 dl2 show ?case by (auto simp add: lcm_pos_int) qed (auto simp add: lcm_pos_int) lemma a_\: assumes linp: "iszlfm p (a #bs)" and d: "d_\ p l" and lp: "l > 0" shows "iszlfm (a_\ p l) (a #bs) \ d_\ (a_\ p l) 1 \ (Ifm (real_of_int (l * x) #bs) (a_\ p l) = Ifm ((real_of_int x)#bs) p)" using linp d proof (induct p rule: iszlfm.induct) case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) < (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0)" using mult_less_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" using mult_le_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) > (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e > 0)" using zero_less_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (8 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) = (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = 0)" using mult_eq_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) also fix k have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)" using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp also have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp next case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using mult_div_mod_eq [where a="l" and b="c"] by simp hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) also fix k have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)" using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp also have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp qed (simp_all add: numbound0_I[where bs="bs" and b="real_of_int (l * x)" and b'="real_of_int x"] isint_Mul del: of_int_mult) lemma a_\_ex: assumes linp: "iszlfm p (a#bs)" and d: "d_\ p l" and lp: "l>0" shows "(\ x. l dvd x \ Ifm (real_of_int x #bs) (a_\ p l)) = (\ (x::int). Ifm (real_of_int x#bs) p)" (is "(\ x. l dvd x \ ?P x) = (\ x. ?P' x)") proof- have "(\ x. l dvd x \ ?P x) = (\ (x::int). ?P (l*x))" using unity_coeff_ex[where l="l" and P="?P", simplified] by simp also have "\ = (\ (x::int). ?P' x)" using a_\[OF linp d lp] by simp finally show ?thesis . qed lemma \: assumes lp: "iszlfm p (a#bs)" and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" and nob: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real_of_int x = b + real_of_int j)" and p: "Ifm (real_of_int x#bs) p" (is "?P x") shows "?P (x - d)" using lp u d dp nob p proof(induct p rule: iszlfm.induct) case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 5 show ?case by (simp del: of_int_minus) next case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 6 show ?case by (simp del: of_int_minus) next case (7 c e) hence p: "Ifm (real_of_int x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp_all let ?e = "Inum (real_of_int x # bs) e" from ie1 have ie: "real_of_int \?e\ = ?e" using isint_iff[where n="e" and bs="a#bs"] numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"] by (simp add: isint_iff) {assume "real_of_int (x-d) +?e > 0" hence ?case using c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] by (simp del: of_int_minus)} moreover {assume H: "\ real_of_int (x-d) + ?e > 0" let ?v="Neg e" have vb: "?v \ set (\ (Gt (CN 0 c e)))" by simp from 7(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e + real_of_int j)" by auto from H p have "real_of_int x + ?e > 0 \ real_of_int x + ?e \ real_of_int d" by (simp add: c1) hence "real_of_int (x + \?e\) > real_of_int (0::int) \ real_of_int (x + \?e\) \ real_of_int d" using ie by simp hence "x + \?e\ \ 1 \ x + \?e\ \ d" by simp hence "\ (j::int) \ {1 .. d}. j = x + \?e\" by simp hence "\ (j::int) \ {1 .. d}. real_of_int x = real_of_int (- \?e\ + j)" by force hence "\ (j::int) \ {1 .. d}. real_of_int x = - ?e + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by auto} ultimately show ?case by blast next case (8 c e) hence p: "Ifm (real_of_int x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" from ie1 have ie: "real_of_int \?e\ = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) {assume "real_of_int (x-d) +?e \ 0" hence ?case using c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] by (simp del: of_int_minus)} moreover {assume H: "\ real_of_int (x-d) + ?e \ 0" let ?v="Sub (C (- 1)) e" have vb: "?v \ set (\ (Ge (CN 0 c e)))" by simp from 8(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e - 1 + real_of_int j)" by auto from H p have "real_of_int x + ?e \ 0 \ real_of_int x + ?e < real_of_int d" by (simp add: c1) hence "real_of_int (x + \?e\) \ real_of_int (0::int) \ real_of_int (x + \?e\) < real_of_int d" using ie by simp hence "x + \?e\ + 1 \ 1 \ x + \?e\ + 1 \ d" by simp hence "\ (j::int) \ {1 .. d}. j = x + \?e\ + 1" by simp hence "\ (j::int) \ {1 .. d}. x= - \?e\ - 1 + j" by (simp add: algebra_simps) hence "\ (j::int) \ {1 .. d}. real_of_int x= real_of_int (- \?e\ - 1 + j)" by presburger hence "\ (j::int) \ {1 .. d}. real_of_int x= - ?e - 1 + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by simp } ultimately show ?case by blast next case (3 c e) hence p: "Ifm (real_of_int x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" let ?v="(Sub (C (- 1)) e)" have vb: "?v \ set (\ (Eq (CN 0 c e)))" by simp from p have "real_of_int x= - ?e" by (simp add: c1) with 3(5) show ?case using dp by simp (erule ballE[where x="1"], simp_all add:algebra_simps numbound0_I[OF bn,where b="real_of_int x"and b'="a"and bs="bs"]) next case (4 c e)hence p: "Ifm (real_of_int x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" let ?v="Neg e" have vb: "?v \ set (\ (NEq (CN 0 c e)))" by simp {assume "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e \ 0" hence ?case by (simp add: c1)} moreover {assume H: "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e = 0" hence "real_of_int x = - Inum ((real_of_int (x -d)) # bs) e + real_of_int d" by simp hence "real_of_int x = - Inum (a # bs) e + real_of_int d" by (simp add: numbound0_I[OF bn,where b="real_of_int x - real_of_int d"and b'="a"and bs="bs"]) with 4(5) have ?case using dp by simp} ultimately show ?case by blast next case (9 j c e) hence p: "Ifm (real_of_int x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" from 9 have "isint e (a #bs)" by simp hence ie: "real_of_int \?e\ = ?e" using isint_iff[where n="e" and bs="(real_of_int x)#bs"] numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp from c1 ie[symmetric] have "?p x = (real_of_int j rdvd real_of_int (x + \?e\))" by simp also have "\ = (j dvd x + \?e\)" using int_rdvd_real[where i="j" and x="real_of_int (x + \?e\)"] by simp also have "\ = (j dvd x - d + \?e\)" using dvd_period[OF id, where x="x" and c="-1" and t="\?e\"] by simp also have "\ = (real_of_int j rdvd real_of_int (x - d + \?e\))" using int_rdvd_real[where i="j" and x="real_of_int (x - d + \?e\)",symmetric, simplified] ie by simp also have "\ = (real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp finally show ?case using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp next case (10 j c e) hence p: "Ifm (real_of_int x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" from 10 have "isint e (a#bs)" by simp hence ie: "real_of_int \?e\ = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp from c1 ie[symmetric] have "?p x = (\ real_of_int j rdvd real_of_int (x + \?e\))" by simp also have "\ = (\ j dvd x + \?e\)" using int_rdvd_real[where i="j" and x="real_of_int (x + \?e\)"] by simp also have "\ = (\ j dvd x - d + \?e\)" using dvd_period[OF id, where x="x" and c="-1" and t="\?e\"] by simp also have "\ = (\ real_of_int j rdvd real_of_int (x - d + \?e\))" using int_rdvd_real[where i="j" and x="real_of_int (x - d + \?e\)",symmetric, simplified] ie by simp also have "\ = (\ real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp finally show ?case using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int (x - d)" and b'="real_of_int x"] simp del: of_int_diff) lemma \': assumes lp: "iszlfm p (a #bs)" and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") proof(clarify) fix x assume nb:"?b" and px: "?P x" hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real_of_int x = b + real_of_int j)" by auto from \[OF lp u d dp nb2 px] show "?P (x -d )" . qed lemma \_int: assumes lp: "iszlfm p bs" shows "\ b\ set (\ p). isint b bs" using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) -lemma cpmi_eq: "0 < D \ (\z::int. \x. x < z \ (P x = P1 x)) -\ \x. \(\(j::int) \ {1..D}. \(b::int) \ B. P(b+j)) \ P (x) \ P (x - D) -\ (\(x::int). \(k::int). ((P1 x)= (P1 (x-k*D)))) -\ (\(x::int). P(x)) = ((\(j::int) \ {1..D} . (P1(j))) | (\(j::int) \ {1..D}. \(b::int) \ B. P (b+j)))" -apply(rule iffI) -prefer 2 -apply(drule minusinfinity) -apply assumption+ -apply(fastforce) -apply clarsimp -apply(subgoal_tac "\k. 0<=k \ \x. P x \ P (x - k*D)") -apply(frule_tac x = x and z=z in decr_lemma) -apply(subgoal_tac "P1(x - (\x - z\ + 1) * D)") -prefer 2 -apply(subgoal_tac "0 <= (\x - z\ + 1)") -prefer 2 apply arith - apply fastforce -apply(drule (1) periodic_finite_ex) -apply blast -apply(blast dest:decr_mult_lemma) -done - +lemma cpmi_eq: + assumes "0 < D" + and "\z::int. \xx. \ (\j\{1..D}. \b\B. P (b + j)) \ P x \ P (x - D)" + and "\x k. P1 x = P1 (x - k * D)" + shows "(\x. P x) \ ((\j\{1..D}. P1 j) \ (\j\{1..D}. \b\B. P (b + j)))" (is "_=?R") +proof + assume L: "\x. P x" + show "(\j\{1..D}. P1 j) \ (\j\{1..D}. \b\B. P (b + j))" + proof clarsimp + assume \
: "\j\{1..D}. \b\B. \ P (b + j)" + then have "\k. 0\k \ \x. P x \ P (x - k*D)" + by (simp add: assms decr_mult_lemma) + with L \
assms show "\j\{1..D}. P1 j" + using periodic_finite_ex [OF \D>0\, of P1] + by (metis abs_1 abs_add_abs abs_ge_zero decr_lemma) + qed +next + assume ?R then show "\x. P x" + using decr_lemma assms by blast +qed theorem cp_thm: assumes lp: "iszlfm p (a #bs)" and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" shows "(\ (x::int). Ifm (real_of_int x #bs) p) = (\ j\ {1.. d}. Ifm (real_of_int j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p))" (is "(\ (x::int). ?P (real_of_int x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real_of_int j)))") proof- from minusinf_inf[OF lp] have th: "\(z::int). \x_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real_of_int \?I b\ = ?I b" by simp from B[rule_format] have "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\j\?D. \b\ ?B. ?P (real_of_int \?I b\ + real_of_int j))" by simp also have "\ = (\j\?D. \b\ ?B. ?P (real_of_int (\?I b\ + j)))" by simp also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" by blast finally have BB': "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" by blast hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j))) \ ?P (real_of_int x) \ ?P (real_of_int (x - d))" using \'[OF lp u d dp] by blast from minusinf_repeats[OF d lp] have th3: "\ x k. ?M x = ?M (x-k*d)" by simp from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast qed (* Reddy and Loveland *) fun \ :: "fm \ (num \ int) list" (* Compute the Reddy and Loveland Bset*) where "\ (And p q) = (\ p @ \ q)" | "\ (Or p q) = (\ p @ \ q)" | "\ (Eq (CN 0 c e)) = [(Sub (C (- 1)) e,c)]" | "\ (NEq (CN 0 c e)) = [(Neg e,c)]" | "\ (Lt (CN 0 c e)) = []" | "\ (Le (CN 0 c e)) = []" | "\ (Gt (CN 0 c e)) = [(Neg e, c)]" | "\ (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]" | "\ p = []" fun \_\:: "fm \ num \ int \ fm" (* Performs the modified substitution of Reddy and Loveland*) where "\_\ (And p q) = (\ (t,k). And (\_\ p (t,k)) (\_\ q (t,k)))" | "\_\ (Or p q) = (\ (t,k). Or (\_\ p (t,k)) (\_\ q (t,k)))" | "\_\ (Eq (CN 0 c e)) = (\ (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) else (Eq (Add (Mul c t) (Mul k e))))" | "\_\ (NEq (CN 0 c e)) = (\ (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) else (NEq (Add (Mul c t) (Mul k e))))" | "\_\ (Lt (CN 0 c e)) = (\ (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) else (Lt (Add (Mul c t) (Mul k e))))" | "\_\ (Le (CN 0 c e)) = (\ (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) else (Le (Add (Mul c t) (Mul k e))))" | "\_\ (Gt (CN 0 c e)) = (\ (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) else (Gt (Add (Mul c t) (Mul k e))))" | "\_\ (Ge (CN 0 c e)) = (\ (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) else (Ge (Add (Mul c t) (Mul k e))))" | "\_\ (Dvd i (CN 0 c e)) =(\ (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) else (Dvd (i*k) (Add (Mul c t) (Mul k e))))" | "\_\ (NDvd i (CN 0 c e))=(\ (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) else (NDvd (i*k) (Add (Mul c t) (Mul k e))))" | "\_\ p = (\ (t,k). p)" fun \_\ :: "fm \ (num \ int) list" where "\_\ (And p q) = (\_\ p @ \_\ q)" | "\_\ (Or p q) = (\_\ p @ \_\ q)" | "\_\ (Eq (CN 0 c e)) = [(Add (C (- 1)) e,c)]" | "\_\ (NEq (CN 0 c e)) = [(e,c)]" | "\_\ (Lt (CN 0 c e)) = [(e,c)]" | "\_\ (Le (CN 0 c e)) = [(Add (C (- 1)) e,c)]" | "\_\ p = []" (* Simulates normal substituion by modifying the formula see correctness theorem *) definition \ :: "fm \ int \ num \ fm" where "\ p k t \ And (Dvd k t) (\_\ p (t,k))" lemma \_\: assumes linp: "iszlfm p (real_of_int (x::int)#bs)" and kpos: "real_of_int k > 0" and tnb: "numbound0 t" and tint: "isint t (real_of_int x#bs)" and kdt: "k dvd \Inum (b'#bs) t\" shows "Ifm (real_of_int x#bs) (\_\ p (t,k)) = (Ifm ((real_of_int (\Inum (b'#bs) t\ div k))#bs) p)" (is "?I (real_of_int x) (?s p) = (?I (real_of_int (\?N b' t\ div k)) p)" is "_ = (?I ?tk p)") using linp kpos tnb proof(induct p rule: \_\.induct) case (3 c e) from 3 have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from kpos have knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Eq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k = 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (4 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from kpos have knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (NEq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (5 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Lt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k < 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (6 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Le (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (7 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Gt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k > 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (8 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Ge (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (9 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Dvd i (CN 0 c e))) = (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k)" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast next case (10 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int \?N (real_of_int x) t\ = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (NDvd i (CN 0 c e))) = (\ (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k))" using real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast qed (simp_all add: bound0_I[where bs="bs" and b="real_of_int (\?N b' t\ div k)" and b'="real_of_int x"] numbound0_I[where bs="bs" and b="real_of_int (\?N b' t\ div k)" and b'="real_of_int x"]) lemma \_\_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t" shows "bound0 (\_\ p (t,k))" using lp by (induct p rule: iszlfm.induct, auto simp add: nb) lemma \_l: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "\ (b,k) \ set (\ p). k >0 \ numbound0 b \ isint b (real_of_int i#bs)" using lp by (induct p rule: \.induct, auto simp add: isint_sub isint_neg) lemma \_\_l: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "\ (b,k) \ set (\_\ p). k >0 \ numbound0 b \ isint b (real_of_int i#bs)" using lp isint_add [OF isint_c[where j="- 1"],where bs="real_of_int i#bs"] by (induct p rule: \_\.induct, auto) lemma \: assumes lp: "iszlfm p (real_of_int (i::int) #bs)" and pi: "Ifm (real_of_int i#bs) p" and d: "d_\ p d" and dp: "d > 0" and nob: "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. real_of_int (c*i) \ Inum (real_of_int i#bs) e + real_of_int j" (is "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. _ \ ?N i e + _") shows "Ifm (real_of_int(i - d)#bs) p" using lp pi d nob proof(induct p rule: iszlfm.induct) case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" and pi: "real_of_int (c*i) = - 1 - ?N i e + real_of_int (1::int)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ -1 - ?N i e + real_of_int j" by simp+ from mult_strict_left_mono[OF dp cp] have one:"1 \ {1 .. c*d}" by auto from nob[rule_format, where j="1", OF one] pi show ?case by simp next case (4 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - ?N i e + real_of_int j" by simp+ {assume "real_of_int (c*i) \ - ?N i e + real_of_int (c*d)" with numbound0_I[OF nb, where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] have ?case by (simp add: algebra_simps)} moreover {assume pi: "real_of_int (c*i) = - ?N i e + real_of_int (c*d)" from mult_strict_left_mono[OF dp cp] have d: "(c*d) \ {1 .. c*d}" by simp from nob[rule_format, where j="c*d", OF d] pi have ?case by simp } ultimately show ?case by blast next case (5 c e) hence cp: "c > 0" by simp from 5 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] of_int_mult] show ?case using 5 dp apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] algebra_simps del: mult_pos_pos) by (metis add.right_neutral of_int_0_less_iff of_int_mult pos_add_strict) next case (6 c e) hence cp: "c > 0" by simp from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] of_int_mult] show ?case using 6 dp - apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] - algebra_simps del: mult_pos_pos) - using order_trans by fastforce + by (fastforce simp: numbound0_I[where bs="bs" and b="i - real_of_int d" and b'="i"] + algebra_simps intro: order_trans) next case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - ?N i e + real_of_int j" and pi: "real_of_int (c*i) + ?N i e > 0" and cp': "real_of_int c >0" by simp+ let ?fe = "\?N i e\" from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c > 0" by (simp add: algebra_simps) from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) > real_of_int (0::int)" by simp hence pi': "c*i + ?fe > 0" by (simp only: of_int_less_iff[symmetric]) have "real_of_int (c*i) + ?N i e > real_of_int (c*d) \ real_of_int (c*i) + ?N i e \ real_of_int (c*d)" by auto moreover {assume "real_of_int (c*i) + ?N i e > real_of_int (c*d)" hence ?case by (simp add: algebra_simps numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} moreover {assume H:"real_of_int (c*i) + ?N i e \ real_of_int (c*d)" with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \ real_of_int (c*d)" by simp hence pid: "c*i + ?fe \ c*d" by (simp only: of_int_le_iff) with pi' have "\ j1\ {1 .. c*d}. c*i + ?fe = j1" by auto hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = - ?N i e + real_of_int j1" unfolding Bex_def using ei[simplified isint_iff] by fastforce with nob have ?case by blast } ultimately show ?case by blast next case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - 1 - ?N i e + real_of_int j" and pi: "real_of_int (c*i) + ?N i e \ 0" and cp': "real_of_int c >0" by simp+ let ?fe = "\?N i e\" from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c \ 0" by (simp add: algebra_simps) from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \ real_of_int (0::int)" by simp hence pi': "c*i + 1 + ?fe \ 1" by (simp only: of_int_le_iff[symmetric]) have "real_of_int (c*i) + ?N i e \ real_of_int (c*d) \ real_of_int (c*i) + ?N i e < real_of_int (c*d)" by auto moreover {assume "real_of_int (c*i) + ?N i e \ real_of_int (c*d)" hence ?case by (simp add: algebra_simps numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} moreover {assume H:"real_of_int (c*i) + ?N i e < real_of_int (c*d)" with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) < real_of_int (c*d)" by simp hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: of_int_le_iff) with pi' have "\ j1\ {1 .. c*d}. c*i + 1+ ?fe = j1" by auto hence "\ j1\ {1 .. c*d}. real_of_int (c*i) + 1= - ?N i e + real_of_int j1" unfolding Bex_def using ei[simplified isint_iff] by fastforce hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = (- ?N i e + real_of_int j1) - 1" by (simp only: algebra_simps) hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = - 1 - ?N i e + real_of_int j1" by (simp add: algebra_simps) with nob have ?case by blast } ultimately show ?case by blast next case (9 j c e) hence p: "real_of_int j rdvd real_of_int (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ let ?e = "Inum (real_of_int i # bs) e" from 9 have "isint e (real_of_int i #bs)" by simp hence ie: "real_of_int \?e\ = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp from ie[symmetric] have "?p i = (real_of_int j rdvd real_of_int (c*i + \?e\))" by simp also have "\ = (j dvd c*i + \?e\)" using int_rdvd_iff [where i="j" and t="c*i + \?e\"] by simp also have "\ = (j dvd c*i - c*d + \?e\)" using dvd_period[OF id, where x="c*i" and c="-c" and t="\?e\"] by simp also have "\ = (real_of_int j rdvd real_of_int (c*i - c*d + \?e\))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + \?e\)",symmetric, simplified] ie by simp also have "\ = (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) finally show ?case using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) next case (10 j c e) hence p: "\ (real_of_int j rdvd real_of_int (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ let ?e = "Inum (real_of_int i # bs) e" from 10 have "isint e (real_of_int i #bs)" by simp hence ie: "real_of_int \?e\ = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp from ie[symmetric] have "?p i = (\ (real_of_int j rdvd real_of_int (c*i + \?e\)))" by simp also have "\ \ \ (j dvd c*i + \?e\)" using int_rdvd_iff [where i="j" and t="c*i + \?e\"] by simp also have "\ \ \ (j dvd c*i - c*d + \?e\)" using dvd_period[OF id, where x="c*i" and c="-c" and t="\?e\"] by simp also have "\ \ \ (real_of_int j rdvd real_of_int (c*i - c*d + \?e\))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + \?e\)",symmetric, simplified] ie by simp also have "\ \ \ (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) finally show ?case using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]) lemma \_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" shows "bound0 (\ p k t)" using \_\_nb[OF lp nb] nb by (simp add: \_def) lemma \': assumes lp: "iszlfm p (a #bs)" and d: "d_\ p d" and dp: "d > 0" shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") proof(clarify) fix x assume nob1:"?b x" and px: "?P x" from iszlfm_gen[OF lp, rule_format, where y="real_of_int x"] have lp': "iszlfm p (real_of_int x#bs)". have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real_of_int (c * x) \ Inum (real_of_int x # bs) e + real_of_int j" proof(clarify) fix e c j assume ecR: "(e,c) \ set (\ p)" and jD: "j\ {1 .. c*d}" and cx: "real_of_int (c*x) = Inum (real_of_int x#bs) e + real_of_int j" let ?e = "Inum (real_of_int x#bs) e" from \_l[OF lp'] ecR have ei:"isint e (real_of_int x#bs)" and cp:"c>0" and nb:"numbound0 e" by auto from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" . from cx ei[simplified isint_iff] have "real_of_int (c*x) = real_of_int (\?e\ + j)" by simp hence cx: "c*x = \?e\ + j" by (simp only: of_int_eq_iff) hence cdej:"c dvd \?e\ + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp) hence "real_of_int c rdvd real_of_int (\?e\ + j)" by (simp only: int_rdvd_iff) hence rcdej: "real_of_int c rdvd ?e + real_of_int j" by (simp add: ei[simplified isint_iff]) from cx have "(c*x) div c = (\?e\ + j) div c" by simp with cp have "x = (\?e\ + j) div c" by simp with px have th: "?P ((\?e\ + j) div c)" by auto from cp have cp': "real_of_int c > 0" by simp from cdej have cdej': "c dvd \Inum (real_of_int x#bs) (Add e (C j))\" by simp from nb have nb': "numbound0 (Add e (C j))" by simp have ji: "isint (C j) (real_of_int x#bs)" by (simp add: isint_def) from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real_of_int x#bs)" . from th \_\[where b'="real_of_int x", OF lp' cp' nb' ei' cdej',symmetric] have "Ifm (real_of_int x#bs) (\_\ p (Add e (C j), c))" by simp with rcdej have th: "Ifm (real_of_int x#bs) (\ p c (Add e (C j)))" by (simp add: \_def) from th bound0_I[OF \_nb[OF lp nb', where k="c"],where bs="bs" and b="real_of_int x" and b'="a"] have "Ifm (a#bs) (\ p c (Add e (C j)))" by blast with ecR jD nob1 show "False" by blast qed from \[OF lp' px d dp nob] show "?P (x -d )" . qed lemma rl_thm: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" is "?lhs = (?MD \ ?RD)" is "?lhs = ?rhs") proof- let ?d= "\ p" from \[OF lp] have d:"d_\ p ?d" and dp: "?d > 0" by auto { assume H:"?MD" hence th:"\ (x::int). ?MP x" by blast from H minusinf_ex[OF lp th] have ?thesis by blast} moreover { fix e c j assume exR:"(e,c) \ ?R" and jD:"j\ {1 .. c*?d}" and spx:"?SP c e j" from exR \_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real_of_int i#bs)" and cp: "c > 0" by auto have "isint (C j) (real_of_int i#bs)" by (simp add: isint_iff) with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real_of_int i"]] have eji:"isint (Add e (C j)) (real_of_int i#bs)" by simp from nb have nb': "numbound0 (Add e (C j))" by simp from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real_of_int i"] have spx': "Ifm (real_of_int i # bs) (\ p c (Add e (C j)))" by blast from spx' have rcdej:"real_of_int c rdvd (Inum (real_of_int i#bs) (Add e (C j)))" and sr:"Ifm (real_of_int i#bs) (\_\ p (Add e (C j),c))" by (simp add: \_def)+ from rcdej eji[simplified isint_iff] have "real_of_int c rdvd real_of_int \Inum (real_of_int i#bs) (Add e (C j))\" by simp hence cdej:"c dvd \Inum (real_of_int i#bs) (Add e (C j))\" by (simp only: int_rdvd_iff) from cp have cp': "real_of_int c > 0" by simp from \_\[OF lp cp' nb' eji cdej] spx' have "?P (\Inum (real_of_int i # bs) (Add e (C j))\ div c)" by (simp add: \_def) hence ?lhs by blast with exR jD spx have ?thesis by blast} moreover { fix x assume px: "?P x" and nob: "\ ?RD" from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" . from \'[OF lp' d dp, rule_format, OF nob] have th:"\ x. ?P x \ ?P (x - ?d)" by blast from minusinf_inf[OF lp] obtain z where z:"\ xx - z\ + 1 \ 0" by arith from decr_lemma[OF dp,where x="x" and z="z"] decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\ x. ?MP x" by auto with minusinf_bex[OF lp] px nob have ?thesis by blast} ultimately show ?thesis by blast qed lemma mirror_\_\: assumes lp: "iszlfm p (a#bs)" shows "(\ (t,k). (Inum (a#bs) t, k)) ` set (\_\ p) = (\ (t,k). (Inum (a#bs) t,k)) ` set (\ (mirror p))" using lp by (induct p rule: mirror.induct) (simp_all add: split_def image_Un) text \The \\\ part\ text\Linearity for fm where Bound 0 ranges over \\\\ fun isrlfm :: "fm \ bool" (* Linearity test for fm *) where "isrlfm (And p q) = (isrlfm p \ isrlfm q)" | "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" | "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm (Le (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm (Gt (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" | "isrlfm p = (isatom p \ (bound0 p))" definition fp :: "fm \ int \ num \ int \ fm" where "fp p n s j \ (if n > 0 then (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) else (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))" (* splits the bounded from the unbounded part*) fun rsplit0 :: "num \ (fm \ int \ num) list" where "rsplit0 (Bound 0) = [(T,1,C 0)]" | "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b in map (\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\acs,b\bcs])" | "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" | "rsplit0 (Neg a) = map (\ (p,n,s). (p,-n,Neg s)) (rsplit0 a)" | "rsplit0 (Floor a) = concat (map (\ (p,n,s). if n=0 then [(p,0,Floor s)] else (map (\ j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then [0 .. n] else [n .. 0]))) (rsplit0 a))" | "rsplit0 (CN 0 c a) = map (\ (p,n,s). (p,n+c,s)) (rsplit0 a)" | "rsplit0 (CN m c a) = map (\ (p,n,s). (p,n,CN m c s)) (rsplit0 a)" | "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)" | "rsplit0 (Mul c a) = map (\ (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)" | "rsplit0 t = [(T,0,t)]" lemma conj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (conj p q)" using conj_def by (cases p, auto) lemma disj_rl[simp]: "isrlfm p \ isrlfm q \ isrlfm (disj p q)" using disj_def by (cases p, auto) lemma rsplit0_cs: shows "\ (p,n,s) \ set (rsplit0 t). (Ifm (x#bs) p \ (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \ numbound0 s \ isrlfm p" (is "\ (p,n,s) \ ?SS t. (?I p \ ?N t = ?N (CN 0 n s)) \ _ \ _ ") proof(induct t rule: rsplit0.induct) case (5 a) let ?p = "\ (p,n,s) j. fp p n s j" let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))" let ?J = "\ n. if n>0 then [0..n] else [n..0]" let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n])))" proof- fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" thus "(\(a, b, c)\M. set (f (a, b, c))) = (\(a, b, c)\M. set (g a b c))" by (auto simp add: split_def) qed have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) [n..0]" by auto hence U3: "(\ ((\(p,n,s). set (?ff (p,n,s))) ` {(p,n,s). (p,n,s) \ ?SS a \ n<0})) = (\ ((\(p,n,s). set (map (?f(p,n,s)) [n..0])) ` {(p,n,s). (p,n,s)\ ?SS a\n<0}))" proof - fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" thus "(\(a, b, c)\M. set (f (a, b, c))) = (\(a, b, c)\M. set (g a b c))" by (auto simp add: split_def) qed have "?SS (Floor a) = \ ((\x. set (?ff x)) ` ?SS a)" by auto also have "\ = \ ((\ (p,n,s). set (?ff (p,n,s))) ` ?SS a)" by blast also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" by (auto split: if_splits) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un (UNION {(p,n,s). (p,n,s)\ ?SS a\n>0} (\(p,n,s). set(map(?f(p,n,s)) [0..n]))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" by (simp only: set_map set_upto list.set) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" - by blast + by force finally have FS: "?SS (Floor a) = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast show ?case proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -) fix p n s let ?ths = "(?I p \ (?N (Floor a) = ?N (CN 0 n s))) \ numbound0 s \ isrlfm p" assume "(\ba. (p, 0, ba) \ set (rsplit0 a) \ n = 0 \ s = Floor ba) \ (\ab ac ba. (ab, ac, ba) \ set (rsplit0 a) \ 0 < ac \ (\j. p = fp ab ac ba j \ n = 0 \ s = Add (Floor ba) (C j) \ 0 \ j \ j \ ac)) \ (\ab ac ba. (ab, ac, ba) \ set (rsplit0 a) \ ac < 0 \ (\j. p = fp ab ac ba j \ n = 0 \ s = Add (Floor ba) (C j) \ ac \ j \ j \ 0))" moreover { fix s' assume "(p, 0, s') \ ?SS a" and "n = 0" and "s = Floor s'" hence ?ths using 5(1) by auto } moreover { fix p' n' s' j assume pns: "(p', n', s') \ ?SS a" and np: "0 < n'" and p_def: "p = ?p (p',n',s') j" and n0: "n = 0" and s_def: "s = (Add (Floor s') (C j))" and jp: "0 \ j" and jn: "j \ n'" from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ numbound0 s' \ isrlfm p'" by blast hence nb: "numbound0 s'" by simp from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp let ?nxs = "CN 0 n' s'" let ?l = "\?N s'\ + j" from H have "?I (?p (p',n',s') j) \ (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: fp_def np algebra_simps) also have "\ \ \?N ?nxs\ = ?l \ ?N a = ?N ?nxs" using floor_eq_iff[where x="?N ?nxs" and a="?l"] by simp moreover have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by blast with s_def n0 p_def nb nf have ?ths by auto} moreover { fix p' n' s' j assume pns: "(p', n', s') \ ?SS a" and np: "n' < 0" and p_def: "p = ?p (p',n',s') j" and n0: "n = 0" and s_def: "s = (Add (Floor s') (C j))" and jp: "n' \ j" and jn: "j \ 0" from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ numbound0 s' \ isrlfm p'" by blast hence nb: "numbound0 s'" by simp from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp let ?nxs = "CN 0 n' s'" let ?l = "\?N s'\ + j" from H have "?I (?p (p',n',s') j) \ (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: np fp_def algebra_simps) also have "\ \ \?N ?nxs\ = ?l \ ?N a = ?N ?nxs" using floor_eq_iff[where x="?N ?nxs" and a="?l"] by simp moreover have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by blast with s_def n0 p_def nb nf have ?ths by auto} ultimately show ?ths by fastforce qed next case (3 a b) then show ?case by auto qed (auto simp add: Let_def split_def algebra_simps) lemma real_in_int_intervals: assumes xb: "real_of_int m \ x \ x < real_of_int ((n::int) + 1)" shows "\ j\ {m.. n}. real_of_int j \ x \ x < real_of_int (j+1)" (is "\ j\ ?N. ?P j") by (rule bexI[where P="?P" and x="\x\" and A="?N"]) (auto simp add: floor_less_iff[where x="x" and z="n+1", simplified] xb[simplified] floor_mono[where x="real_of_int m" and y="x", OF conjunct1[OF xb], simplified floor_of_int[where z="m"]]) lemma rsplit0_complete: assumes xp:"0 \ x" and x1:"x < 1" shows "\ (p,n,s) \ set (rsplit0 t). Ifm (x#bs) p" (is "\ (p,n,s) \ ?SS t. ?I p") proof(induct t rule: rsplit0.induct) case (2 a b) then have "\ (pa,na,sa) \ ?SS a. ?I pa" by auto then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\ ?SS a \ ?I pa" by blast with 2 have "\ (pb,nb,sb) \ ?SS b. ?I pb" by blast then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\ ?SS b \ ?I pb" by blast from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \ set[(x,y). x\rsplit0 a, y\rsplit0 b]" by (auto) let ?f="(\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))" from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \ ?SS (Add a b)" by (simp add: Let_def) hence "(And pa pb, na +nb, Add sa sb) \ ?SS (Add a b)" by simp moreover from pa pb have "?I (And pa pb)" by simp ultimately show ?case by blast next case (5 a) let ?p = "\ (p,n,s) j. fp p n s j" let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))" let ?J = "\ n. if n>0 then [0..n] else [n..0]" let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n])))" proof- fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" thus "(\(a, b, c)\M. set (f (a, b, c))) = (\(a, b, c)\M. set (g a b c))" by (auto simp add: split_def) qed have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) [n..0]" by auto hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) [n..0])))" proof- fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g assume "\ (a,b,c) \ M. f (a,b,c) = g a b c" thus "(\(a, b, c)\M. set (f (a, b, c))) = (\(a, b, c)\M. set (g a b c))" by (auto simp add: split_def) qed have "?SS (Floor a) = \ ((\x. set (?ff x)) ` ?SS a)" by auto also have "\ = \ ((\ (p,n,s). set (?ff (p,n,s))) ` ?SS a)" by blast also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" by (auto split: if_splits) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n]))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" by (simp only: set_map set_upto list.set) also have "\ = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" - by blast + by force finally have FS: "?SS (Floor a) = ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast from 5 have "\ (p,n,s) \ ?SS a. ?I p" by auto then obtain "p" "n" "s" where pns: "(p,n,s) \ ?SS a \ ?I p" by blast let ?N = "\ t. Inum (x#bs) t" from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \ numbound0 s \ isrlfm p" by auto have "n=0 \ n >0 \ n <0" by arith moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto } moreover { assume np: "n > 0" from of_int_floor_le[of "?N s"] have "?N (Floor s) \ ?N s" by simp also from mult_left_mono[OF xp] np have "?N s \ real_of_int n * x + ?N s" by simp finally have "?N (Floor s) \ real_of_int n * x + ?N s" . moreover {from x1 np have "real_of_int n *x + ?N s < real_of_int n + ?N s" by simp also from real_of_int_floor_add_one_gt[where r="?N s"] have "\ < real_of_int n + ?N (Floor s) + 1" by simp finally have "real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp} ultimately have "?N (Floor s) \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp hence th: "0 \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (n+1)" by simp from real_in_int_intervals th have "\ j\ {0 .. n}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp hence "\ j\ {0 .. n}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) hence "\ j\ {0.. n}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def np algebra_simps) then obtain "j" where j_def: "j\ {0 .. n} \ ?I (?p (p,n,s) j)" by blast hence "\x \ {?p (p,n,s) j |j. 0\ j \ j \ n }. ?I x" by auto hence ?case using pns by (simp only: FS,simp add: bex_Un) (rule disjI2, rule disjI1,rule exI [where x="p"], rule exI [where x="n"],rule exI [where x="s"],simp_all add: np) } moreover { assume nn: "n < 0" hence np: "-n >0" by simp from of_int_floor_le[of "?N s"] have "?N (Floor s) + 1 > ?N s" by simp moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real_of_int n * x + ?N s" by simp ultimately have "?N (Floor s) + 1 > real_of_int n * x + ?N s" by arith moreover {from x1 nn have "real_of_int n *x + ?N s \ real_of_int n + ?N s" by simp moreover from of_int_floor_le[of "?N s"] have "real_of_int n + ?N s \ real_of_int n + ?N (Floor s)" by simp ultimately have "real_of_int n *x + ?N s \ ?N (Floor s) + real_of_int n" by (simp only: algebra_simps)} ultimately have "?N (Floor s) + real_of_int n \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (1::int)" by simp hence th: "real_of_int n \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (1::int)" by simp have th1: "\ (a::real). (- a > 0) = (a < 0)" by auto have th2: "\ (a::real). (0 \ - a) = (a \ 0)" by auto from real_in_int_intervals th have "\ j\ {n .. 0}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp hence "\ j\ {n .. 0}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) hence "\ j\ {n .. 0}. 0 \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j) \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) hence "\ j\ {n.. 0}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def nn algebra_simps del: diff_less_0_iff_less diff_le_0_iff_le) then obtain "j" where j_def: "j\ {n .. 0} \ ?I (?p (p,n,s) j)" by blast hence "\x \ {?p (p,n,s) j |j. n\ j \ j \ 0 }. ?I x" by auto hence ?case using pns by (simp only: FS,simp add: bex_Un) (rule disjI2, rule disjI2,rule exI [where x="p"], rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn) } ultimately show ?case by blast qed (auto simp add: Let_def split_def) (* Linearize a formula where Bound 0 ranges over [0,1) *) definition rsplit :: "(int \ num \ fm) \ num \ fm" where "rsplit f a \ foldr disj (map (\ (\, n, s). conj \ (f n s)) (rsplit0 a)) F" lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\ x \ set xs. Ifm bs (f x))" by(induct xs, simp_all) lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\ x \ set xs. Ifm bs (f x))" by(induct xs, simp_all) lemma foldr_disj_map_rlfm: assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" and \: "\ (\,n,s) \ set xs. numbound0 s \ isrlfm \" shows "isrlfm (foldr disj (map (\ (\, n, s). conj \ (f n s)) xs) F)" using lf \ by (induct xs, auto) lemma rsplit_ex: "Ifm bs (rsplit f a) = (\ (\,n,s) \ set (rsplit0 a). Ifm bs (conj \ (f n s)))" using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def) lemma rsplit_l: assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" shows "isrlfm (rsplit f a)" proof- from rsplit0_cs[where t="a"] have th: "\ (\,n,s) \ set (rsplit0 a). numbound0 s \ isrlfm \" by blast from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp qed lemma rsplit: assumes xp: "x \ 0" and x1: "x < 1" and f: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))" shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)" proof(auto) let ?I = "\x p. Ifm (x#bs) p" let ?N = "\ x t. Inum (x#bs) t" assume "?I x (rsplit f a)" hence "\ (\,n,s) \ set (rsplit0 a). ?I x (And \ (f n s))" using rsplit_ex by simp then obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and "?I x (And \ (f n s))" by blast hence \: "?I x \" and fns: "?I x (f n s)" by auto from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \ have th: "(?N x a = ?N x (CN 0 n s)) \ numbound0 s" by auto from f[rule_format, OF th] fns show "?I x (g a)" by simp next let ?I = "\x p. Ifm (x#bs) p" let ?N = "\ x t. Inum (x#bs) t" assume ga: "?I x (g a)" from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and fx: "?I x \" by blast from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto with ga f have "?I x (f n s)" by auto with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto qed definition lt :: "int \ num \ fm" where lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) else (Gt (CN 0 (-c) (Neg t))))" definition le :: "int \ num \ fm" where le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) else (Ge (CN 0 (-c) (Neg t))))" definition gt :: "int \ num \ fm" where gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) else (Lt (CN 0 (-c) (Neg t))))" definition ge :: "int \ num \ fm" where ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) else (Le (CN 0 (-c) (Neg t))))" definition eq :: "int \ num \ fm" where eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) else (Eq (CN 0 (-c) (Neg t))))" definition neq :: "int \ num \ fm" where neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) else (NEq (CN 0 (-c) (Neg t))))" lemma lt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)" (is "\ a n s . ?N a = ?N (CN 0 n s) \ _\ ?I (lt n s) = ?I (Lt a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def)) (cases "n > 0", simp_all add: lt_def algebra_simps myless[of _ "0"]) qed lemma lt_l: "isrlfm (rsplit lt a)" by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def, case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma le_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (le n s) = Ifm (x#bs) (Le a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (le n s) = ?I (Le a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def)) (cases "n > 0", simp_all add: le_def algebra_simps myle[of _ "0"]) qed lemma le_l: "isrlfm (rsplit le a)" by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat",simp_all) lemma gt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (gt n s) = ?I (Gt a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def)) (cases "n > 0", simp_all add: gt_def algebra_simps myless[of _ "0"]) qed lemma gt_l: "isrlfm (rsplit gt a)" by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma ge_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\ a n s . ?N a = ?N (CN 0 n s) \ _ \ ?I (ge n s) = ?I (Ge a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def)) (cases "n > 0", simp_all add: ge_def algebra_simps myle[of _ "0"]) qed lemma ge_l: "isrlfm (rsplit ge a)" by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma eq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (eq n s) = ?I (Eq a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps) qed lemma eq_l: "isrlfm (rsplit eq a)" by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) (case_tac s, simp_all, rename_tac nat a b, case_tac"nat", simp_all) lemma neq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (neq n s) = ?I (NEq a)") proof(clarify) fix a n s bs assume H: "?N a = ?N (CN 0 n s)" show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def algebra_simps) qed lemma neq_l: "isrlfm (rsplit neq a)" by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) (case_tac s, simp_all, rename_tac nat a b, case_tac"nat", simp_all) lemma small_le: assumes u0:"0 \ u" and u1: "u < 1" shows "(-u \ real_of_int (n::int)) = (0 \ n)" using u0 u1 by auto lemma small_lt: assumes u0:"0 \ u" and u1: "u < 1" shows "(real_of_int (n::int) < real_of_int (m::int) - u) = (n < m)" using u0 u1 by auto lemma rdvd01_cs: assumes up: "u \ 0" and u1: "u<1" and np: "real_of_int n > 0" shows "(real_of_int (i::int) rdvd real_of_int (n::int) * u - s) = (\ j\ {0 .. n - 1}. real_of_int n * u = s - real_of_int \s\ + real_of_int j \ real_of_int i rdvd real_of_int (j - \s\))" (is "?lhs = ?rhs") proof- let ?ss = "s - real_of_int \s\" from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" by (auto simp: floor_less_cancel) from np have n0: "real_of_int n \ 0" by simp from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] have nu0:"real_of_int n * u - s \ -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto from int_rdvd_real[where i="i" and x="real_of_int (n::int) * u - s"] have "real_of_int i rdvd real_of_int n * u - s = (i dvd \real_of_int n * u - s\ \ (real_of_int \real_of_int n * u - s\ = real_of_int n * u - s ))" (is "_ = (?DE)" is "_ = (?D \ ?E)") by simp also have "\ = (?DE \ real_of_int (\real_of_int n * u - s\ + \s\) \ -?ss \ real_of_int (\real_of_int n * u - s\ + \s\) < real_of_int n - ?ss)" (is "_=(?DE \real_of_int ?a \ _ \ real_of_int ?a < _)") using nu0 nun by auto also have "\ = (?DE \ ?a \ 0 \ ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. ?a = j))" by simp also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. real_of_int (\real_of_int n * u - s\) = real_of_int j - real_of_int \s\ ))" by (simp only: algebra_simps of_int_diff[symmetric] of_int_eq_iff) also have "\ = ((\ j\ {0 .. (n - 1)}. real_of_int n * u - s = real_of_int j - real_of_int \s\ \ real_of_int i rdvd real_of_int n * u - s))" using int_rdvd_iff[where i="i" and t="\real_of_int n * u - s\"] by (auto cong: conj_cong) also have "\ = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps ) finally show ?thesis . qed definition DVDJ:: "int \ int \ num \ fm" where DVDJ_def: "DVDJ i n s = (foldr disj (map (\ j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) F)" definition NDVDJ:: "int \ int \ num \ fm" where NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\ j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) T)" lemma DVDJ_DVD: assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" proof- let ?f = "\ j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" let ?s= "Inum (x#bs) s" from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np DVDJ_def) also have "\ = (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int \- ?s\ + real_of_int j \ real_of_int i rdvd real_of_int (j - \- ?s\))" by (simp add: algebra_simps) also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] have "\ = (real_of_int i rdvd real_of_int n * x - (-?s))" by simp finally show ?thesis by simp qed lemma NDVDJ_NDVD: assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" proof- let ?f = "\ j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))" let ?s= "Inum (x#bs) s" from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np NDVDJ_def) also have "\ = (\ (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int \- ?s\ + real_of_int j \ real_of_int i rdvd real_of_int (j - \- ?s\)))" by (simp add: algebra_simps) also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] have "\ = (\ (real_of_int i rdvd real_of_int n * x - (-?s)))" by simp finally show ?thesis by simp qed lemma foldr_disj_map_rlfm2: assumes lf: "\ n . isrlfm (f n)" shows "isrlfm (foldr disj (map f xs) F)" using lf by (induct xs, auto) lemma foldr_And_map_rlfm2: assumes lf: "\ n . isrlfm (f n)" shows "isrlfm (foldr conj (map f xs) T)" using lf by (induct xs, auto) lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" shows "isrlfm (DVDJ i n s)" proof- let ?f="\j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" have th: "\ j. isrlfm (?f j)" using nb np by auto from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp qed lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" shows "isrlfm (NDVDJ i n s)" proof- let ?f="\j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))" have th: "\ j. isrlfm (?f j)" using nb np by auto from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto qed definition DVD :: "int \ int \ num \ fm" where DVD_def: "DVD i c t = (if i=0 then eq c t else if c = 0 then (Dvd i t) else if c >0 then DVDJ \i\ c t else DVDJ \i\ (-c) (Neg t))" definition NDVD :: "int \ int \ num \ fm" where "NDVD i c t = (if i=0 then neq c t else if c = 0 then (NDvd i t) else if c >0 then NDVDJ \i\ c t else NDVDJ \i\ (-c) (Neg t))" lemma DVD_mono: assumes xp: "0\ x" and x1: "x < 1" shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (DVD i n s) = ?I (Dvd i a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" let ?th = "?I (DVD i n s) = ?I (Dvd i a)" have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] by (simp add: DVD_def rdvd_left_0_eq)} moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } moreover {assume inz: "i\0" and "n<0" hence ?th by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast qed lemma NDVD_mono: assumes xp: "0\ x" and x1: "x < 1" shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (NDVD i n s) = ?I (NDvd i a)") proof(clarify) fix a n s assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" let ?th = "?I (NDVD i n s) = ?I (NDvd i a)" have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] by (simp add: NDVD_def rdvd_left_0_eq)} moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } moreover {assume inz: "i\0" and "n<0" hence ?th by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast qed lemma DVD_l: "isrlfm (rsplit (DVD i) a)" by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)" by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) fun rlfm :: "fm \ fm" where "rlfm (And p q) = conj (rlfm p) (rlfm q)" | "rlfm (Or p q) = disj (rlfm p) (rlfm q)" | "rlfm (Imp p q) = disj (rlfm (Not p)) (rlfm q)" | "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (Not p)) (rlfm (Not q)))" | "rlfm (Lt a) = rsplit lt a" | "rlfm (Le a) = rsplit le a" | "rlfm (Gt a) = rsplit gt a" | "rlfm (Ge a) = rsplit ge a" | "rlfm (Eq a) = rsplit eq a" | "rlfm (NEq a) = rsplit neq a" | "rlfm (Dvd i a) = rsplit (\ t. DVD i t) a" | "rlfm (NDvd i a) = rsplit (\ t. NDVD i t) a" | "rlfm (Not (And p q)) = disj (rlfm (Not p)) (rlfm (Not q))" | "rlfm (Not (Or p q)) = conj (rlfm (Not p)) (rlfm (Not q))" | "rlfm (Not (Imp p q)) = conj (rlfm p) (rlfm (Not q))" | "rlfm (Not (Iff p q)) = disj (conj(rlfm p) (rlfm(Not q))) (conj(rlfm(Not p)) (rlfm q))" | "rlfm (Not (Not p)) = rlfm p" | "rlfm (Not T) = F" | "rlfm (Not F) = T" | "rlfm (Not (Lt a)) = simpfm (rlfm (Ge a))" | "rlfm (Not (Le a)) = simpfm (rlfm (Gt a))" | "rlfm (Not (Gt a)) = simpfm (rlfm (Le a))" | "rlfm (Not (Ge a)) = simpfm (rlfm (Lt a))" | "rlfm (Not (Eq a)) = simpfm (rlfm (NEq a))" | "rlfm (Not (NEq a)) = simpfm (rlfm (Eq a))" | "rlfm (Not (Dvd i a)) = simpfm (rlfm (NDvd i a))" | "rlfm (Not (NDvd i a)) = simpfm (rlfm (Dvd i a))" | "rlfm p = p" lemma bound0at_l : "\isatom p ; bound0 p\ \ isrlfm p" by (induct p rule: isrlfm.induct, auto) lemma simpfm_rl: "isrlfm p \ isrlfm (simpfm p)" proof (induct p) case (Lt a) hence "bound0 (Lt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" using simpfm_bound0 by blast have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with Lt a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (Le a) hence "bound0 (Le a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover { assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" using simpfm_bound0 by blast have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with Le a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (Gt a) hence "bound0 (Gt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a, simp_all, rename_tac nat a b,case_tac "nat", simp_all) moreover {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" using simpfm_bound0 by blast have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1: "numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with Gt a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (Ge a) hence "bound0 (Ge a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover { assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" using simpfm_bound0 by blast have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with Ge a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (Eq a) hence "bound0 (Eq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover { assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" using simpfm_bound0 by blast have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with Eq a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (NEq a) hence "bound0 (NEq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" using simpfm_bound0 by blast have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']] have "0 < c div numgcd (CN 0 c (simpnum e))" by simp } with NEq a have ?case by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" using simpfm_bound0 by blast have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) with bn bound0at_l show ?case by blast next case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" using simpfm_bound0 by blast have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) with bn bound0at_l show ?case by blast qed(auto simp add: conj_def imp_def disj_def iff_def Let_def) lemma rlfm_I: assumes qfp: "qfree p" and xp: "0 \ x" and x1: "x < 1" shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \ isrlfm (rlfm p)" using qfp by (induct p rule: rlfm.induct) (auto simp add: rsplit[OF xp x1 lt_mono] lt_l rsplit[OF xp x1 le_mono] le_l rsplit[OF xp x1 gt_mono] gt_l rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl) lemma rlfm_l: assumes qfp: "qfree p" shows "isrlfm (rlfm p)" using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l by (induct p rule: rlfm.induct) (auto simp add: simpfm_rl) (* Operations needed for Ferrante and Rackoff *) lemma rminusinf_inf: assumes lp: "isrlfm p" shows "\ z. \ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") using lp proof (induct p rule: minusinf.induct) case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto next case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto next case (3 c e) from 3 have nb: "numbound0 e" by simp from 3 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp thus ?case by blast next case (4 c e) from 4 have nb: "numbound0 e" by simp from 4 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp thus ?case by blast next case (5 c e) from 5 have nb: "numbound0 e" by simp from 5 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp thus ?case by blast next case (6 c e) from 6 have nb: "numbound0 e" by simp from 6 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp thus ?case by blast next case (7 c e) from 7 have nb: "numbound0 e" by simp from 7 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp thus ?case by blast next case (8 c e) from 8 have nb: "numbound0 e" by simp from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp thus ?case by blast qed simp_all lemma rplusinf_inf: assumes lp: "isrlfm p" shows "\ z. \ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\ z. \ x. ?P z x p") using lp proof (induct p rule: isrlfm.induct) case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto next case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto next case (3 c e) from 3 have nb: "numbound0 e" by simp from 3 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp thus ?case by blast next case (4 c e) from 4 have nb: "numbound0 e" by simp from 4 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp thus ?case by blast next case (5 c e) from 5 have nb: "numbound0 e" by simp from 5 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp thus ?case by blast next case (6 c e) from 6 have nb: "numbound0 e" by simp from 6 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp thus ?case by blast next case (7 c e) from 7 have nb: "numbound0 e" by simp from 7 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp thus ?case by blast next case (8 c e) from 8 have nb: "numbound0 e" by simp from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp thus ?case by blast qed simp_all lemma rminusinf_bound0: assumes lp: "isrlfm p" shows "bound0 (minusinf p)" using lp by (induct p rule: minusinf.induct) simp_all lemma rplusinf_bound0: assumes lp: "isrlfm p" shows "bound0 (plusinf p)" using lp by (induct p rule: plusinf.induct) simp_all lemma rminusinf_ex: assumes lp: "isrlfm p" and ex: "Ifm (a#bs) (minusinf p)" shows "\ x. Ifm (x#bs) p" proof- from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex have th: "\ x. Ifm (x#bs) (minusinf p)" by auto from rminusinf_inf[OF lp, where bs="bs"] obtain z where z_def: "\x x. Ifm (x#bs) p" proof- from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex have th: "\ x. Ifm (x#bs) (plusinf p)" by auto from rplusinf_inf[OF lp, where bs="bs"] obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp moreover have "z + 1 > z" by simp ultimately show ?thesis using z_def by auto qed fun \:: "fm \ (num \ int) list" where "\ (And p q) = (\ p @ \ q)" | "\ (Or p q) = (\ p @ \ q)" | "\ (Eq (CN 0 c e)) = [(Neg e,c)]" | "\ (NEq (CN 0 c e)) = [(Neg e,c)]" | "\ (Lt (CN 0 c e)) = [(Neg e,c)]" | "\ (Le (CN 0 c e)) = [(Neg e,c)]" | "\ (Gt (CN 0 c e)) = [(Neg e,c)]" | "\ (Ge (CN 0 c e)) = [(Neg e,c)]" | "\ p = []" fun \ :: "fm \ num \ int \ fm" where "\ (And p q) = (\ (t,n). And (\ p (t,n)) (\ q (t,n)))" | "\ (Or p q) = (\ (t,n). Or (\ p (t,n)) (\ q (t,n)))" | "\ (Eq (CN 0 c e)) = (\ (t,n). Eq (Add (Mul c t) (Mul n e)))" | "\ (NEq (CN 0 c e)) = (\ (t,n). NEq (Add (Mul c t) (Mul n e)))" | "\ (Lt (CN 0 c e)) = (\ (t,n). Lt (Add (Mul c t) (Mul n e)))" | "\ (Le (CN 0 c e)) = (\ (t,n). Le (Add (Mul c t) (Mul n e)))" | "\ (Gt (CN 0 c e)) = (\ (t,n). Gt (Add (Mul c t) (Mul n e)))" | "\ (Ge (CN 0 c e)) = (\ (t,n). Ge (Add (Mul c t) (Mul n e)))" | "\ p = (\ (t,n). p)" lemma \_I: assumes lp: "isrlfm p" and np: "real_of_int n > 0" and nbt: "numbound0 t" shows "(Ifm (x#bs) (\ p (t,n)) = Ifm (((Inum (x#bs) t)/(real_of_int n))#bs) p) \ bound0 (\ p (t,n))" (is "(?I x (\ p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") using lp proof(induct p rule: \.induct) case (5 c e) from 5 have cp: "c >0" and nb: "numbound0 e" by simp_all have "?I ?u (Lt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) < 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) < 0)" by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) < 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (6 c e) from 6 have cp: "c >0" and nb: "numbound0 e" by simp_all have "?I ?u (Le (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (7 c e) from 7 have cp: "c >0" and nb: "numbound0 e" by simp_all have "?I ?u (Gt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) > 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) > 0)" by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) > 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (8 c e) from 8 have cp: "c >0" and nb: "numbound0 e" by simp_all have "?I ?u (Ge (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (3 c e) from 3 have cp: "c >0" and nb: "numbound0 e" by simp_all from np have np: "real_of_int n \ 0" by simp have "?I ?u (Eq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) = 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) = 0)" by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) = 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (4 c e) from 4 have cp: "c >0" and nb: "numbound0 e" by simp_all from np have np: "real_of_int n \ 0" by simp have "?I ?u (NEq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified div_0]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int n" and b'="x"]) lemma \_l: assumes lp: "isrlfm p" shows "\ (t,k) \ set (\ p). numbound0 t \ k >0" using lp by(induct p rule: \.induct) auto lemma rminusinf_\: assumes lp: "isrlfm p" and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real_of_int m" (is "\ (s,m) \ ?U p. x \ ?N a s / real_of_int m") proof- have "\ (s,m) \ set (\ p). real_of_int m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed lemma rplusinf_\: assumes lp: "isrlfm p" and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real_of_int m" (is "\ (s,m) \ ?U p. x \ ?N a s / real_of_int m") proof- have "\ (s,m) \ set (\ p). real_of_int m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed lemma lin_dense: assumes lp: "isrlfm p" and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real_of_int n) ` set (\ p)" (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real_of_int n ) ` (?U p)") and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" and ly: "l < y" and yu: "y < u" shows "Ifm (y#bs) p" using lp px noS proof (induct p rule: isrlfm.induct) case (5 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 5 have "x * real_of_int c + ?N x e < 0" by (simp add: algebra_simps) hence pxc: "x < (- ?N x e) / real_of_int c" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) from 5 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto moreover {assume y: "y < (-?N x e)/ real_of_int c" hence "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} moreover {assume y: "y > (- ?N x e) / real_of_int c" with yu have eu: "u > (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (6 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 6 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) from 6 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto moreover {assume y: "y < (-?N x e)/ real_of_int c" hence "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} moreover {assume y: "y > (- ?N x e) / real_of_int c" with yu have eu: "u > (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (7 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 7 have "x * real_of_int c + ?N x e > 0" by (simp add: algebra_simps) hence pxc: "x > (- ?N x e) / real_of_int c" by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) from 7 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto moreover {assume y: "y > (-?N x e)/ real_of_int c" hence "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} moreover {assume y: "y < (- ?N x e) / real_of_int c" with ly have eu: "l < (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (8 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 8 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) from 8 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto moreover {assume y: "y > (-?N x e)/ real_of_int c" hence "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} moreover {assume y: "y < (- ?N x e) / real_of_int c" with ly have eu: "l < (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (3 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from cp have cnz: "real_of_int c \ 0" by simp from 3 have "x * real_of_int c + ?N x e = 0" by (simp add: algebra_simps) hence pxc: "x = (- ?N x e) / real_of_int c" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) from 3 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with lx xu have yne: "x \ - ?N x e / real_of_int c" by auto with pxc show ?case by simp next case (4 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from cp have cnz: "real_of_int c \ 0" by simp from 4 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto hence "y* real_of_int c \ -?N x e" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp hence "y* real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by (simp add: algebra_simps) qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"]) lemma rinf_\: assumes lp: "isrlfm p" and nmi: "\ (Ifm (x#bs) (minusinf p))" (is "\ (Ifm (x#bs) (?M p))") and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") shows "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p" proof- let ?N = "\ x t. Inum (x#bs) t" let ?U = "set (\ p)" from ex obtain a where pa: "?I a p" by blast from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi have nmi': "\ (?I a (?M p))" by simp from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi have npi': "\ (?I a (?P p))" by simp have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((?N a l/real_of_int n + ?N a s /real_of_int m) / 2) p" proof- let ?M = "(\ (t,c). ?N a t / real_of_int c) ` ?U" have fM: "finite ?M" by auto from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real_of_int n \ a \ ?N x s / real_of_int m" by blast then obtain "t" "n" "s" "m" where tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" and xs1: "a \ ?N x s / real_of_int m" and tx1: "a \ ?N x t / real_of_int n" by blast from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real_of_int m" and tx: "a \ ?N a t / real_of_int n" by auto from tnU have Mne: "?M \ {}" by auto hence Une: "?U \ {}" by simp let ?l = "Min ?M" let ?u = "Max ?M" have linM: "?l \ ?M" using fM Mne by simp have uinM: "?u \ ?M" using fM Mne by simp have tnM: "?N a t / real_of_int n \ ?M" using tnU by auto have smM: "?N a s / real_of_int m \ ?M" using smU by auto have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto have "?l \ ?N a t / real_of_int n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp have "?N a s / real_of_int m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] have "(\ s\ ?M. ?I s p) \ (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . moreover { fix u assume um: "u\ ?M" and pu: "?I u p" hence "\ (tu,nu) \ ?U. u = ?N a tu / real_of_int nu" by auto then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real_of_int nu" by blast have "(u + u) / 2 = u" by auto with pu tuu have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) p" by simp with tuU have ?thesis by blast} moreover{ assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" by blast from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real_of_int t1n" by auto then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real_of_int t1n" by blast from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real_of_int t2n" by auto then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real_of_int t2n" by blast from t1x xt2 have t1t2: "t1 < t2" by simp let ?u = "(t1 + t2) / 2" from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . with t1uU t2uU t1u t2u have ?thesis by blast} ultimately show ?thesis by blast qed then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p" by blast from lnU smU \_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) p" by simp with lnU smU show ?thesis by auto qed (* The Ferrante - Rackoff Theorem *) theorem fr_eq: assumes lp: "isrlfm p" shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real_of_int n + (Inum (x#bs) s) / real_of_int m) /2)#bs) p))" (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") proof assume px: "\ x. ?I x p" have "?M \ ?P \ (\ ?M \ \ ?P)" by blast moreover {assume "?M \ ?P" hence "?D" by blast} moreover {assume nmi: "\ ?M" and npi: "\ ?P" from rinf_\[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} ultimately show "?D" by blast next assume "?D" moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } moreover {assume f:"?F" hence "?E" by blast} ultimately show "?E" by blast qed lemma fr_eq_\: assumes lp: "isrlfm p" shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (\ p). \ (s,l) \ set (\ p). Ifm (x#bs) (\ p (Add(Mul l t) (Mul k s) , 2*k*l))))" (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") proof assume px: "\ x. ?I x p" have "?M \ ?P \ (\ ?M \ \ ?P)" by blast moreover {assume "?M \ ?P" hence "?D" by blast} moreover {assume nmi: "\ ?M" and npi: "\ ?P" let ?f ="\ (t,n). Inum (x#bs) t / real_of_int n" let ?N = "\ t. Inum (x#bs) t" {fix t n s m assume "(t,n)\ set (\ p)" and "(s,m) \ set (\ p)" with \_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) /2) p" by (simp only: st[symmetric])} with rinf_\[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} ultimately show "?D" by blast next assume "?D" moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" with \_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int k > 0" and snb: "numbound0 s" and mp:"real_of_int l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" from np mp have mnp: "real_of_int (2*k*l) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp from \_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} ultimately show "?E" by blast qed text\The overall Part\ lemma real_ex_int_real01: shows "(\ (x::real). P x) = (\ (i::int) (u::real). 0\ u \ u< 1 \ P (real_of_int i + u))" proof(auto) fix x assume Px: "P x" let ?i = "\x\" let ?u = "x - real_of_int ?i" have "x = real_of_int ?i + ?u" by simp hence "P (real_of_int ?i + ?u)" using Px by simp moreover have "real_of_int ?i \ x" using of_int_floor_le by simp hence "0 \ ?u" by arith moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real_of_int i + u))" by blast qed fun exsplitnum :: "num \ num" where "exsplitnum (C c) = (C c)" | "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)" | "exsplitnum (Bound n) = Bound (n+1)" | "exsplitnum (Neg a) = Neg (exsplitnum a)" | "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) " | "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) " | "exsplitnum (Mul c a) = Mul c (exsplitnum a)" | "exsplitnum (Floor a) = Floor (exsplitnum a)" | "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))" | "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)" | "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)" fun exsplit :: "fm \ fm" where "exsplit (Lt a) = Lt (exsplitnum a)" | "exsplit (Le a) = Le (exsplitnum a)" | "exsplit (Gt a) = Gt (exsplitnum a)" | "exsplit (Ge a) = Ge (exsplitnum a)" | "exsplit (Eq a) = Eq (exsplitnum a)" | "exsplit (NEq a) = NEq (exsplitnum a)" | "exsplit (Dvd i a) = Dvd i (exsplitnum a)" | "exsplit (NDvd i a) = NDvd i (exsplitnum a)" | "exsplit (And p q) = And (exsplit p) (exsplit q)" | "exsplit (Or p q) = Or (exsplit p) (exsplit q)" | "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)" | "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)" | "exsplit (Not p) = Not (exsplit p)" | "exsplit p = p" lemma exsplitnum: "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t" by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps) lemma exsplit: assumes qfp: "qfree p" shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p" using qfp exsplitnum[where x="x" and y="y" and bs="bs"] by(induct p rule: exsplit.induct) simp_all lemma splitex: assumes qf: "qfree p" shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real_of_int i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") proof- have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real_of_int i)#bs) (exsplit p))" by auto also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real_of_int i + x) #bs) p)" by (simp only: exsplit[OF qf] ac_simps) also have "\ = (\ x. Ifm (x#bs) p)" by (simp only: real_ex_int_real01[where P="\ x. Ifm (x#bs) p"]) finally show ?thesis by simp qed (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) definition ferrack01 :: "fm \ fm" where "ferrack01 p \ (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); U = remdups(map simp_num_pair (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs (\ p')))) in decr (evaldjf (\ p') U ))" lemma fr_eq_01: assumes qf: "qfree p" shows "(\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\ (t,n) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \ (s,m) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) (Add (Mul m t) (Mul n s), 2*n*m)))" (is "(\ x. ?I x ?q) = ?F") proof- let ?rq = "rlfm ?q" let ?M = "?I x (minusinf ?rq)" let ?P = "?I x (plusinf ?rq)" have MF: "?M = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def) by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all) - have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def) + have PF: "?P = False" + apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def) by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all) have "(\ x. ?I x ?q ) = ((?I x (minusinf ?rq)) \ (?I x (plusinf ?rq )) \ (\ (t,n) \ set (\ ?rq). \ (s,m) \ set (\ ?rq ). ?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))))" (is "(\ x. ?I x ?q) = (?M \ ?P \ ?F)" is "?E = ?D") proof assume "\ x. ?I x ?q" then obtain x where qx: "?I x ?q" by blast hence xp: "0\ x" and x1: "x< 1" and px: "?I x p" by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf]) from qx have "?I x ?rq " by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto from qf have qfq:"isrlfm ?rq" by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) with lqx fr_eq_\[OF qfq] show "?M \ ?P \ ?F" by blast next assume D: "?D" let ?U = "set (\ ?rq )" from MF PF D have "?F" by auto then obtain t n s m where aU:"(t,n) \ ?U" and bU:"(s,m)\ ?U" and rqx: "?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] by (auto simp add: rsplit_def lt_def ge_def) from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by (auto simp add: split_def) let ?st = "Add (Mul m t) (Mul n s)" from tnb snb have stnb: "numbound0 ?st" by simp from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute) from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx have "\ x. ?I x ?rq" by auto thus "?E" using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def) qed with MF PF show ?thesis by blast qed lemma \_cong_aux: assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" shows "((\ (t,n). Inum (x#bs) t /real_of_int n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (set U \ set U))" (is "?lhs = ?rhs") proof(auto) fix t n s m assume "((t,n),(s,m)) \ set (alluopairs U)" hence th: "((t,n),(s,m)) \ (set U \ set U)" using alluopairs_set1[where xs="U"] by blast let ?N = "\ t. Inum (x#bs) t" let ?st= "Add (Mul m t) (Mul n s)" from Ul th have mnz: "m \ 0" by auto from Ul th have nnz: "n \ 0" by auto have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) thus "(real_of_int m * Inum (x # bs) t + real_of_int n * Inum (x # bs) s) / (2 * real_of_int n * real_of_int m) \ (\((t, n), s, m). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` (set U \ set U)"using mnz nnz th - apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) - by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute) + by (force simp add: th add_divide_distrib algebra_simps split_def image_def) next fix t n s m assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" let ?N = "\ t. Inum (x#bs) t" let ?st= "Add (Mul m t) (Mul n s)" from Ul smU have mnz: "m \ 0" by auto from Ul tnU have nnz: "n \ 0" by auto have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" have Pc:"\ a b. ?P a b = ?P b a" by auto from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast from alluopairs_ex[OF Pc, where xs="U"] tnU smU have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" by blast then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" and Pts': "?P (t',n') (s',m')" by blast from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto let ?st' = "Add (Mul m' t') (Mul n' s')" have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m')/2 = ?N ?st' / real_of_int (2*n'*m')" using mnz' nnz' by (simp add: algebra_simps add_divide_distrib) from Pts' have "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" by simp also have "\ = ((\(t, n). Inum (x # bs) t / real_of_int n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') finally show "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2 \ (\(t, n). Inum (x # bs) t / real_of_int n) ` (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` set (alluopairs U)" using ts'_U by blast qed lemma \_cong: assumes lp: "isrlfm p" and UU': "((\ (t,n). Inum (x#bs) t /real_of_int n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") and U: "\ (t,n) \ U. numbound0 t \ n > 0" and U': "\ (t,n) \ U'. numbound0 t \ n > 0" shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (\ p (t,n)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and Pst: "Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))" by blast let ?N = "\ t. Inum (x#bs) t" from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" by auto (rule_tac x="(a,b)" in bexI, auto) then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp from conjunct1[OF \_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) then show ?rhs using tnU' by auto next assume ?rhs then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (\ p (t', n'))" by blast from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" by auto (rule_tac x="(a,b)" in bexI, auto) then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and th: "?f (t',n') = ?g((t,n),(s,m)) "by blast let ?N = "\ t. Inum (x#bs) t" from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from \_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp with \_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast qed lemma ferrack01: assumes qf: "qfree p" shows "((\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \ qfree (ferrack01 p)" (is "(?lhs = ?rhs) \ _") proof- let ?I = "\ x p. Ifm (x#bs) p" fix x let ?N = "\ t. Inum (x#bs) t" let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)" let ?U = "\ ?q" let ?Up = "alluopairs ?U" let ?g = "\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" let ?S = "map ?g ?Up" let ?SS = "map simp_num_pair ?S" let ?Y = "remdups ?SS" let ?f= "(\ (t,n). ?N t / real_of_int n)" let ?h = "\ ((t,n),(s,m)). (?N t/real_of_int n + ?N s/ real_of_int m) /2" let ?F = "\ p. \ a \ set (\ p). \ b \ set (\ p). ?I x (\ p (?g(a,b)))" let ?ep = "evaldjf (\ ?q) ?Y" from rlfm_l[OF qf] have lq: "isrlfm ?q" by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def) from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp from \_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . from U_l UpU have "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " by (auto) have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" proof- { fix t n assume tnY: "(t,n) \ set ?Y" hence "(t,n) \ set ?SS" by simp hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" by (auto simp add: split_def simp del: map_map) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) then obtain t' n' where tn'S: "(t',n') \ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto from simp_num_pair_l[OF tnb np tns] have "numbound0 t \ n > 0" . } thus ?thesis by blast qed have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" proof- from simp_num_pair_ci[where bs="x#bs"] have "\x. (?f o simp_num_pair) x = ?f x" by auto hence th: "?f o simp_num_pair = ?f" using ext by blast have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_comp comp_assoc) also have "\ = (?f ` set ?S)" by (simp add: th) also have "\ = ((?f o ?g) ` set ?Up)" by (simp only: set_map o_def image_comp) also have "\ = (?h ` (set ?U \ set ?U))" using \_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_comp] by blast finally show ?thesis . qed have "\ (t,n) \ set ?Y. bound0 (\ ?q (t,n))" proof- { fix t n assume tnY: "(t,n) \ set ?Y" with Y_l have tnb: "numbound0 t" and np: "real_of_int n > 0" by auto from \_I[OF lq np tnb] have "bound0 (\ ?q (t,n))" by simp} thus ?thesis by blast qed hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="\ ?q"] by auto from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q" by (simp only: split_def fst_conv snd_conv) also have "\ = (\ (t,n) \ set ?Y. ?I x (\ ?q (t,n)))" using \_cong[OF lq YU U_l Y_l] by (simp only: split_def fst_conv snd_conv) also have "\ = (Ifm (x#bs) ?ep)" using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\ ?q",symmetric] by (simp only: split_def prod.collapse) also have "\ = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def) from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def) with lr show ?thesis by blast qed lemma cp_thm': assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and up: "d_\ p 1" and dd: "d_\ p d" and dp: "d > 0" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. d}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real_of_int i#bs)) ` set (\ p). Ifm ((b+real_of_int j)#bs) p))" using cp_thm[OF lp up dd dp] by auto definition unit :: "fm \ fm \ num list \ int" where "unit p \ (let p' = zlfm p ; l = \ p' ; q = And (Dvd l (CN 0 1 (C 0))) (a_\ p' l); d = \ q; B = remdups (map simpnum (\ q)) ; a = remdups (map simpnum (\ q)) in if length B \ length a then (q,B,d) else (mirror q, a,d))" lemma unit: assumes qf: "qfree p" shows "\ q B d. unit p = (q,B,d) \ ((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ (Inum (real_of_int i#bs)) ` set B = (Inum (real_of_int i#bs)) ` set (\ q) \ d_\ q 1 \ d_\ q d \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ b\ set B. numbound0 b)" proof- fix q B d assume qBd: "unit p = (q,B,d)" let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ Inum (real_of_int i#bs) ` set B = Inum (real_of_int i#bs) ` set (\ q) \ d_\ q 1 \ d_\ q d \ 0 < d \ iszlfm q (real_of_int i # bs) \ (\ b\ set B. numbound0 b)" let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?p' = "zlfm p" let ?l = "\ ?p'" let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\ ?p' ?l)" let ?d = "\ ?q" let ?B = "set (\ ?q)" let ?B'= "remdups (map simpnum (\ ?q))" let ?A = "set (\ ?q)" let ?A'= "remdups (map simpnum (\ ?q))" from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?p' = ?I i p" by auto from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] have lp': "\ (i::int). iszlfm ?p' (real_of_int i#bs)" by simp hence lp'': "iszlfm ?p' (real_of_int (i::int)#bs)" by simp from lp' \[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d_\ ?p' ?l" by auto from a_\_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) from lp'' lp a_\[OF lp'' dl lp] have lq:"iszlfm ?q (real_of_int i#bs)" and uq: "d_\ ?q 1" by (auto simp add: isint_def) from \[OF lq] have dp:"?d >0" and dd: "d_\ ?q ?d" by blast+ let ?N = "\ t. Inum (real_of_int (i::int)#bs) t" have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp) also have "\ = ?N ` ?B" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have BB': "?N ` set ?B' = ?N ` ?B" . have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp) also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_numbound0[OF lq] have B_nb:"\ b\ set ?B'. numbound0 b" by simp from \_l[OF lq] have A_nb: "\ b\ set ?A'. numbound0 b" by simp { assume "length ?B' \ length ?A'" hence q:"q=?q" and "B = ?B'" and d:"d = ?d" using qBd by (auto simp add: Let_def unit_def) with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" and bn: "\b\ set B. numbound0 b" by simp+ with pq_ex dp uq dd lq q d have ?thes by simp } moreover { assume "\ (length ?B' \ length ?A')" hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" using qBd by (auto simp add: Let_def unit_def) with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" and bn: "\b\ set B. numbound0 b" by simp+ from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp from lq uq q mirror_d_\ [where p="?q" and bs="bs" and a="real_of_int i"] have lq': "iszlfm q (real_of_int i#bs)" and uq: "d_\ q 1" by auto from \[OF lq'] mirror_\[OF lq] q d have dq:"d_\ q d " by auto from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp } ultimately show ?thes by blast qed (* Cooper's Algorithm *) definition cooper :: "fm \ fm" where "cooper p \ (let (q,B,d) = unit p; js = [1..d]; mq = simpfm (minusinf q); md = evaldjf (\ j. simpfm (subst0 (C j) mq)) js in if md = T then T else (let qd = evaldjf (\ t. simpfm (subst0 t q)) (remdups (map (\ (b,j). simpnum (Add b (C j))) [(b,j). b\B,j\js])) in decr (disj md qd)))" lemma cooper: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" (is "(?lhs = ?rhs) \ _") proof- let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "fst (unit p)" let ?B = "fst (snd(unit p))" let ?d = "snd (snd (unit p))" let ?js = "[1..?d]" let ?mq = "minusinf ?q" let ?smq = "simpfm ?mq" let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" fix i let ?N = "\ t. Inum (real_of_int (i::int)#bs) t" let ?bjs = "[(b,j). b\?B,j\?js]" let ?sbjs = "map (\ (b,j). simpnum (Add b (C j))) ?bjs" let ?qd = "evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs)" have qbf:"unit p = (?q,?B,?d)" by simp from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and B:"?N ` set ?B = ?N ` set (\ ?q)" and uq:"d_\ ?q 1" and dd: "d_\ ?q ?d" and dp: "?d > 0" and lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ b\ set ?B. numbound0 b" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" by (auto simp only: subst0_bound0[OF qfmq]) hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" by auto from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp from Bn jsnb have "\ (b,j) \ set ?bjs. numbound0 (Add b (C j))" by simp hence "\ (b,j) \ set ?bjs. numbound0 (simpnum (Add b (C j)))" using simpnum_numbound0 by blast hence "\ t \ set ?sbjs. numbound0 t" by simp hence "\ t \ set (remdups ?sbjs). bound0 (subst0 t ?q)" using subst0_bound0[OF qfq] by auto hence th': "\ t \ set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))" using simpfm_bound0 by blast from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real_of_int j)#bs) ?q))" by auto also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real_of_int j)#bs) ?q))" by auto also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci) also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" by (auto simp add: split_def) also have "\ = ((\ j\ set ?js. (\ j. ?I i (simpfm (subst0 (C j) ?smq))) j) \ (\ t \ set (remdups ?sbjs). (\ t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] Inum.simps subst0_I[OF qfmq] set_remdups) also have "\ = ((?I i (evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js)) \ (?I i (evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex) finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by simp hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp {assume mdT: "?md = T" hence cT:"cooper p = T" by (simp only: cooper_def unit_def split_def Let_def if_True) simp from mdT mdqd have lhs:"?lhs" by auto from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) with lhs cT have ?thesis by simp } moreover {assume mdT: "?md \ T" hence "cooper p = decr (disj ?md ?qd)" by (simp only: cooper_def unit_def split_def Let_def if_False) with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } ultimately show ?thesis by blast qed lemma DJcooper: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" proof- from cooper have cqf: "\ p. qfree p \ qfree (cooper p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" by (simp add: DJ_def evaldjf_ex) also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using cooper disjuncts_qf[OF qf] by blast also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast qed (* Redy and Loveland *) lemma \_\_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" shows "Ifm (a#bs) (\_\ p (t,c)) = Ifm (a#bs) (\_\ p (t',c))" using lp by (induct p rule: iszlfm.induct, auto simp add: tt') lemma \_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" shows "Ifm (a#bs) (\ p c t) = Ifm (a#bs) (\ p c t')" by (simp add: \_def tt' \_\_cong[OF lp tt']) lemma \_cong: assumes lp: "iszlfm p (a#bs)" and RR: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" shows "(\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))) = (\ (e,c) \ set (\ p). \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j))))" (is "?lhs = ?rhs") proof let ?d = "\ p" assume ?lhs then obtain e c j where ecR: "(e,c) \ R" and jD:"j \ {1 .. c*?d}" and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" by auto hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" using RR by simp hence "\ (e',c') \ set (\ p). Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto then obtain e' c' where ecRo:"(e',c') \ set (\ p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'" and cc':"c = c'" by blast from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp - from ecRo jD px' show ?rhs apply (auto simp: cc') - by (rule_tac x="(e', c')" in bexI,simp_all) - (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) + from ecRo jD px' show ?rhs + using cc' by blast next let ?d = "\ p" assume ?rhs then obtain e c j where ecR: "(e,c) \ set (\ p)" and jD:"j \ {1 .. c*?d}" and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" by auto hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp hence "\ (e',c') \ R. Inum (a#bs) e = Inum (a#bs) e' \ c = c'" by auto then obtain e' c' where ecRo:"(e',c') \ R" and ee':"Inum (a#bs) e = Inum (a#bs) e'" and cc':"c = c'" by blast from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp - from ecRo jD px' show ?lhs apply (auto simp: cc') - by (rule_tac x="(e', c')" in bexI,simp_all) - (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) + from ecRo jD px' show ?lhs + using cc' by blast qed lemma rl_thm': assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and R: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp definition chooset :: "fm \ fm \ ((num\int) list) \ int" where "chooset p \ (let q = zlfm p ; d = \ q; B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; a = remdups (map (\ (t,k). (simpnum t,k)) (\_\ q)) in if length B \ length a then (q,B,d) else (mirror q, a,d))" lemma chooset: assumes qf: "qfree p" shows "\ q B d. chooset p = (q,B,d) \ ((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" proof- fix q B d assume qBd: "chooset p = (q,B,d)" let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "zlfm p" let ?d = "\ ?q" let ?B = "set (\ ?q)" let ?f = "\ (t,k). (simpnum t,k)" let ?B'= "remdups (map ?f (\ ?q))" let ?A = "set (\_\ ?q)" let ?A'= "remdups (map ?f (\_\ ?q))" from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?q = ?I i p" by auto hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real_of_int i"] have lq: "iszlfm ?q (real_of_int (i::int)#bs)" . from \[OF lq] have dp:"?d >0" by blast let ?N = "\ (t,c). (Inum (real_of_int (i::int)#bs) t,c)" have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_comp) also have "\ = ?N ` ?B" by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have BB': "?N ` set ?B' = ?N ` ?B" . have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp) also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_l[OF lq] have B_nb:"\ (e,c)\ set ?B'. numbound0 e \ c > 0" by (simp add: split_def) from \_\_l[OF lq] have A_nb: "\ (e,c)\ set ?A'. numbound0 e \ c > 0" by (simp add: split_def) {assume "length ?B' \ length ?A'" hence q:"q=?q" and "B = ?B'" and d:"d = ?d" using qBd by (auto simp add: Let_def chooset_def) with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto with pq_ex dp lq q d have ?thes by simp} moreover {assume "\ (length ?B' \ length ?A')" hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" using qBd by (auto simp add: Let_def chooset_def) with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp from lq q mirror_l [where p="?q" and bs="bs" and a="real_of_int i"] have lq': "iszlfm q (real_of_int i#bs)" by auto from mirror_\[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp } ultimately show ?thes by blast qed definition stage :: "fm \ int \ (num \ int) \ fm" where "stage p d \ (\ (e,c). evaldjf (\ j. simpfm (\ p c (Add e (C j)))) [1..c*d])" lemma stage: shows "Ifm bs (stage p d (e,c)) = (\ j\{1 .. c*d}. Ifm bs (\ p c (Add e (C j))))" by (unfold stage_def split_def ,simp only: evaldjf_ex simpfm) simp lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e" shows "bound0 (stage p d (e,c))" proof- let ?f = "\ j. simpfm (\ p c (Add e (C j)))" have th: "\ j\ set [1..c*d]. bound0 (?f j)" proof fix j from nb have nb':"numbound0 (Add e (C j))" by simp from simpfm_bound0[OF \_nb[OF lp nb', where k="c"]] show "bound0 (simpfm (\ p c (Add e (C j))))" . qed from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp qed definition redlove :: "fm \ fm" where "redlove p \ (let (q,B,d) = chooset p; mq = simpfm (minusinf q); md = evaldjf (\ j. simpfm (subst0 (C j) mq)) [1..d] in if md = T then T else (let qd = evaldjf (stage q d) B in decr (disj md qd)))" lemma redlove: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" (is "(?lhs = ?rhs) \ _") proof- let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "fst (chooset p)" let ?B = "fst (snd(chooset p))" let ?d = "snd (snd (chooset p))" let ?js = "[1..?d]" let ?mq = "minusinf ?q" let ?smq = "simpfm ?mq" let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" fix i let ?N = "\ (t,k). (Inum (real_of_int (i::int)#bs) t,k)" let ?qd = "evaldjf (stage ?q ?d) ?B" have qbf:"chooset p = (?q,?B,?d)" by simp from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ (e,c)\ set ?B. numbound0 e \ c > 0" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" by (auto simp only: subst0_bound0[OF qfmq]) hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" by auto from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp from Bn stage_nb[OF lq] have th:"\ x \ set ?B. bound0 (stage ?q ?d x)" by auto from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" . from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex rl_thm'[OF lq B]] dd have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real_of_int i#bs) (\ ?q c (Add e (C j)))))" by auto also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" by (simp add: stage split_def) also have "\ = ((\ j\ {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \ ?I i ?qd)" by (simp add: evaldjf_ex subst0_I[OF qfmq]) finally have mdqd:"?lhs = (?I i ?md \ ?I i ?qd)" by (simp only: evaldjf_ex set_upto simpfm) also have "\ = (?I i (disj ?md ?qd))" by simp also have "\ = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . {assume mdT: "?md = T" hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def) from mdT have lhs:"?lhs" using mdqd by simp from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def) with lhs cT have ?thesis by simp } moreover {assume mdT: "?md \ T" hence "redlove p = decr (disj ?md ?qd)" by (simp add: redlove_def chooset_def split_def Let_def) with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } ultimately show ?thesis by blast qed lemma DJredlove: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" proof- from redlove have cqf: "\ p. qfree p \ qfree (redlove p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" by (simp add: DJ_def evaldjf_ex) also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using redlove disjuncts_qf[OF qf] by blast also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast qed lemma exsplit_qf: assumes qf: "qfree p" shows "qfree (exsplit p)" using qf by (induct p rule: exsplit.induct, auto) definition mircfr :: "fm \ fm" where "mircfr = DJ cooper o ferrack01 o simpfm o exsplit" definition mirlfr :: "fm \ fm" where "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit" lemma mircfr: "\ bs p. qfree p \ qfree (mircfr p) \ Ifm bs (mircfr p) = Ifm bs (E p)" proof(clarsimp simp del: Ifm.simps) fix bs p assume qf: "qfree p" show "qfree (mircfr p)\(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) qed qed lemma mirlfr: "\ bs p. qfree p \ qfree(mirlfr p) \ Ifm bs (mirlfr p) = Ifm bs (E p)" proof(clarsimp simp del: Ifm.simps) fix bs p assume qf: "qfree p" show "qfree (mirlfr p)\(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) qed qed definition mircfrqe:: "fm \ fm" where "mircfrqe p = qelim (prep p) mircfr" definition mirlfrqe:: "fm \ fm" where "mirlfrqe p = qelim (prep p) mirlfr" theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \ qfree (mircfrqe p)" using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def) theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \ qfree (mirlfrqe p)" using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def) definition "problem1 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))" definition "problem2 = A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0))))" definition "problem3 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))" definition "problem4 = E (And (Ge (Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg (Bound 0))))))" ML_val \@{code mircfrqe} @{code problem1}\ ML_val \@{code mirlfrqe} @{code problem1}\ ML_val \@{code mircfrqe} @{code problem2}\ ML_val \@{code mirlfrqe} @{code problem2}\ ML_val \@{code mircfrqe} @{code problem3}\ ML_val \@{code mirlfrqe} @{code problem3}\ ML_val \@{code mircfrqe} @{code problem4}\ ML_val \@{code mirlfrqe} @{code problem4}\ (*code_reflect Mir functions mircfrqe mirlfrqe file "mir.ML"*) oracle mirfr_oracle = \ let val mk_C = @{code C} o @{code int_of_integer}; val mk_Dvd = @{code Dvd} o apfst @{code int_of_integer}; val mk_Bound = @{code Bound} o @{code nat_of_integer}; fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (=) vs t of NONE => error "Variable not found in the list!" | SOME n => mk_Bound n) | num_of_term vs \<^term>\of_int (0::int)\ = mk_C 0 | num_of_term vs \<^term>\of_int (1::int)\ = mk_C 1 | num_of_term vs \<^term>\0::real\ = mk_C 0 | num_of_term vs \<^term>\1::real\ = mk_C 1 | num_of_term vs \<^term>\- 1::real\ = mk_C (~ 1) | num_of_term vs (Bound i) = mk_Bound i | num_of_term vs \<^Const_>\uminus \<^Type>\real\ for t'\ = @{code Neg} (num_of_term vs t') | num_of_term vs \<^Const_>\plus \<^Type>\real\ for t1 t2\ = @{code Add} (num_of_term vs t1, num_of_term vs t2) | num_of_term vs \<^Const_>\minus \<^Type>\real\ for t1 t2\ = @{code Sub} (num_of_term vs t1, num_of_term vs t2) | num_of_term vs \<^Const_>\times \<^Type>\real\ for t1 t2\ = (case num_of_term vs t1 of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported Multiplication") | num_of_term vs \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\numeral \<^Type>\int\ for t'\\ = mk_C (HOLogic.dest_numeral t') | num_of_term vs \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\uminus \<^Type>\int\ for \<^Const_>\numeral \<^Type>\int\ for t'\\\ = mk_C (~ (HOLogic.dest_numeral t')) | num_of_term vs \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\floor \<^Type>\real\ for t'\\ = @{code Floor} (num_of_term vs t') | num_of_term vs \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\ceiling \<^Type>\real\ for t'\\ = @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) | num_of_term vs \<^Const_>\numeral \<^Type>\real\ for t'\ = mk_C (HOLogic.dest_numeral t') | num_of_term vs \<^Const_>\uminus \<^Type>\real\ for \<^Const_>\numeral \<^Type>\real\ for t'\\ = mk_C (~ (HOLogic.dest_numeral t')) | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term \<^context> t); fun fm_of_term vs \<^Const_>\True\ = @{code T} | fm_of_term vs \<^Const_>\False\ = @{code F} | fm_of_term vs \<^Const_>\less \<^Type>\real\ for t1 t2\ = @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs \<^Const_>\less_eq \<^Type>\real\ for t1 t2\ = @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs \<^Const_>\HOL.eq \<^Type>\real\ for t1 t2\ = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs \<^Const_>\rdvd for \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\numeral \<^Type>\int\ for t1\\ t2\ = mk_Dvd (HOLogic.dest_numeral t1, num_of_term vs t2) | fm_of_term vs \<^Const_>\rdvd for \<^Const_>\of_int \<^Type>\real\ for \<^Const_>\uminus \<^Type>\int\ for \<^Const_>\numeral \<^Type>\int\ for t1\\\ t2\ = mk_Dvd (~ (HOLogic.dest_numeral t1), num_of_term vs t2) | fm_of_term vs \<^Const_>\HOL.eq \<^Type>\bool\ for t1 t2\ = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs \<^Const_>\HOL.conj for t1 t2\ = @{code And} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs \<^Const_>\HOL.disj for t1 t2\ = @{code Or} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs \<^Const_>\HOL.implies for t1 t2\ = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) | fm_of_term vs \<^Const_>\HOL.Not for t'\ = @{code Not} (fm_of_term vs t') | fm_of_term vs \<^Const_>\Ex _ for \Abs (xn, xT, p)\\ = @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) | fm_of_term vs \<^Const_>\All _ for \Abs (xn, xT, p)\\ = @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term \<^context> t); fun term_of_num vs (@{code C} i) = \<^Const>\of_int \<^Type>\real\ for \HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)\\ | term_of_num vs (@{code Bound} n) = let val m = @{code integer_of_nat} n; in fst (the (find_first (fn (_, q) => m = q) vs)) end | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) = \<^Const>\of_int \<^Type>\real\ for \<^Const>\ceiling \<^Type>\real\ for \term_of_num vs t'\\\ | term_of_num vs (@{code Neg} t') = \<^Const>\uminus \<^Type>\real\ for \term_of_num vs t'\\ | term_of_num vs (@{code Add} (t1, t2)) = \<^Const>\plus \<^Type>\real\ for \term_of_num vs t1\ \term_of_num vs t2\\ | term_of_num vs (@{code Sub} (t1, t2)) = \<^Const>\minus \<^Type>\real\ for \term_of_num vs t1\ \term_of_num vs t2\\ | term_of_num vs (@{code Mul} (i, t2)) = \<^Const>\times \<^Type>\real\ for \term_of_num vs (@{code C} i)\ \term_of_num vs t2\\ | term_of_num vs (@{code Floor} t) = \<^Const>\of_int \<^Type>\real\ for \<^Const>\floor \<^Type>\real\ for \term_of_num vs t\\\ | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); fun term_of_fm vs @{code T} = \<^Const>\True\ | term_of_fm vs @{code F} = \<^Const>\False\ | term_of_fm vs (@{code Lt} t) = \<^Const>\less \<^Type>\real\ for \term_of_num vs t\ \<^term>\0::real\\ | term_of_fm vs (@{code Le} t) = \<^Const>\less_eq \<^Type>\real\ for \term_of_num vs t\ \<^term>\0::real\\ | term_of_fm vs (@{code Gt} t) = \<^Const>\less \<^Type>\real\ for \<^term>\0::real\ \term_of_num vs t\\ | term_of_fm vs (@{code Ge} t) = \<^Const>\less_eq \<^Type>\real\ for \<^term>\0::real\ \term_of_num vs t\\ | term_of_fm vs (@{code Eq} t) = \<^Const>\HOL.eq \<^Type>\real\ for \term_of_num vs t\ \<^term>\0::real\\ | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code Not} (@{code Eq} t)) | term_of_fm vs (@{code Dvd} (i, t)) = \<^Const>\rdvd for \term_of_num vs (@{code C} i)\ \term_of_num vs t\\ | term_of_fm vs (@{code NDvd} (i, t)) = term_of_fm vs (@{code Not} (@{code Dvd} (i, t))) | term_of_fm vs (@{code Not} t') = HOLogic.Not $ term_of_fm vs t' | term_of_fm vs (@{code And} (t1, t2)) = HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 | term_of_fm vs (@{code Or} (t1, t2)) = HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 | term_of_fm vs (@{code Imp} (t1, t2)) = HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 | term_of_fm vs (@{code Iff} (t1, t2)) = \<^Const>\HOL.eq \<^Type>\bool\ for \term_of_fm vs t1\ \term_of_fm vs t2\\; in fn (ctxt, t) => let val fs = Misc_Legacy.term_frees t; val vs = map_index swap fs; (*If quick_and_dirty then run without proof generation as oracle*) val qe = if Config.get ctxt quick_and_dirty then @{code mircfrqe} else @{code mirlfrqe}; val t' = term_of_fm vs (qe (fm_of_term vs t)); in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end end \ lemmas iff_real_of_int = of_int_eq_iff [where 'a = real, symmetric] of_int_less_iff [where 'a = real, symmetric] of_int_le_iff [where 'a = real, symmetric] ML_file \mir_tac.ML\ method_setup mir = \ Scan.lift (Args.mode "no_quantify") >> (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q))) \ "decision procedure for MIR arithmetic" lemma "\x::real. (\x\ = \x\ \ (x = real_of_int \x\))" by mir lemma "\x::real. real_of_int (2::int)*x - (real_of_int (1::int)) < real_of_int \x\ + real_of_int \x\ \ real_of_int \x\ + real_of_int \x\ \ real_of_int (2::int)*x + (real_of_int (1::int))" by mir lemma "\x::real. 2*\x\ \ \2*x\ \ \2*x\ \ 2*\x+1\" by mir lemma "\x::real. \y \ x. (\x\ = \y\)" by mir lemma "\(x::real) (y::real). \x\ = \y\ \ 0 \ \y - x\ \ \y - x\ \ 1" by mir end diff --git a/src/HOL/Decision_Procs/Rat_Pair.thy b/src/HOL/Decision_Procs/Rat_Pair.thy --- a/src/HOL/Decision_Procs/Rat_Pair.thy +++ b/src/HOL/Decision_Procs/Rat_Pair.thy @@ -1,616 +1,589 @@ (* Title: HOL/Decision_Procs/Rat_Pair.thy Author: Amine Chaieb *) section \Rational numbers as pairs\ theory Rat_Pair imports Complex_Main begin type_synonym Num = "int \ int" abbreviation Num0_syn :: Num (\0\<^sub>N\) where "0\<^sub>N \ (0, 0)" abbreviation Numi_syn :: "int \ Num" (\'((_)')\<^sub>N\) where "(i)\<^sub>N \ (i, 1)" definition isnormNum :: "Num \ bool" where "isnormNum = (\(a, b). if a = 0 then b = 0 else b > 0 \ gcd a b = 1)" definition normNum :: "Num \ Num" where "normNum = (\(a,b). (if a = 0 \ b = 0 then (0, 0) else (let g = gcd a b in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))" declare gcd_dvd1[presburger] gcd_dvd2[presburger] lemma normNum_isnormNum [simp]: "isnormNum (normNum x)" proof - obtain a b where x: "x = (a, b)" by (cases x) consider "a = 0 \ b = 0" | "a \ 0" "b \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis by (simp add: x normNum_def isnormNum_def) next case ab: 2 let ?g = "gcd a b" let ?a' = "a div ?g" let ?b' = "b div ?g" let ?g' = "gcd ?a' ?b'" from ab have "?g \ 0" by simp with gcd_ge_0_int[of a b] have gpos: "?g > 0" by arith have gdvd: "?g dvd a" "?g dvd b" by arith+ from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] ab have nz': "?a' \ 0" "?b' \ 0" by - (rule notI, simp)+ from ab have stupid: "a \ 0 \ b \ 0" by arith from div_gcd_coprime[OF stupid] have gp1: "?g' = 1" by simp from ab consider "b < 0" | "b > 0" by arith then show ?thesis proof cases case b: 1 have False if b': "?b' \ 0" proof - from gpos have th: "?g \ 0" by arith from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)] show ?thesis using b by arith qed then have b': "?b' < 0" by (presburger add: linorder_not_le[symmetric]) from ab(1) nz' b b' gp1 show ?thesis by (simp add: x isnormNum_def normNum_def Let_def split_def) next case b: 2 then have "?b' \ 0" by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos]) with nz' have b': "?b' > 0" by arith from b b' ab(1) nz' gp1 show ?thesis by (simp add: x isnormNum_def normNum_def Let_def split_def) qed qed qed text \Arithmetic over Num\ definition Nadd :: "Num \ Num \ Num" (infixl \+\<^sub>N\ 60) where "Nadd = (\(a, b) (a', b'). if a = 0 \ b = 0 then normNum (a', b') else if a' = 0 \ b' = 0 then normNum (a, b) else normNum (a * b' + b * a', b * b'))" definition Nmul :: "Num \ Num \ Num" (infixl \*\<^sub>N\ 60) where "Nmul = (\(a, b) (a', b'). let g = gcd (a * a') (b * b') in (a * a' div g, b * b' div g))" definition Nneg :: "Num \ Num" (\~\<^sub>N\) where "Nneg = (\(a, b). (- a, b))" definition Nsub :: "Num \ Num \ Num" (infixl \-\<^sub>N\ 60) where "Nsub = (\a b. a +\<^sub>N ~\<^sub>N b)" definition Ninv :: "Num \ Num" where "Ninv = (\(a, b). if a < 0 then (- b, \a\) else (b, a))" definition Ndiv :: "Num \ Num \ Num" (infixl \\
\<^sub>N\ 60) where "Ndiv = (\a b. a *\<^sub>N Ninv b)" lemma Nneg_normN[simp]: "isnormNum x \ isnormNum (~\<^sub>N x)" by (simp add: isnormNum_def Nneg_def split_def) lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)" by (simp add: Nadd_def split_def) lemma Nsub_normN[simp]: "isnormNum y \ isnormNum (x -\<^sub>N y)" by (simp add: Nsub_def split_def) lemma Nmul_normN[simp]: assumes xn: "isnormNum x" and yn: "isnormNum y" shows "isnormNum (x *\<^sub>N y)" proof - obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) consider "a = 0" | "a' = 0" | "a \ 0" "a' \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis using xn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) next case 2 then show ?thesis using yn x y by (simp add: isnormNum_def Let_def Nmul_def split_def) next case aa': 3 then have bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def) from bp have "x *\<^sub>N y = normNum (a * a', b * b')" using x y aa' bp by (simp add: Nmul_def Let_def split_def normNum_def) then show ?thesis by simp qed qed lemma Ninv_normN[simp]: "isnormNum x \ isnormNum (Ninv x)" - apply (simp add: Ninv_def isnormNum_def split_def) - apply (cases "fst x = 0") - apply (auto simp add: gcd.commute) - done + by (simp add: Ninv_def isnormNum_def split_def gcd.commute split: if_split_asm) lemma isnormNum_int[simp]: "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \ 0 \ isnormNum (i)\<^sub>N" by (simp_all add: isnormNum_def) text \Relations over Num\ definition Nlt0:: "Num \ bool" (\0>\<^sub>N\) where "Nlt0 = (\(a, b). a < 0)" definition Nle0:: "Num \ bool" (\0\\<^sub>N\) where "Nle0 = (\(a, b). a \ 0)" definition Ngt0:: "Num \ bool" (\0<\<^sub>N\) where "Ngt0 = (\(a, b). a > 0)" definition Nge0:: "Num \ bool" (\0\\<^sub>N\) where "Nge0 = (\(a, b). a \ 0)" definition Nlt :: "Num \ Num \ bool" (infix \<\<^sub>N\ 55) where "Nlt = (\a b. 0>\<^sub>N (a -\<^sub>N b))" definition Nle :: "Num \ Num \ bool" (infix \\\<^sub>N\ 55) where "Nle = (\a b. 0\\<^sub>N (a -\<^sub>N b))" definition "INum = (\(a, b). of_int a / of_int b)" lemma INum_int [simp]: "INum (i)\<^sub>N = (of_int i ::'a::field)" "INum 0\<^sub>N = (0::'a::field)" by (simp_all add: INum_def) lemma isnormNum_unique[simp]: assumes na: "isnormNum x" and nb: "isnormNum y" shows "(INum x ::'a::field_char_0) = INum y \ x = y" (is "?lhs = ?rhs") proof obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) consider "a = 0 \ b = 0 \ a' = 0 \ b' = 0" | "a \ 0" "b \ 0" "a' \ 0" "b' \ 0" by blast then show ?rhs if H: ?lhs proof cases case 1 then show ?thesis using na nb H by (simp add: x y INum_def split_def isnormNum_def split: if_split_asm) next case 2 with na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def) from H \b \ 0\ \b' \ 0\ have eq: "a * b' = a' * b" by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult) from \a \ 0\ \a' \ 0\ na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1" by (simp_all add: x y isnormNum_def add: gcd.commute) then have "coprime a b" "coprime b a" "coprime a' b'" "coprime b' a'" by (simp_all add: coprime_iff_gcd_eq_1) from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'" - apply - - apply algebra - apply algebra - apply simp - apply algebra - done + by (algebra|simp)+ then have eq1: "b = b'" using pos \coprime b a\ \coprime b' a'\ by (auto simp add: coprime_dvd_mult_left_iff intro: associated_eqI) with eq have "a = a'" using pos by simp with \b = b'\ show ?thesis by (simp add: x y) qed show ?lhs if ?rhs using that by simp qed lemma isnormNum0[simp]: "isnormNum x \ INum x = (0::'a::field_char_0) \ x = 0\<^sub>N" unfolding INum_int(2)[symmetric] by (rule isnormNum_unique) simp_all lemma of_int_div_aux: assumes "d \ 0" shows "(of_int x ::'a::field_char_0) / of_int d = of_int (x div d) + (of_int (x mod d)) / of_int d" proof - let ?t = "of_int (x div d) * (of_int d ::'a) + of_int (x mod d)" let ?f = "\x. x / of_int d" have "x = (x div d) * d + x mod d" by auto then have eq: "of_int x = ?t" by (simp only: of_int_mult[symmetric] of_int_add [symmetric]) then have "of_int x / of_int d = ?t / of_int d" using cong[OF refl[of ?f] eq] by simp then show ?thesis by (simp add: add_divide_distrib algebra_simps \d \ 0\) qed lemma of_int_div: fixes d :: int assumes "d \ 0" "d dvd n" shows "(of_int (n div d) ::'a::field_char_0) = of_int n / of_int d" using assms of_int_div_aux [of d n, where ?'a = 'a] by simp lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::field_char_0)" proof - obtain a b where x: "x = (a, b)" by (cases x) consider "a = 0 \ b = 0" | "a \ 0" "b \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis by (simp add: x INum_def normNum_def split_def Let_def) next case ab: 2 let ?g = "gcd a b" from ab have g: "?g \ 0"by simp from of_int_div[OF g, where ?'a = 'a] show ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) qed qed lemma INum_normNum_iff: "(INum x ::'a::field_char_0) = INum y \ normNum x = normNum y" (is "?lhs \ _") proof - have "normNum x = normNum y \ (INum (normNum x) :: 'a) = INum (normNum y)" by (simp del: normNum) also have "\ = ?lhs" by simp finally show ?thesis by simp qed lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: field_char_0)" proof - let ?z = "0::'a" obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) consider "a = 0 \ a'= 0 \ b =0 \ b' = 0" | "a \ 0" "a'\ 0" "b \ 0" "b' \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis - apply (cases "a = 0") - apply (simp_all add: x y Nadd_def) - apply (cases "b = 0") - apply (simp_all add: INum_def) - apply (cases "a'= 0") - apply simp_all - apply (cases "b'= 0") - apply simp_all - done + by (auto simp: x y INum_def Nadd_def normNum_def Let_def of_int_div) next case neq: 2 show ?thesis proof (cases "a * b' + b * a' = 0") case True then have "of_int (a * b' + b * a') / (of_int b * of_int b') = ?z" by simp then have "of_int b' * of_int a / (of_int b * of_int b') + of_int b * of_int a' / (of_int b * of_int b') = ?z" by (simp add: add_divide_distrib) then have th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using neq by simp from True neq show ?thesis by (simp add: x y th Nadd_def normNum_def INum_def split_def) next case False let ?g = "gcd (a * b' + b * a') (b * b')" have gz: "?g \ 0" using False by simp show ?thesis using neq False gz of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * b' + b * a'" "b * b'"]] of_int_div [where ?'a = 'a, OF gz gcd_dvd2 [of "a * b' + b * a'" "b * b'"]] by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps) qed qed qed lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a::field_char_0)" proof - let ?z = "0::'a" obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) consider "a = 0 \ a' = 0 \ b = 0 \ b' = 0" | "a \ 0" "a' \ 0" "b \ 0" "b' \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis by (auto simp add: x y Nmul_def INum_def) next case neq: 2 let ?g = "gcd (a * a') (b * b')" have gz: "?g \ 0" using neq by simp from neq of_int_div [where ?'a = 'a, OF gz gcd_dvd1 [of "a * a'" "b * b'"]] of_int_div [where ?'a = 'a , OF gz gcd_dvd2 [of "a * a'" "b * b'"]] show ?thesis by (simp add: Nmul_def x y Let_def INum_def) qed qed lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x :: 'a::field)" by (simp add: Nneg_def split_def INum_def) lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a::field_char_0)" by (simp add: Nsub_def split_def) lemma Ninv[simp]: "INum (Ninv x) = (1 :: 'a::field) / INum x" by (simp add: Ninv_def INum_def split_def) lemma Ndiv[simp]: "INum (x \
\<^sub>N y) = INum x / (INum y :: 'a::field_char_0)" by (simp add: Ndiv_def) lemma Nlt0_iff[simp]: assumes nx: "isnormNum x" shows "((INum x :: 'a::{field_char_0,linordered_field}) < 0) = 0>\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) show ?thesis proof (cases "a = 0") case True then show ?thesis by (simp add: x Nlt0_def INum_def) next case False then have b: "(of_int b::'a) > 0" using nx by (simp add: x isnormNum_def) from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"] show ?thesis by (simp add: x Nlt0_def INum_def) qed qed lemma Nle0_iff[simp]: assumes nx: "isnormNum x" shows "((INum x :: 'a::{field_char_0,linordered_field}) \ 0) = 0\\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) show ?thesis proof (cases "a = 0") case True then show ?thesis by (simp add: x Nle0_def INum_def) next case False then have b: "(of_int b :: 'a) > 0" using nx by (simp add: x isnormNum_def) from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"] show ?thesis by (simp add: x Nle0_def INum_def) qed qed lemma Ngt0_iff[simp]: assumes nx: "isnormNum x" shows "((INum x :: 'a::{field_char_0,linordered_field}) > 0) = 0<\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) show ?thesis proof (cases "a = 0") case True then show ?thesis by (simp add: x Ngt0_def INum_def) next case False then have b: "(of_int b::'a) > 0" using nx by (simp add: x isnormNum_def) from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"] show ?thesis by (simp add: x Ngt0_def INum_def) qed qed lemma Nge0_iff[simp]: assumes nx: "isnormNum x" shows "(INum x :: 'a::{field_char_0,linordered_field}) \ 0 \ 0\\<^sub>N x" proof - obtain a b where x: "x = (a, b)" by (cases x) show ?thesis proof (cases "a = 0") case True then show ?thesis by (simp add: x Nge0_def INum_def) next case False then have b: "(of_int b::'a) > 0" using nx by (simp add: x isnormNum_def) from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"] show ?thesis by (simp add: x Nge0_def INum_def) qed qed lemma Nlt_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y" shows "((INum x :: 'a::{field_char_0,linordered_field}) < INum y) \ x <\<^sub>N y" proof - let ?z = "0::'a" have "((INum x ::'a) < INum y) \ INum (x -\<^sub>N y) < ?z" using nx ny by simp also have "\ \ 0>\<^sub>N (x -\<^sub>N y)" using Nlt0_iff[OF Nsub_normN[OF ny]] by simp finally show ?thesis by (simp add: Nlt_def) qed lemma Nle_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y" shows "((INum x :: 'a::{field_char_0,linordered_field}) \ INum y) \ x \\<^sub>N y" proof - have "((INum x ::'a) \ INum y) \ INum (x -\<^sub>N y) \ (0::'a)" using nx ny by simp also have "\ \ 0\\<^sub>N (x -\<^sub>N y)" using Nle0_iff[OF Nsub_normN[OF ny]] by simp finally show ?thesis by (simp add: Nle_def) qed lemma Nadd_commute: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "x +\<^sub>N y = y +\<^sub>N x" proof - have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp with isnormNum_unique[OF n] show ?thesis by simp qed lemma [simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "(0, b) +\<^sub>N y = normNum y" and "(a, 0) +\<^sub>N y = normNum y" and "x +\<^sub>N (0, b) = normNum x" and "x +\<^sub>N (a, 0) = normNum x" - apply (simp add: Nadd_def split_def) - apply (simp add: Nadd_def split_def) - apply (subst Nadd_commute) - apply (simp add: Nadd_def split_def) - apply (subst Nadd_commute) - apply (simp add: Nadd_def split_def) - done + by (simp_all add: Nadd_def normNum_def split_def) lemma normNum_nilpotent_aux[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" assumes nx: "isnormNum x" shows "normNum x = x" proof - let ?a = "normNum x" have n: "isnormNum ?a" by simp have th: "INum ?a = (INum x ::'a)" by simp with isnormNum_unique[OF n nx] show ?thesis by simp qed lemma normNum_nilpotent[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "normNum (normNum x) = normNum x" by simp lemma normNum0[simp]: "normNum (0, b) = 0\<^sub>N" "normNum (a, 0) = 0\<^sub>N" by (simp_all add: normNum_def) lemma normNum_Nadd: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp lemma Nadd_normNum1[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "normNum x +\<^sub>N y = x +\<^sub>N y" proof - have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp also have "\ = INum (x +\<^sub>N y)" by simp finally show ?thesis using isnormNum_unique[OF n] by simp qed lemma Nadd_normNum2[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "x +\<^sub>N normNum y = x +\<^sub>N y" proof - have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp also have "\ = INum (x +\<^sub>N y)" by simp finally show ?thesis using isnormNum_unique[OF n] by simp qed lemma Nadd_assoc: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)" proof - have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp with isnormNum_unique[OF n] show ?thesis by simp qed lemma Nmul_commute: "isnormNum x \ isnormNum y \ x *\<^sub>N y = y *\<^sub>N x" by (simp add: Nmul_def split_def Let_def gcd.commute mult.commute) lemma Nmul_assoc: assumes "SORT_CONSTRAINT('a::field_char_0)" assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z" shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)" proof - from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))" by simp_all have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp with isnormNum_unique[OF n] show ?thesis by simp qed lemma Nsub0: assumes "SORT_CONSTRAINT('a::field_char_0)" assumes x: "isnormNum x" and y: "isnormNum y" shows "x -\<^sub>N y = 0\<^sub>N \ x = y" proof - from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"] have "x -\<^sub>N y = 0\<^sub>N \ INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)" by simp also have "\ \ INum x = (INum y :: 'a)" by simp also have "\ \ x = y" using x y by simp finally show ?thesis . qed lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N" by (simp_all add: Nmul_def Let_def split_def) lemma Nmul_eq0[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" assumes nx: "isnormNum x" and ny: "isnormNum y" shows "x*\<^sub>N y = 0\<^sub>N \ x = 0\<^sub>N \ y = 0\<^sub>N" proof - obtain a b where x: "x = (a, b)" by (cases x) obtain a' b' where y: "y = (a', b')" by (cases y) have n0: "isnormNum 0\<^sub>N" by simp show ?thesis using nx ny - apply (simp only: isnormNum_unique[where ?'a = 'a, OF Nmul_normN[OF nx ny] n0, symmetric] - Nmul[where ?'a = 'a]) - apply (simp add: x y INum_def split_def isnormNum_def split: if_split_asm) - done + by (metis (no_types, opaque_lifting) INum_int(2) Nmul Nmul0(1) Nmul0(2) isnormNum0 mult_eq_0_iff) qed lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c" by (simp add: Nneg_def split_def) lemma Nmul1[simp]: "isnormNum c \ (1)\<^sub>N *\<^sub>N c = c" "isnormNum c \ c *\<^sub>N (1)\<^sub>N = c" - apply (simp_all add: Nmul_def Let_def split_def isnormNum_def) - apply (cases "fst c = 0", simp_all, cases c, simp_all)+ - done + by (simp add: Nmul_def Let_def split_def isnormNum_def, metis div_0 div_by_1 surjective_pairing)+ end