author | wenzelm |
Thu, 31 May 2007 12:06:31 +0200 | |
changeset 23146 | 0bc590051d95 |
parent 23145 | 5d8faadf3ecf |
child 23147 | a5db2f7d7654 |
--- a/src/HOL/Integ/Presburger.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1279 +0,0 @@ -(* Title: HOL/Integ/Presburger.thy - ID: $Id$ - Author: Amine Chaieb, Tobias Nipkow and Stefan Berghofer, TU Muenchen - -File containing necessary theorems for the proof -generation for Cooper Algorithm -*) - -header {* Presburger Arithmetic: Cooper's Algorithm *} - -theory Presburger -imports NatSimprocs "../SetInterval" -uses - ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") - ("reflected_presburger.ML") ("reflected_cooper.ML") ("presburger.ML") -begin - -text {* Theorem for unitifying the coeffitients of @{text x} in an existential formula*} - -theorem unity_coeff_ex: "(\<exists>x::int. P (l * x)) = (\<exists>x. l dvd (1*x+0) \<and> P x)" - apply (rule iffI) - apply (erule exE) - apply (rule_tac x = "l * x" in exI) - apply simp - apply (erule exE) - apply (erule conjE) - apply (erule dvdE) - apply (rule_tac x = k in exI) - apply simp - done - -lemma uminus_dvd_conv: "(d dvd (t::int)) = (-d dvd t)" -apply(unfold dvd_def) -apply(rule iffI) -apply(clarsimp) -apply(rename_tac k) -apply(rule_tac x = "-k" in exI) -apply simp -apply(clarsimp) -apply(rename_tac k) -apply(rule_tac x = "-k" in exI) -apply simp -done - -lemma uminus_dvd_conv': "(d dvd (t::int)) = (d dvd -t)" -apply(unfold dvd_def) -apply(rule iffI) -apply(clarsimp) -apply(rule_tac x = "-k" in exI) -apply simp -apply(clarsimp) -apply(rule_tac x = "-k" in exI) -apply simp -done - - - -text {*Theorems for the combination of proofs of the equality of @{text P} and @{text P_m} for integers @{text x} less than some integer @{text z}.*} - -theorem eq_minf_conjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> - \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> - \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))" - apply (erule exE)+ - apply (rule_tac x = "min z1 z2" in exI) - apply simp - done - - -theorem eq_minf_disjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> - \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> - \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))" - - apply (erule exE)+ - apply (rule_tac x = "min z1 z2" in exI) - apply simp - done - - -text {*Theorems for the combination of proofs of the equality of @{text P} and @{text P_m} for integers @{text x} greather than some integer @{text z}.*} - -theorem eq_pinf_conjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> - \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> - \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))" - apply (erule exE)+ - apply (rule_tac x = "max z1 z2" in exI) - apply simp - done - - -theorem eq_pinf_disjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> - \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> - \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))" - apply (erule exE)+ - apply (rule_tac x = "max z1 z2" in exI) - apply simp - done - -text {* - \medskip Theorems for the combination of proofs of the modulo @{text - D} property for @{text "P plusinfinity"} - - FIXME: This is THE SAME theorem as for the @{text minusinf} version, - but with @{text "+k.."} instead of @{text "-k.."} In the future - replace these both with only one. *} - -theorem modd_pinf_conjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow> - \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow> - \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x+k*d) \<and> B (x+k*d))" - by simp - -theorem modd_pinf_disjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow> - \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow> - \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x+k*d) \<or> B (x+k*d))" - by simp - -text {* - This is one of the cases where the simplifed formula is prooved to - habe some property (in relation to @{text P_m}) but we need to prove - the property for the original formula (@{text P_m}) - - FIXME: This is exaclty the same thm as for @{text minusinf}. *} - -lemma pinf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x)) ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) " - by blast - - -text {* - \medskip Theorems for the combination of proofs of the modulo @{text D} - property for @{text "P minusinfinity"} *} - -theorem modd_minf_conjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow> - \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow> - \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x-k*d) \<and> B (x-k*d))" - by simp - -theorem modd_minf_disjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow> - \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow> - \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x-k*d) \<or> B (x-k*d))" - by simp - -text {* - This is one of the cases where the simplifed formula is prooved to - have some property (in relation to @{text P_m}) but we need to - prove the property for the original formula (@{text P_m}). *} - -lemma minf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x)) ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) " - by blast - -text {* - Theorem needed for proving at runtime divide properties using the - arithmetic tactic (which knows only about modulo = 0). *} - -lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))" - by(simp add:dvd_def zmod_eq_0_iff) - -text {* - \medskip Theorems used for the combination of proof for the - backwards direction of Cooper's Theorem. They rely exclusively on - Predicate calculus.*} - -lemma not_ast_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P1(x) --> P1(x + d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<or> P2(x)) --> (P1(x + d) \<or> P2(x + d))) " - by blast - - -lemma not_ast_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a- j)) --> P1(x) --> P1(x + d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<and> P2(x)) --> (P1(x + d) -\<and> P2(x + d))) " - by blast - -lemma not_ast_p_Q_elim: " -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->P(x) --> P(x + d)) -==> ( P = Q ) -==> (ALL x. ~(EX (j::int) : {1..d}. EX (a::int) : A. P(a - j)) -->P(x) --> P(x + d))" - by blast - -text {* - \medskip Theorems used for the combination of proof for the - backwards direction of Cooper's Theorem. They rely exclusively on - Predicate calculus.*} - -lemma not_bst_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<or> P2(x)) --> (P1(x - d) -\<or> P2(x-d))) " - by blast - -lemma not_bst_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d)) -==> -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<and> P2(x)) --> (P1(x - d) -\<and> P2(x-d))) " - by blast - -lemma not_bst_p_Q_elim: " -(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->P(x) --> P(x - d)) -==> ( P = Q ) -==> (ALL x. ~(EX (j::int) : {1..d}. EX (b::int) : B. P(b+j)) -->P(x) --> P(x - d))" - by blast - -text {* \medskip This is the first direction of Cooper's Theorem. *} -lemma cooper_thm: "(R --> (EX x::int. P x)) ==> (Q -->(EX x::int. P x )) ==> ((R|Q) --> (EX x::int. P x )) " - by blast - -text {* - \medskip The full Cooper's Theorem in its equivalence Form. Given - the premises it is trivial too, it relies exclusively on prediacte calculus.*} -lemma cooper_eq_thm: "(R --> (EX x::int. P x)) ==> (Q -->(EX x::int. P x )) ==> ((~Q) ---> (EX x::int. P x ) --> R) ==> (EX x::int. P x) = R|Q " - by blast - -text {* - \medskip Some of the atomic theorems generated each time the atom - does not depend on @{text x}, they are trivial.*} - -lemma fm_eq_minf: "EX z::int. ALL x. x < z --> (P = P) " - by blast - -lemma fm_modd_minf: "ALL (x::int). ALL (k::int). (P = P)" - by blast - -lemma not_bst_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> fm --> fm" - by blast - -lemma fm_eq_pinf: "EX z::int. ALL x. z < x --> (P = P) " - by blast - -text {* The next two thms are the same as the @{text minusinf} version. *} - -lemma fm_modd_pinf: "ALL (x::int). ALL (k::int). (P = P)" - by blast - -lemma not_ast_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> fm --> fm" - by blast - -text {* Theorems to be deleted from simpset when proving simplified formulaes. *} - -lemma P_eqtrue: "(P=True) = P" - by iprover - -lemma P_eqfalse: "(P=False) = (~P)" - by iprover - -text {* - \medskip Theorems for the generation of the bachwards direction of - Cooper's Theorem. - - These are the 6 interesting atomic cases which have to be proved relying on the - properties of B-set and the arithmetic and contradiction proofs. *} - -lemma not_bst_p_lt: "0 < (d::int) ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ( 0 < -x + a) --> (0 < -(x - d) + a )" - by arith - -lemma not_bst_p_gt: "\<lbrakk> (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 < (x) + a) --> ( 0 < (x - d) + a)" -apply clarsimp -apply(rule ccontr) -apply(drule_tac x = "x+a" in bspec) -apply(simp add:atLeastAtMost_iff) -apply(drule_tac x = "-a" in bspec) -apply assumption -apply(simp) -done - -lemma not_bst_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a - 1 \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 = x + a) --> (0 = (x - d) + a )" -apply clarsimp -apply(subgoal_tac "x = -a") - prefer 2 apply arith -apply(drule_tac x = "1" in bspec) -apply(simp add:atLeastAtMost_iff) -apply(drule_tac x = "-a- 1" in bspec) -apply assumption -apply(simp) -done - - -lemma not_bst_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(0 = x + a) --> ~(0 = (x - d) + a)" -apply clarsimp -apply(subgoal_tac "x = -a+d") - prefer 2 apply arith -apply(drule_tac x = "d" in bspec) -apply(simp add:atLeastAtMost_iff) -apply(drule_tac x = "-a" in bspec) -apply assumption -apply(simp) -done - - -lemma not_bst_p_dvd: "(d1::int) dvd d ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> d1 dvd (x + a) --> d1 dvd ((x - d) + a )" -apply(clarsimp simp add:dvd_def) -apply(rename_tac m) -apply(rule_tac x = "m - k" in exI) -apply(simp add:int_distrib) -done - -lemma not_bst_p_ndvd: "(d1::int) dvd d ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(d1 dvd (x + a)) --> ~(d1 dvd ((x - d) + a ))" -apply(clarsimp simp add:dvd_def) -apply(rename_tac m) -apply(erule_tac x = "m + k" in allE) -apply(simp add:int_distrib) -done - -text {* - \medskip Theorems for the generation of the bachwards direction of - Cooper's Theorem. - - These are the 6 interesting atomic cases which have to be proved - relying on the properties of A-set ant the arithmetic and - contradiction proofs. *} - -lemma not_ast_p_gt: "0 < (d::int) ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ( 0 < x + t) --> (0 < (x + d) + t )" - by arith - -lemma not_ast_p_lt: "\<lbrakk>0 < d ;(t::int) \<in> A \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 < -x + t) --> ( 0 < -(x + d) + t)" - apply clarsimp - apply (rule ccontr) - apply (drule_tac x = "t-x" in bspec) - apply simp - apply (drule_tac x = "t" in bspec) - apply assumption - apply simp - done - -lemma not_ast_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t + 1 \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 = x + t) --> (0 = (x + d) + t )" - apply clarsimp - apply (drule_tac x="1" in bspec) - apply simp - apply (drule_tac x="- t + 1" in bspec) - apply assumption - apply(subgoal_tac "x = -t") - prefer 2 apply arith - apply simp - done - -lemma not_ast_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t \<rbrakk> \<Longrightarrow> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(0 = x + t) --> ~(0 = (x + d) + t)" - apply clarsimp - apply (subgoal_tac "x = -t-d") - prefer 2 apply arith - apply (drule_tac x = "d" in bspec) - apply simp - apply (drule_tac x = "-t" in bspec) - apply assumption - apply simp - done - -lemma not_ast_p_dvd: "(d1::int) dvd d ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> d1 dvd (x + t) --> d1 dvd ((x + d) + t )" - apply(clarsimp simp add:dvd_def) - apply(rename_tac m) - apply(rule_tac x = "m + k" in exI) - apply(simp add:int_distrib) - done - -lemma not_ast_p_ndvd: "(d1::int) dvd d ==> - ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(d1 dvd (x + t)) --> ~(d1 dvd ((x + d) + t ))" - apply(clarsimp simp add:dvd_def) - apply(rename_tac m) - apply(erule_tac x = "m - k" in allE) - apply(simp add:int_distrib) - done - -text {* - \medskip These are the atomic cases for the proof generation for the - modulo @{text D} property for @{text "P plusinfinity"} - - They are fully based on arithmetics. *} - -lemma dvd_modd_pinf: "((d::int) dvd d1) ==> - (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x+k*d1 + t))))" - apply(clarsimp simp add:dvd_def) - apply(rule iffI) - apply(clarsimp) - apply(rename_tac n m) - apply(rule_tac x = "m + n*k" in exI) - apply(simp add:int_distrib) - apply(clarsimp) - apply(rename_tac n m) - apply(rule_tac x = "m - n*k" in exI) - apply(simp add:int_distrib mult_ac) - done - -lemma not_dvd_modd_pinf: "((d::int) dvd d1) ==> - (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x+k*d1 + t))))" - apply(clarsimp simp add:dvd_def) - apply(rule iffI) - apply(clarsimp) - apply(rename_tac n m) - apply(erule_tac x = "m - n*k" in allE) - apply(simp add:int_distrib mult_ac) - apply(clarsimp) - apply(rename_tac n m) - apply(erule_tac x = "m + n*k" in allE) - apply(simp add:int_distrib mult_ac) - done - -text {* - \medskip These are the atomic cases for the proof generation for the - equivalence of @{text P} and @{text "P plusinfinity"} for integers - @{text x} greater than some integer @{text z}. - - They are fully based on arithmetics. *} - -lemma eq_eq_pinf: "EX z::int. ALL x. z < x --> (( 0 = x +t ) = False )" - apply(rule_tac x = "-t" in exI) - apply simp - done - -lemma neq_eq_pinf: "EX z::int. ALL x. z < x --> ((~( 0 = x +t )) = True )" - apply(rule_tac x = "-t" in exI) - apply simp - done - -lemma le_eq_pinf: "EX z::int. ALL x. z < x --> ( 0 < x +t = True )" - apply(rule_tac x = "-t" in exI) - apply simp - done - -lemma len_eq_pinf: "EX z::int. ALL x. z < x --> (0 < -x +t = False )" - apply(rule_tac x = "t" in exI) - apply simp - done - -lemma dvd_eq_pinf: "EX z::int. ALL x. z < x --> ((d dvd (x + t)) = (d dvd (x + t))) " - by simp - -lemma not_dvd_eq_pinf: "EX z::int. ALL x. z < x --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) " - by simp - -text {* - \medskip These are the atomic cases for the proof generation for the - modulo @{text D} property for @{text "P minusinfinity"}. - - They are fully based on arithmetics. *} - -lemma dvd_modd_minf: "((d::int) dvd d1) ==> - (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x-k*d1 + t))))" -apply(clarsimp simp add:dvd_def) -apply(rule iffI) -apply(clarsimp) -apply(rename_tac n m) -apply(rule_tac x = "m - n*k" in exI) -apply(simp add:int_distrib) -apply(clarsimp) -apply(rename_tac n m) -apply(rule_tac x = "m + n*k" in exI) -apply(simp add:int_distrib mult_ac) -done - - -lemma not_dvd_modd_minf: "((d::int) dvd d1) ==> - (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x-k*d1 + t))))" -apply(clarsimp simp add:dvd_def) -apply(rule iffI) -apply(clarsimp) -apply(rename_tac n m) -apply(erule_tac x = "m + n*k" in allE) -apply(simp add:int_distrib mult_ac) -apply(clarsimp) -apply(rename_tac n m) -apply(erule_tac x = "m - n*k" in allE) -apply(simp add:int_distrib mult_ac) -done - -text {* - \medskip These are the atomic cases for the proof generation for the - equivalence of @{text P} and @{text "P minusinfinity"} for integers - @{text x} less than some integer @{text z}. - - They are fully based on arithmetics. *} - -lemma eq_eq_minf: "EX z::int. ALL x. x < z --> (( 0 = x +t ) = False )" -apply(rule_tac x = "-t" in exI) -apply simp -done - -lemma neq_eq_minf: "EX z::int. ALL x. x < z --> ((~( 0 = x +t )) = True )" -apply(rule_tac x = "-t" in exI) -apply simp -done - -lemma le_eq_minf: "EX z::int. ALL x. x < z --> ( 0 < x +t = False )" -apply(rule_tac x = "-t" in exI) -apply simp -done - - -lemma len_eq_minf: "EX z::int. ALL x. x < z --> (0 < -x +t = True )" -apply(rule_tac x = "t" in exI) -apply simp -done - -lemma dvd_eq_minf: "EX z::int. ALL x. x < z --> ((d dvd (x + t)) = (d dvd (x + t))) " - by simp - -lemma not_dvd_eq_minf: "EX z::int. ALL x. x < z --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) " - by simp - -text {* - \medskip This Theorem combines whithnesses about @{text "P - minusinfinity"} to show one component of the equivalence proof for - Cooper's Theorem. - - FIXME: remove once they are part of the distribution. *} - -theorem int_ge_induct[consumes 1,case_names base step]: - assumes ge: "k \<le> (i::int)" and - base: "P(k)" and - step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)" - shows "P i" -proof - - { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k <= i \<Longrightarrow> P i" - proof (induct n) - case 0 - hence "i = k" by arith - thus "P i" using base by simp - next - case (Suc n) - hence "n = nat((i - 1) - k)" by arith - moreover - have ki1: "k \<le> i - 1" using Suc.prems by arith - ultimately - have "P(i - 1)" by(rule Suc.hyps) - from step[OF ki1 this] show ?case by simp - qed - } - from this ge show ?thesis by fast -qed - -theorem int_gr_induct[consumes 1,case_names base step]: - assumes gr: "k < (i::int)" and - base: "P(k+1)" and - step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)" - shows "P i" -apply(rule int_ge_induct[of "k + 1"]) - using gr apply arith - apply(rule base) -apply(rule step) - apply simp+ -done - -lemma decr_lemma: "0 < (d::int) \<Longrightarrow> x - (abs(x-z)+1) * d < z" -apply(induct rule: int_gr_induct) - apply simp -apply (simp add:int_distrib) -done - -lemma incr_lemma: "0 < (d::int) \<Longrightarrow> z < x + (abs(x-z)+1) * d" -apply(induct rule: int_gr_induct) - apply simp -apply (simp add:int_distrib) -done - -lemma minusinfinity: - assumes "0 < d" and - P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and - ePeqP1: "EX z::int. ALL x. x < z \<longrightarrow> (P x = P1 x)" - shows "(EX x. P1 x) \<longrightarrow> (EX x. P x)" -proof - assume eP1: "EX x. P1 x" - then obtain x where P1: "P1 x" .. - from ePeqP1 obtain z where P1eqP: "ALL x. x < z \<longrightarrow> (P x = P1 x)" .. - let ?w = "x - (abs(x-z)+1) * d" - show "EX x. P x" - proof - have w: "?w < z" by(rule decr_lemma) - have "P1 x = P1 ?w" using P1eqP1 by blast - also have "\<dots> = P(?w)" using w P1eqP by blast - finally show "P ?w" using P1 by blast - qed -qed - -text {* - \medskip This Theorem combines whithnesses about @{text "P - minusinfinity"} to show one component of the equivalence proof for - Cooper's Theorem. *} - -lemma plusinfinity: - assumes "0 < d" and - P1eqP1: "ALL (x::int) (k::int). P1 x = P1 (x + k * d)" and - ePeqP1: "EX z::int. ALL x. z < x --> (P x = P1 x)" - shows "(EX x::int. P1 x) --> (EX x::int. P x)" -proof - assume eP1: "EX x. P1 x" - then obtain x where P1: "P1 x" .. - from ePeqP1 obtain z where P1eqP: "ALL x. z < x \<longrightarrow> (P x = P1 x)" .. - let ?w = "x + (abs(x-z)+1) * d" - show "EX x. P x" - proof - have w: "z < ?w" by(rule incr_lemma) - have "P1 x = P1 ?w" using P1eqP1 by blast - also have "\<dots> = P(?w)" using w P1eqP by blast - finally show "P ?w" using P1 by blast - qed -qed - -text {* - \medskip Theorem for periodic function on discrete sets. *} - -lemma minf_vee: - assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)" - shows "(EX x. P x) = (EX j : {1..d}. P j)" - (is "?LHS = ?RHS") -proof - assume ?LHS - then obtain x where P: "P x" .. - have "x mod d = x - (x div d)*d" - by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) - hence Pmod: "P x = P(x mod d)" using modd by simp - show ?RHS - proof (cases) - assume "x mod d = 0" - hence "P 0" using P Pmod by simp - moreover have "P 0 = P(0 - (-1)*d)" using modd by blast - ultimately have "P d" by simp - moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff) - ultimately show ?RHS .. - next - assume not0: "x mod d \<noteq> 0" - have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) - moreover have "x mod d : {1..d}" - proof - - have "0 \<le> x mod d" by(rule pos_mod_sign) - moreover have "x mod d < d" by(rule pos_mod_bound) - ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) - qed - ultimately show ?RHS .. - qed -next - assume ?RHS thus ?LHS by blast -qed - -text {* - \medskip Theorem for periodic function on discrete sets. *} - -lemma pinf_vee: - assumes dpos: "0 < (d::int)" and modd: "ALL (x::int) (k::int). P x = P (x+k*d)" - shows "(EX x::int. P x) = (EX (j::int) : {1..d} . P j)" - (is "?LHS = ?RHS") -proof - assume ?LHS - then obtain x where P: "P x" .. - have "x mod d = x + (-(x div d))*d" - by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) - hence Pmod: "P x = P(x mod d)" using modd by (simp only:) - show ?RHS - proof (cases) - assume "x mod d = 0" - hence "P 0" using P Pmod by simp - moreover have "P 0 = P(0 + 1*d)" using modd by blast - ultimately have "P d" by simp - moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff) - ultimately show ?RHS .. - next - assume not0: "x mod d \<noteq> 0" - have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) - moreover have "x mod d : {1..d}" - proof - - have "0 \<le> x mod d" by(rule pos_mod_sign) - moreover have "x mod d < d" by(rule pos_mod_bound) - ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) - qed - ultimately show ?RHS .. - qed -next - assume ?RHS thus ?LHS by blast -qed - -lemma decr_mult_lemma: - assumes dpos: "(0::int) < d" and - minus: "ALL x::int. P x \<longrightarrow> P(x - d)" and - knneg: "0 <= k" - shows "ALL x. P x \<longrightarrow> P(x - k*d)" -using knneg -proof (induct rule:int_ge_induct) - case base thus ?case by simp -next - case (step i) - show ?case - proof - fix x - have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast - also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)" - using minus[THEN spec, of "x - i * d"] - by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric]) - ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast - qed -qed - -lemma incr_mult_lemma: - assumes dpos: "(0::int) < d" and - plus: "ALL x::int. P x \<longrightarrow> P(x + d)" and - knneg: "0 <= k" - shows "ALL x. P x \<longrightarrow> P(x + k*d)" -using knneg -proof (induct rule:int_ge_induct) - case base thus ?case by simp -next - case (step i) - show ?case - proof - fix x - have "P x \<longrightarrow> P (x + i * d)" using step.hyps by blast - also have "\<dots> \<longrightarrow> P(x + (i + 1) * d)" - using plus[THEN spec, of "x + i * d"] - by (simp add:int_distrib zadd_ac) - ultimately show "P x \<longrightarrow> P(x + (i + 1) * d)" by blast - qed -qed - -lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x)) -==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) -==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) -==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" -apply(rule iffI) -prefer 2 -apply(drule minusinfinity) -apply assumption+ -apply(fastsimp) -apply clarsimp -apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)") -apply(frule_tac x = x and z=z in decr_lemma) -apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)") -prefer 2 -apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") -prefer 2 apply arith - apply fastsimp -apply(drule (1) minf_vee) -apply blast -apply(blast dest:decr_mult_lemma) -done - -text {* Cooper Theorem, plus infinity version. *} -lemma cppi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. z < x --> (P x = P1 x)) -==> ALL x.~(EX (j::int) : {1..D}. EX (a::int) : A. P(a - j)) --> P (x) --> P (x + D) -==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x+k*D)))) -==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)))" - apply(rule iffI) - prefer 2 - apply(drule plusinfinity) - apply assumption+ - apply(fastsimp) - apply clarsimp - apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x + k*D)") - apply(frule_tac x = x and z=z in incr_lemma) - apply(subgoal_tac "P1(x + (\<bar>x - z\<bar> + 1) * D)") - prefer 2 - apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") - prefer 2 apply arith - apply fastsimp - apply(drule (1) pinf_vee) - apply blast - apply(blast dest:incr_mult_lemma) - done - - -text {* - \bigskip Theorems for the quantifier elminination Functions. *} - -lemma qe_ex_conj: "(EX (x::int). A x) = R - ==> (EX (x::int). P x) = (Q & (EX x::int. A x)) - ==> (EX (x::int). P x) = (Q & R)" -by blast - -lemma qe_ex_nconj: "(EX (x::int). P x) = (True & Q) - ==> (EX (x::int). P x) = Q" -by blast - -lemma qe_conjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 & Q1) = (P2 & Q2)" -by blast - -lemma qe_disjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 | Q1) = (P2 | Q2)" -by blast - -lemma qe_impI: "P1 = P2 ==> Q1 = Q2 ==> (P1 --> Q1) = (P2 --> Q2)" -by blast - -lemma qe_eqI: "P1 = P2 ==> Q1 = Q2 ==> (P1 = Q1) = (P2 = Q2)" -by blast - -lemma qe_Not: "P = Q ==> (~P) = (~Q)" -by blast - -lemma qe_ALL: "(EX x. ~P x) = R ==> (ALL x. P x) = (~R)" -by blast - -text {* \bigskip Theorems for proving NNF *} - -lemma nnf_im: "((~P) = P1) ==> (Q=Q1) ==> ((P --> Q) = (P1 | Q1))" -by blast - -lemma nnf_eq: "((P & Q) = (P1 & Q1)) ==> (((~P) & (~Q)) = (P2 & Q2)) ==> ((P = Q) = ((P1 & Q1)|(P2 & Q2)))" -by blast - -lemma nnf_nn: "(P = Q) ==> ((~~P) = Q)" - by blast -lemma nnf_ncj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P & Q)) = (P1 | Q1))" -by blast - -lemma nnf_ndj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P | Q)) = (P1 & Q1))" -by blast -lemma nnf_nim: "(P = P1) ==> ((~Q) = Q1) ==> ((~(P --> Q)) = (P1 & Q1))" -by blast -lemma nnf_neq: "((P & (~Q)) = (P1 & Q1)) ==> (((~P) & Q) = (P2 & Q2)) ==> ((~(P = Q)) = ((P1 & Q1)|(P2 & Q2)))" -by blast -lemma nnf_sdj: "((A & (~B)) = (A1 & B1)) ==> ((C & (~D)) = (C1 & D1)) ==> (A = (~C)) ==> ((~((A & B) | (C & D))) = ((A1 & B1) | (C1 & D1)))" -by blast - - -lemma qe_exI2: "A = B ==> (EX (x::int). A(x)) = (EX (x::int). B(x))" - by simp - -lemma qe_exI: "(!!x::int. A x = B x) ==> (EX (x::int). A(x)) = (EX (x::int). B(x))" - by iprover - -lemma qe_ALLI: "(!!x::int. A x = B x) ==> (ALL (x::int). A(x)) = (ALL (x::int). B(x))" - by iprover - -lemma cp_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j))) -==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j))) " -by blast - -lemma cppi_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j))) -==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j))) " -by blast - - -lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})" -apply(simp add:atLeastAtMost_def atLeast_def atMost_def) -apply(fastsimp) -done - -text {* \bigskip Theorems required for the @{text adjustcoeffitienteq} *} - -lemma ac_dvd_eq: assumes not0: "0 ~= (k::int)" -shows "((m::int) dvd (c*n+t)) = (k*m dvd ((k*c)*n+(k*t)))" (is "?P = ?Q") -proof - assume ?P - thus ?Q - apply(simp add:dvd_def) - apply clarify - apply(rename_tac d) - apply(drule_tac f = "op * k" in arg_cong) - apply(simp only:int_distrib) - apply(rule_tac x = "d" in exI) - apply(simp only:mult_ac) - done -next - assume ?Q - then obtain d where "k * c * n + k * t = (k*m)*d" by(fastsimp simp:dvd_def) - hence "(c * n + t) * k = (m*d) * k" by(simp add:int_distrib mult_ac) - hence "((c * n + t) * k) div k = ((m*d) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"]) - hence "c*n+t = m*d" by(simp add: zdiv_zmult_self1[OF not0[symmetric]]) - thus ?P by(simp add:dvd_def) -qed - -lemma ac_lt_eq: assumes gr0: "0 < (k::int)" -shows "((m::int) < (c*n+t)) = (k*m <((k*c)*n+(k*t)))" (is "?P = ?Q") -proof - assume P: ?P - show ?Q using zmult_zless_mono2[OF P gr0] by(simp add: int_distrib mult_ac) -next - assume ?Q - hence "0 < k*(c*n + t - m)" by(simp add: int_distrib mult_ac) - with gr0 have "0 < (c*n + t - m)" by(simp add: zero_less_mult_iff) - thus ?P by(simp) -qed - -lemma ac_eq_eq : assumes not0: "0 ~= (k::int)" shows "((m::int) = (c*n+t)) = (k*m =((k*c)*n+(k*t)) )" (is "?P = ?Q") -proof - assume ?P - thus ?Q - apply(drule_tac f = "op * k" in arg_cong) - apply(simp only:int_distrib) - done -next - assume ?Q - hence "m * k = (c*n + t) * k" by(simp add:int_distrib mult_ac) - hence "((m) * k) div k = ((c*n + t) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"]) - thus ?P by(simp add: zdiv_zmult_self1[OF not0[symmetric]]) -qed - -lemma ac_pi_eq: assumes gr0: "0 < (k::int)" shows "(~((0::int) < (c*n + t))) = (0 < ((-k)*c)*n + ((-k)*t + k))" -proof - - have "(~ (0::int) < (c*n + t)) = (0<1-(c*n + t))" by arith - also have "(1-(c*n + t)) = (-1*c)*n + (-t+1)" by(simp add: int_distrib mult_ac) - also have "0<(-1*c)*n + (-t+1) = (0 < (k*(-1*c)*n) + (k*(-t+1)))" by(rule ac_lt_eq[of _ 0,OF gr0,simplified]) - also have "(k*(-1*c)*n) + (k*(-t+1)) = ((-k)*c)*n + ((-k)*t + k)" by(simp add: int_distrib mult_ac) - finally show ?thesis . -qed - -lemma binminus_uminus_conv: "(a::int) - b = a + (-b)" -by arith - -lemma linearize_dvd: "(t::int) = t1 ==> (d dvd t) = (d dvd t1)" -by simp - -lemma lf_lt: "(l::int) = ll ==> (r::int) = lr ==> (l < r) =(ll < lr)" -by simp - -lemma lf_eq: "(l::int) = ll ==> (r::int) = lr ==> (l = r) =(ll = lr)" -by simp - -lemma lf_dvd: "(l::int) = ll ==> (r::int) = lr ==> (l dvd r) =(ll dvd lr)" -by simp - -text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*} - -theorem all_nat: "(\<forall>x::nat. P x) = (\<forall>x::int. 0 <= x \<longrightarrow> P (nat x))" - by (simp split add: split_nat) - - -theorem zdiff_int_split: "P (int (x - y)) = - ((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))" - apply (case_tac "y \<le> x") - apply (simp_all add: zdiff_int) - done - - -theorem number_of1: "(0::int) <= number_of n \<Longrightarrow> (0::int) <= number_of (n BIT b)" - by simp - -theorem number_of2: "(0::int) <= Numeral0" by simp - -theorem Suc_plus1: "Suc n = n + 1" by simp - -text {* - \medskip Specific instances of congruence rules, to prevent - simplifier from looping. *} - -theorem imp_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::int) \<longrightarrow> P) = (0 <= x \<longrightarrow> P')" - by simp - -theorem conj_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::int) \<and> P) = (0 <= x \<and> P')" - by (simp cong: conj_cong) - - (* Theorems used in presburger.ML for the computation simpset*) - (* FIXME: They are present in Float.thy, so may be Float.thy should be lightened.*) - -lemma lift_bool: "x \<Longrightarrow> x=True" - by simp - -lemma nlift_bool: "~x \<Longrightarrow> x=False" - by simp - -lemma not_false_eq_true: "(~ False) = True" by simp - -lemma not_true_eq_false: "(~ True) = False" by simp - - -lemma int_eq_number_of_eq: - "(((number_of v)::int) = (number_of w)) = iszero ((number_of (v + (uminus w)))::int)" - by simp -lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" - by (simp only: iszero_number_of_Pls) - -lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" - by simp - -lemma int_iszero_number_of_0: "iszero ((number_of (w BIT bit.B0))::int) = iszero ((number_of w)::int)" - by simp - -lemma int_iszero_number_of_1: "\<not> iszero ((number_of (w BIT bit.B1))::int)" - by simp - -lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)" - by simp - -lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))" - by simp - -lemma int_neg_number_of_Min: "neg (-1::int)" - by simp - -lemma int_neg_number_of_BIT: "neg ((number_of (w BIT x))::int) = neg ((number_of w)::int)" - by simp - -lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))" - by simp -lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)" - by simp - -lemma int_number_of_diff_sym: - "((number_of v)::int) - number_of w = number_of (v + (uminus w))" - by simp - -lemma int_number_of_mult_sym: - "((number_of v)::int) * number_of w = number_of (v * w)" - by simp - -lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)" - by simp -lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" - by simp - -lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" - by simp - -lemma mult_left_one: "1 * a = (a::'a::semiring_1)" - by simp - -lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" - by simp - -lemma int_pow_0: "(a::int)^(Numeral0) = 1" - by simp - -lemma int_pow_1: "(a::int)^(Numeral1) = a" - by simp - -lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" - by simp - -lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" - by simp - -lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" - by simp - -lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" - by simp - -lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" - by simp - -lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" -proof - - have 1:"((-1)::nat) = 0" - by simp - show ?thesis by (simp add: 1) -qed - -use "cooper_dec.ML" -use "reflected_presburger.ML" -use "reflected_cooper.ML" -oracle - presburger_oracle ("term") = ReflectedCooper.presburger_oracle - -use "cooper_proof.ML" -use "qelim.ML" -use "presburger.ML" - -setup "Presburger.setup" - - -subsection {* Code generator setup *} - -text {* - Presburger arithmetic is convenient to prove some - of the following code lemmas on integer numerals: -*} - -lemma eq_Pls_Pls: - "Numeral.Pls = Numeral.Pls \<longleftrightarrow> True" by rule+ - -lemma eq_Pls_Min: - "Numeral.Pls = Numeral.Min \<longleftrightarrow> False" - unfolding Pls_def Min_def by auto - -lemma eq_Pls_Bit0: - "Numeral.Pls = Numeral.Bit k bit.B0 \<longleftrightarrow> Numeral.Pls = k" - unfolding Pls_def Bit_def bit.cases by auto - -lemma eq_Pls_Bit1: - "Numeral.Pls = Numeral.Bit k bit.B1 \<longleftrightarrow> False" - unfolding Pls_def Bit_def bit.cases by arith - -lemma eq_Min_Pls: - "Numeral.Min = Numeral.Pls \<longleftrightarrow> False" - unfolding Pls_def Min_def by auto - -lemma eq_Min_Min: - "Numeral.Min = Numeral.Min \<longleftrightarrow> True" by rule+ - -lemma eq_Min_Bit0: - "Numeral.Min = Numeral.Bit k bit.B0 \<longleftrightarrow> False" - unfolding Min_def Bit_def bit.cases by arith - -lemma eq_Min_Bit1: - "Numeral.Min = Numeral.Bit k bit.B1 \<longleftrightarrow> Numeral.Min = k" - unfolding Min_def Bit_def bit.cases by auto - -lemma eq_Bit0_Pls: - "Numeral.Bit k bit.B0 = Numeral.Pls \<longleftrightarrow> Numeral.Pls = k" - unfolding Pls_def Bit_def bit.cases by auto - -lemma eq_Bit1_Pls: - "Numeral.Bit k bit.B1 = Numeral.Pls \<longleftrightarrow> False" - unfolding Pls_def Bit_def bit.cases by arith - -lemma eq_Bit0_Min: - "Numeral.Bit k bit.B0 = Numeral.Min \<longleftrightarrow> False" - unfolding Min_def Bit_def bit.cases by arith - -lemma eq_Bit1_Min: - "(Numeral.Bit k bit.B1) = Numeral.Min \<longleftrightarrow> Numeral.Min = k" - unfolding Min_def Bit_def bit.cases by auto - -lemma eq_Bit_Bit: - "Numeral.Bit k1 v1 = Numeral.Bit k2 v2 \<longleftrightarrow> - v1 = v2 \<and> k1 = k2" - unfolding Bit_def - apply (cases v1) - apply (cases v2) - apply auto - apply arith - apply (cases v2) - apply auto - apply arith - apply (cases v2) - apply auto -done - -lemma eq_number_of: - "(number_of k \<Colon> int) = number_of l \<longleftrightarrow> k = l" - unfolding number_of_is_id .. - - -lemma less_eq_Pls_Pls: - "Numeral.Pls \<le> Numeral.Pls \<longleftrightarrow> True" by rule+ - -lemma less_eq_Pls_Min: - "Numeral.Pls \<le> Numeral.Min \<longleftrightarrow> False" - unfolding Pls_def Min_def by auto - -lemma less_eq_Pls_Bit: - "Numeral.Pls \<le> Numeral.Bit k v \<longleftrightarrow> Numeral.Pls \<le> k" - unfolding Pls_def Bit_def by (cases v) auto - -lemma less_eq_Min_Pls: - "Numeral.Min \<le> Numeral.Pls \<longleftrightarrow> True" - unfolding Pls_def Min_def by auto - -lemma less_eq_Min_Min: - "Numeral.Min \<le> Numeral.Min \<longleftrightarrow> True" by rule+ - -lemma less_eq_Min_Bit0: - "Numeral.Min \<le> Numeral.Bit k bit.B0 \<longleftrightarrow> Numeral.Min < k" - unfolding Min_def Bit_def by auto - -lemma less_eq_Min_Bit1: - "Numeral.Min \<le> Numeral.Bit k bit.B1 \<longleftrightarrow> Numeral.Min \<le> k" - unfolding Min_def Bit_def by auto - -lemma less_eq_Bit0_Pls: - "Numeral.Bit k bit.B0 \<le> Numeral.Pls \<longleftrightarrow> k \<le> Numeral.Pls" - unfolding Pls_def Bit_def by simp - -lemma less_eq_Bit1_Pls: - "Numeral.Bit k bit.B1 \<le> Numeral.Pls \<longleftrightarrow> k < Numeral.Pls" - unfolding Pls_def Bit_def by auto - -lemma less_eq_Bit_Min: - "Numeral.Bit k v \<le> Numeral.Min \<longleftrightarrow> k \<le> Numeral.Min" - unfolding Min_def Bit_def by (cases v) auto - -lemma less_eq_Bit0_Bit: - "Numeral.Bit k1 bit.B0 \<le> Numeral.Bit k2 v \<longleftrightarrow> k1 \<le> k2" - unfolding Bit_def bit.cases by (cases v) auto - -lemma less_eq_Bit_Bit1: - "Numeral.Bit k1 v \<le> Numeral.Bit k2 bit.B1 \<longleftrightarrow> k1 \<le> k2" - unfolding Bit_def bit.cases by (cases v) auto - -lemma less_eq_Bit1_Bit0: - "Numeral.Bit k1 bit.B1 \<le> Numeral.Bit k2 bit.B0 \<longleftrightarrow> k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_eq_number_of: - "(number_of k \<Colon> int) \<le> number_of l \<longleftrightarrow> k \<le> l" - unfolding number_of_is_id .. - - -lemma less_Pls_Pls: - "Numeral.Pls < Numeral.Pls \<longleftrightarrow> False" by auto - -lemma less_Pls_Min: - "Numeral.Pls < Numeral.Min \<longleftrightarrow> False" - unfolding Pls_def Min_def by auto - -lemma less_Pls_Bit0: - "Numeral.Pls < Numeral.Bit k bit.B0 \<longleftrightarrow> Numeral.Pls < k" - unfolding Pls_def Bit_def by auto - -lemma less_Pls_Bit1: - "Numeral.Pls < Numeral.Bit k bit.B1 \<longleftrightarrow> Numeral.Pls \<le> k" - unfolding Pls_def Bit_def by auto - -lemma less_Min_Pls: - "Numeral.Min < Numeral.Pls \<longleftrightarrow> True" - unfolding Pls_def Min_def by auto - -lemma less_Min_Min: - "Numeral.Min < Numeral.Min \<longleftrightarrow> False" by auto - -lemma less_Min_Bit: - "Numeral.Min < Numeral.Bit k v \<longleftrightarrow> Numeral.Min < k" - unfolding Min_def Bit_def by (auto split: bit.split) - -lemma less_Bit_Pls: - "Numeral.Bit k v < Numeral.Pls \<longleftrightarrow> k < Numeral.Pls" - unfolding Pls_def Bit_def by (auto split: bit.split) - -lemma less_Bit0_Min: - "Numeral.Bit k bit.B0 < Numeral.Min \<longleftrightarrow> k \<le> Numeral.Min" - unfolding Min_def Bit_def by auto - -lemma less_Bit1_Min: - "Numeral.Bit k bit.B1 < Numeral.Min \<longleftrightarrow> k < Numeral.Min" - unfolding Min_def Bit_def by auto - -lemma less_Bit_Bit0: - "Numeral.Bit k1 v < Numeral.Bit k2 bit.B0 \<longleftrightarrow> k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_Bit1_Bit: - "Numeral.Bit k1 bit.B1 < Numeral.Bit k2 v \<longleftrightarrow> k1 < k2" - unfolding Bit_def by (auto split: bit.split) - -lemma less_Bit0_Bit1: - "Numeral.Bit k1 bit.B0 < Numeral.Bit k2 bit.B1 \<longleftrightarrow> k1 \<le> k2" - unfolding Bit_def bit.cases by auto - -lemma less_number_of: - "(number_of k \<Colon> int) < number_of l \<longleftrightarrow> k < l" - unfolding number_of_is_id .. - - -lemmas pred_succ_numeral_code [code func] = - arith_simps(5-12) - -lemmas plus_numeral_code [code func] = - arith_simps(13-17) - arith_simps(26-27) - arith_extra_simps(1) [where 'a = int] - -lemmas minus_numeral_code [code func] = - arith_simps(18-21) - arith_extra_simps(2) [where 'a = int] - arith_extra_simps(5) [where 'a = int] - -lemmas times_numeral_code [code func] = - arith_simps(22-25) - arith_extra_simps(4) [where 'a = int] - -lemmas eq_numeral_code [code func] = - eq_Pls_Pls eq_Pls_Min eq_Pls_Bit0 eq_Pls_Bit1 - eq_Min_Pls eq_Min_Min eq_Min_Bit0 eq_Min_Bit1 - eq_Bit0_Pls eq_Bit1_Pls eq_Bit0_Min eq_Bit1_Min eq_Bit_Bit - eq_number_of - -lemmas less_eq_numeral_code [code func] = less_eq_Pls_Pls less_eq_Pls_Min less_eq_Pls_Bit - less_eq_Min_Pls less_eq_Min_Min less_eq_Min_Bit0 less_eq_Min_Bit1 - less_eq_Bit0_Pls less_eq_Bit1_Pls less_eq_Bit_Min less_eq_Bit0_Bit less_eq_Bit_Bit1 less_eq_Bit1_Bit0 - less_eq_number_of - -lemmas less_numeral_code [code func] = less_Pls_Pls less_Pls_Min less_Pls_Bit0 - less_Pls_Bit1 less_Min_Pls less_Min_Min less_Min_Bit less_Bit_Pls - less_Bit0_Min less_Bit1_Min less_Bit_Bit0 less_Bit1_Bit less_Bit0_Bit1 - less_number_of - -end
--- a/src/HOL/Integ/cooper_dec.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,920 +0,0 @@ -(* Title: HOL/Integ/cooper_dec.ML - ID: $Id$ - Author: Amine Chaieb and Tobias Nipkow, TU Muenchen - -File containing the implementation of Cooper Algorithm -decision procedure (intensively inspired from J.Harrison) -*) - - -signature COOPER_DEC = -sig - exception COOPER - val mk_number : IntInf.int -> term - val zero : term - val one : term - val dest_number : term -> IntInf.int - val is_number : term -> bool - val is_arith_rel : term -> bool - val linear_cmul : IntInf.int -> term -> term - val linear_add : string list -> term -> term -> term - val linear_sub : string list -> term -> term -> term - val linear_neg : term -> term - val lint : string list -> term -> term - val linform : string list -> term -> term - val formlcm : term -> term -> IntInf.int - val adjustcoeff : term -> IntInf.int -> term -> term - val unitycoeff : term -> term -> term - val divlcm : term -> term -> IntInf.int - val bset : term -> term -> term list - val aset : term -> term -> term list - val linrep : string list -> term -> term -> term -> term - val list_disj : term list -> term - val list_conj : term list -> term - val simpl : term -> term - val fv : term -> string list - val negate : term -> term - val operations : (string * (IntInf.int * IntInf.int -> bool)) list - val conjuncts : term -> term list - val disjuncts : term -> term list - val has_bound : term -> bool - val minusinf : term -> term -> term - val plusinf : term -> term -> term - val onatoms : (term -> term) -> term -> term - val evalc : term -> term - val cooper_w : string list -> term -> (term option * term) - val integer_qelim : Term.term -> Term.term -end; - -structure CooperDec : COOPER_DEC = -struct - -(* ========================================================================= *) -(* Cooper's algorithm for Presburger arithmetic. *) -(* ========================================================================= *) -exception COOPER; - - -(* ------------------------------------------------------------------------- *) -(* Lift operations up to numerals. *) -(* ------------------------------------------------------------------------- *) - -(*Assumption : The construction of atomar formulas in linearl arithmetic is based on -relation operations of Type : [IntInf.int,IntInf.int]---> bool *) - -(* ------------------------------------------------------------------------- *) - -(*Function is_arith_rel returns true if and only if the term is an atomar presburger -formula *) -fun is_arith_rel tm = case tm - of Const(p, Type ("fun", [Type ("IntDef.int", []), Type ("fun", [Type ("IntDef.int", []), - Type ("bool", [])])])) $ _ $_ => true - | _ => false; - -(*Function is_arith_rel returns true if and only if the term is an operation of the -form [int,int]---> int*) - -val mk_number = HOLogic.mk_number HOLogic.intT; -val zero = mk_number 0; -val one = mk_number 1; -fun dest_number t = let - val (T, n) = HOLogic.dest_number t - in if T = HOLogic.intT then n else error ("bad typ: " ^ Display.raw_string_of_typ T) end; -val is_number = can dest_number; - -(*maps a unary natural function on a term containing an natural number*) -fun numeral1 f n = mk_number (f (dest_number n)); - -(*maps a binary natural function on 2 term containing natural numbers*) -fun numeral2 f m n = mk_number (f (dest_number m) (dest_number n)); - -(* ------------------------------------------------------------------------- *) -(* Operations on canonical linear terms c1 * x1 + ... + cn * xn + k *) -(* *) -(* Note that we're quite strict: the ci must be present even if 1 *) -(* (but if 0 we expect the monomial to be omitted) and k must be there *) -(* even if it's zero. Thus, it's a constant iff not an addition term. *) -(* ------------------------------------------------------------------------- *) - - -fun linear_cmul n tm = if n = 0 then zero else let fun times n k = n*k in - ( case tm of - (Const(@{const_name HOL.plus},T) $ (Const (@{const_name HOL.times},T1 ) $c1 $ x1) $ rest) => - Const(@{const_name HOL.plus},T) $ ((Const(@{const_name HOL.times},T1) $ (numeral1 (times n) c1) $ x1)) $ (linear_cmul n rest) - |_ => numeral1 (times n) tm) - end ; - - - - -(* Whether the first of two items comes earlier in the list *) -fun earlier [] x y = false - |earlier (h::t) x y =if h = y then false - else if h = x then true - else earlier t x y ; - -fun earlierv vars (Bound i) (Bound j) = i < j - |earlierv vars (Bound _) _ = true - |earlierv vars _ (Bound _) = false - |earlierv vars (Free (x,_)) (Free (y,_)) = earlier vars x y; - - -fun linear_add vars tm1 tm2 = - let fun addwith x y = x + y in - (case (tm1,tm2) of - ((Const (@{const_name HOL.plus},T1) $ ( Const(@{const_name HOL.times},T2) $ c1 $ x1) $ rest1),(Const - (@{const_name HOL.plus},T3)$( Const(@{const_name HOL.times},T4) $ c2 $ x2) $ rest2)) => - if x1 = x2 then - let val c = (numeral2 (addwith) c1 c2) - in - if c = zero then (linear_add vars rest1 rest2) - else (Const(@{const_name HOL.plus},T1) $ (Const(@{const_name HOL.times},T2) $ c $ x1) $ (linear_add vars rest1 rest2)) - end - else - if earlierv vars x1 x2 then (Const(@{const_name HOL.plus},T1) $ - (Const(@{const_name HOL.times},T2)$ c1 $ x1) $ (linear_add vars rest1 tm2)) - else (Const(@{const_name HOL.plus},T1) $ (Const(@{const_name HOL.times},T2) $ c2 $ x2) $ (linear_add vars tm1 rest2)) - |((Const(@{const_name HOL.plus},T1) $ (Const(@{const_name HOL.times},T2) $ c1 $ x1) $ rest1) ,_) => - (Const(@{const_name HOL.plus},T1)$ (Const(@{const_name HOL.times},T2) $ c1 $ x1) $ (linear_add vars - rest1 tm2)) - |(_, (Const(@{const_name HOL.plus},T1) $(Const(@{const_name HOL.times},T2) $ c2 $ x2) $ rest2)) => - (Const(@{const_name HOL.plus},T1) $ (Const(@{const_name HOL.times},T2) $ c2 $ x2) $ (linear_add vars tm1 - rest2)) - | (_,_) => numeral2 (addwith) tm1 tm2) - - end; - -(*To obtain the unary - applyed on a formula*) - -fun linear_neg tm = linear_cmul (0 - 1) tm; - -(*Substraction of two terms *) - -fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); - - -(* ------------------------------------------------------------------------- *) -(* Linearize a term. *) -(* ------------------------------------------------------------------------- *) - -(* linearises a term from the point of view of Variable Free (x,T). -After this fuction the all expressions containig ths variable will have the form - c*Free(x,T) + t where c is a constant ant t is a Term which is not containing - Free(x,T)*) - -fun lint vars tm = if is_number tm then tm else case tm of - (Free (x,T)) => (HOLogic.mk_binop @{const_name HOL.plus} ((HOLogic.mk_binop @{const_name HOL.times} ((mk_number 1),Free (x,T))), zero)) - |(Bound i) => (Const(@{const_name HOL.plus},HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ - (Const(@{const_name HOL.times},HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ (mk_number 1) $ (Bound i)) $ zero) - |(Const(@{const_name HOL.uminus},_) $ t ) => (linear_neg (lint vars t)) - |(Const(@{const_name HOL.plus},_) $ s $ t) => (linear_add vars (lint vars s) (lint vars t)) - |(Const(@{const_name HOL.minus},_) $ s $ t) => (linear_sub vars (lint vars s) (lint vars t)) - |(Const (@{const_name HOL.times},_) $ s $ t) => - let val s' = lint vars s - val t' = lint vars t - in - if is_number s' then (linear_cmul (dest_number s') t') - else if is_number t' then (linear_cmul (dest_number t') s') - - else raise COOPER - end - |_ => raise COOPER; - - - -(* ------------------------------------------------------------------------- *) -(* Linearize the atoms in a formula, and eliminate non-strict inequalities. *) -(* ------------------------------------------------------------------------- *) - -fun mkatom vars p t = Const(p,HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ zero $ (lint vars t); - -fun linform vars (Const ("Divides.dvd",_) $ c $ t) = - if is_number c then - let val c' = (mk_number(abs(dest_number c))) - in (HOLogic.mk_binrel "Divides.dvd" (c,lint vars t)) - end - else (warning "Nonlinear term --- Non numeral leftside at dvd" - ;raise COOPER) - |linform vars (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ s $ t ) = (mkatom vars "op =" (Const (@{const_name HOL.minus},HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s) ) - |linform vars (Const(@{const_name Orderings.less},_)$ s $t ) = (mkatom vars @{const_name Orderings.less} (Const (@{const_name HOL.minus},HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s)) - |linform vars (Const("op >",_) $ s $ t ) = (mkatom vars @{const_name Orderings.less} (Const (@{const_name HOL.minus},HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ s $ t)) - |linform vars (Const(@{const_name Orderings.less_eq},_)$ s $ t ) = - (mkatom vars @{const_name Orderings.less} (Const (@{const_name HOL.minus},HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ (Const(@{const_name HOL.plus},HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $t $(mk_number 1)) $ s)) - |linform vars (Const("op >=",_)$ s $ t ) = - (mkatom vars @{const_name Orderings.less} (Const (@{const_name HOL.minus},HOLogic.intT --> HOLogic.intT --> - HOLogic.intT) $ (Const(@{const_name HOL.plus},HOLogic.intT --> HOLogic.intT --> - HOLogic.intT) $s $(mk_number 1)) $ t)) - - |linform vars fm = fm; - -(* ------------------------------------------------------------------------- *) -(* Post-NNF transformation eliminating negated inequalities. *) -(* ------------------------------------------------------------------------- *) - -fun posineq fm = case fm of - (Const ("Not",_)$(Const(@{const_name Orderings.less},_)$ c $ t)) => - (HOLogic.mk_binrel @{const_name Orderings.less} (zero , (linear_sub [] (mk_number 1) (linear_add [] c t ) ))) - | ( Const ("op &",_) $ p $ q) => HOLogic.mk_conj (posineq p,posineq q) - | ( Const ("op |",_) $ p $ q ) => HOLogic.mk_disj (posineq p,posineq q) - | _ => fm; - - -(* ------------------------------------------------------------------------- *) -(* Find the LCM of the coefficients of x. *) -(* ------------------------------------------------------------------------- *) -(*gcd calculates gcd (a,b) and helps lcm_num calculating lcm (a,b)*) - -(*BEWARE: replaces Library.gcd!! There is also Library.lcm!*) -fun gcd (a:IntInf.int) b = if a=0 then b else gcd (b mod a) a ; -fun lcm_num a b = (abs a*b) div (gcd (abs a) (abs b)); - -fun formlcm x fm = case fm of - (Const (p,_)$ _ $(Const (@{const_name HOL.plus}, _)$(Const (@{const_name HOL.times},_)$ c $ y ) $z ) ) => if - (is_arith_rel fm) andalso (x = y) then (abs(dest_number c)) else 1 - | ( Const ("Not", _) $p) => formlcm x p - | ( Const ("op &",_) $ p $ q) => lcm_num (formlcm x p) (formlcm x q) - | ( Const ("op |",_) $ p $ q )=> lcm_num (formlcm x p) (formlcm x q) - | _ => 1; - -(* ------------------------------------------------------------------------- *) -(* Adjust all coefficients of x in formula; fold in reduction to +/- 1. *) -(* ------------------------------------------------------------------------- *) - -fun adjustcoeff x l fm = - case fm of - (Const(p,_) $d $( Const (@{const_name HOL.plus}, _)$(Const (@{const_name HOL.times},_) $ - c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then - let val m = l div (dest_number c) - val n = (if p = @{const_name Orderings.less} then abs(m) else m) - val xtm = HOLogic.mk_binop @{const_name HOL.times} ((mk_number (m div n)), x) - in - (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop @{const_name HOL.plus} ( xtm ,( linear_cmul n z) )))) - end - else fm - |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeff x l p) - |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeff x l p) $(adjustcoeff x l q) - |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeff x l p)$ (adjustcoeff x l q) - |_ => fm; - -(* ------------------------------------------------------------------------- *) -(* Hence make coefficient of x one in existential formula. *) -(* ------------------------------------------------------------------------- *) - -fun unitycoeff x fm = - let val l = formlcm x fm - val fm' = adjustcoeff x l fm in - if l = 1 then fm' - else - let val xp = (HOLogic.mk_binop @{const_name HOL.plus} - ((HOLogic.mk_binop @{const_name HOL.times} ((mk_number 1), x )), zero)) - in - HOLogic.conj $(HOLogic.mk_binrel "Divides.dvd" ((mk_number l) , xp )) $ (adjustcoeff x l fm) - end - end; - -(* adjustcoeffeq l fm adjusts the coeffitients c_i of x overall in fm to l*) -(* Here l must be a multiple of all c_i otherwise the obtained formula is not equivalent*) -(* -fun adjustcoeffeq x l fm = - case fm of - (Const(p,_) $d $( Const (@{const_name HOL.plus}, _)$(Const (@{const_name HOL.times},_) $ - c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then - let val m = l div (dest_number c) - val n = (if p = @{const_name Orderings.less} then abs(m) else m) - val xtm = (HOLogic.mk_binop @{const_name HOL.times} ((mk_number ((m div n)*l) ), x)) - in (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop @{const_name HOL.plus} ( xtm ,( linear_cmul n z) )))) - end - else fm - |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeffeq x l p) - |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeffeq x l p) $(adjustcoeffeq x l q) - |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeffeq x l p)$ (adjustcoeffeq x l q) - |_ => fm; - - -*) - -(* ------------------------------------------------------------------------- *) -(* The "minus infinity" version. *) -(* ------------------------------------------------------------------------- *) - -fun minusinf x fm = case fm of - (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const - else fm - - |(Const(@{const_name Orderings.less},_) $ c $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z - )) => if (x = y) - then if (pm1 = one) andalso (c = zero) then HOLogic.false_const - else if (dest_number pm1 = ~1) andalso (c = zero) then HOLogic.true_const - else error "minusinf : term not in normal form!!!" - else fm - - |(Const ("Not", _) $ p) => HOLogic.Not $ (minusinf x p) - |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (minusinf x p) $ (minusinf x q) - |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (minusinf x p) $ (minusinf x q) - |_ => fm; - -(* ------------------------------------------------------------------------- *) -(* The "Plus infinity" version. *) -(* ------------------------------------------------------------------------- *) - -fun plusinf x fm = case fm of - (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const - else fm - - |(Const(@{const_name Orderings.less},_) $ c $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z - )) => if (x = y) - then if (pm1 = one) andalso (c = zero) then HOLogic.true_const - else if (dest_number pm1 = ~1) andalso (c = zero) then HOLogic.false_const - else error "plusinf : term not in normal form!!!" - else fm - - |(Const ("Not", _) $ p) => HOLogic.Not $ (plusinf x p) - |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (plusinf x p) $ (plusinf x q) - |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (plusinf x p) $ (plusinf x q) - |_ => fm; - -(* ------------------------------------------------------------------------- *) -(* The LCM of all the divisors that involve x. *) -(* ------------------------------------------------------------------------- *) - -fun divlcm x (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z ) ) = - if x = y then abs(dest_number d) else 1 - |divlcm x ( Const ("Not", _) $ p) = divlcm x p - |divlcm x ( Const ("op &",_) $ p $ q) = lcm_num (divlcm x p) (divlcm x q) - |divlcm x ( Const ("op |",_) $ p $ q ) = lcm_num (divlcm x p) (divlcm x q) - |divlcm x _ = 1; - -(* ------------------------------------------------------------------------- *) -(* Construct the B-set. *) -(* ------------------------------------------------------------------------- *) - -fun bset x fm = case fm of - (Const ("Not", _) $ p) => if (is_arith_rel p) then - (case p of - (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $c2 $y) $a ) ) - => if (is_arith_rel p) andalso (x= y) andalso (c2 = one) andalso (c1 = zero) - then [linear_neg a] - else bset x p - |_ =>[]) - - else bset x p - |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_neg(linear_add [] a (mk_number 1))] else [] - |(Const (@{const_name Orderings.less},_) $ c1$ (Const (@{const_name HOL.plus},_) $(Const (@{const_name HOL.times},_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_neg a] else [] - |(Const ("op &",_) $ p $ q) => (bset x p) union (bset x q) - |(Const ("op |",_) $ p $ q) => (bset x p) union (bset x q) - |_ => []; - -(* ------------------------------------------------------------------------- *) -(* Construct the A-set. *) -(* ------------------------------------------------------------------------- *) - -fun aset x fm = case fm of - (Const ("Not", _) $ p) => if (is_arith_rel p) then - (case p of - (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $c2 $y) $a ) ) - => if (x= y) andalso (c2 = one) andalso (c1 = zero) - then [linear_neg a] - else [] - |_ =>[]) - - else aset x p - |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_sub [] (mk_number 1) a] else [] - |(Const (@{const_name Orderings.less},_) $ c1$ (Const (@{const_name HOL.plus},_) $(Const (@{const_name HOL.times},_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = (mk_number (~1))) then [a] else [] - |(Const ("op &",_) $ p $ q) => (aset x p) union (aset x q) - |(Const ("op |",_) $ p $ q) => (aset x p) union (aset x q) - |_ => []; - - -(* ------------------------------------------------------------------------- *) -(* Replace top variable with another linear form, retaining canonicality. *) -(* ------------------------------------------------------------------------- *) - -fun linrep vars x t fm = case fm of - ((Const(p,_)$ d $ (Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$ c $ y) $ z))) => - if (x = y) andalso (is_arith_rel fm) - then - let val ct = linear_cmul (dest_number c) t - in (HOLogic.mk_binrel p (d, linear_add vars ct z)) - end - else fm - |(Const ("Not", _) $ p) => HOLogic.Not $ (linrep vars x t p) - |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (linrep vars x t p) $ (linrep vars x t q) - |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (linrep vars x t p) $ (linrep vars x t q) - |_ => fm; - -(* ------------------------------------------------------------------------- *) -(* Evaluation of constant expressions. *) -(* ------------------------------------------------------------------------- *) - -(* An other implementation of divides, that covers more cases*) - -exception DVD_UNKNOWN - -fun dvd_op (d, t) = - if not(is_number d) then raise DVD_UNKNOWN - else let - val dn = dest_number d - fun coeffs_of x = case x of - Const(p,_) $ tl $ tr => - if p = @{const_name HOL.plus} then (coeffs_of tl) union (coeffs_of tr) - else if p = @{const_name HOL.times} - then if (is_number tr) - then [(dest_number tr) * (dest_number tl)] - else [dest_number tl] - else [] - |_ => if (is_number t) then [dest_number t] else [] - val ts = coeffs_of t - in case ts of - [] => raise DVD_UNKNOWN - |_ => fold_rev (fn k => fn r => r andalso (k mod dn = 0)) ts true - end; - - -val operations = - [("op =",op=), (@{const_name Orderings.less},IntInf.<), ("op >",IntInf.>), (@{const_name Orderings.less_eq},IntInf.<=) , - ("op >=",IntInf.>=), - ("Divides.dvd",fn (x,y) =>((IntInf.mod(y, x)) = 0))]; - -fun applyoperation (SOME f) (a,b) = f (a, b) - |applyoperation _ (_, _) = false; - -(*Evaluation of constant atomic formulas*) - (*FIXME : This is an optimation but still incorrect !! *) -(* -fun evalc_atom at = case at of - (Const (p,_) $ s $ t) => - (if p="Divides.dvd" then - ((if dvd_op(s,t) then HOLogic.true_const - else HOLogic.false_const) - handle _ => at) - else - case AList.lookup (op =) operations p of - SOME f => ((if (f ((dest_number s),(dest_number t))) then HOLogic.true_const else HOLogic.false_const) - handle _ => at) - | _ => at) - |Const("Not",_)$(Const (p,_) $ s $ t) =>( - case AList.lookup (op =) operations p of - SOME f => ((if (f ((dest_number s),(dest_number t))) then - HOLogic.false_const else HOLogic.true_const) - handle _ => at) - | _ => at) - | _ => at; - -*) - -fun evalc_atom at = case at of - (Const (p,_) $ s $ t) => - ( case AList.lookup (op =) operations p of - SOME f => ((if (f ((dest_number s),(dest_number t))) then HOLogic.true_const - else HOLogic.false_const) - handle _ => at) - | _ => at) - |Const("Not",_)$(Const (p,_) $ s $ t) =>( - case AList.lookup (op =) operations p of - SOME f => ((if (f ((dest_number s),(dest_number t))) - then HOLogic.false_const else HOLogic.true_const) - handle _ => at) - | _ => at) - | _ => at; - - (*Function onatoms apllys function f on the atomic formulas involved in a.*) - -fun onatoms f a = if (is_arith_rel a) then f a else case a of - - (Const ("Not",_) $ p) => if is_arith_rel p then HOLogic.Not $ (f p) - - else HOLogic.Not $ (onatoms f p) - |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (onatoms f p) $ (onatoms f q) - |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (onatoms f p) $ (onatoms f q) - |(Const ("op -->",_) $ p $ q) => HOLogic.imp $ (onatoms f p) $ (onatoms f q) - |((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) => (Const ("op =", [HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)) $ (onatoms f p) $ (onatoms f q) - |(Const("All",_) $ Abs(x,T,p)) => Const("All", [HOLogic.intT --> - HOLogic.boolT] ---> HOLogic.boolT)$ Abs (x ,T, (onatoms f p)) - |(Const("Ex",_) $ Abs(x,T,p)) => Const("Ex", [HOLogic.intT --> HOLogic.boolT]---> HOLogic.boolT) $ Abs( x ,T, (onatoms f p)) - |_ => a; - -val evalc = onatoms evalc_atom; - -(* ------------------------------------------------------------------------- *) -(* Hence overall quantifier elimination. *) -(* ------------------------------------------------------------------------- *) - - -(*list_disj[conj] makes a disj[conj] of a given list. used with conjucts or disjuncts -it liearises iterated conj[disj]unctions. *) - -fun list_disj [] = HOLogic.false_const - | list_disj ps = foldr1 (fn (p, q) => HOLogic.disj $ p $ q) ps; - -fun list_conj [] = HOLogic.true_const - | list_conj ps = foldr1 (fn (p, q) => HOLogic.conj $ p $ q) ps; - - -(*Simplification of Formulas *) - -(*Function q_bnd_chk checks if a quantified Formula makes sens : Means if in -the body of the existential quantifier there are bound variables to the -existential quantifier.*) - -fun has_bound fm =let fun has_boundh fm i = case fm of - Bound n => (i = n) - |Abs (_,_,p) => has_boundh p (i+1) - |t1 $ t2 => (has_boundh t1 i) orelse (has_boundh t2 i) - |_ =>false - -in case fm of - Bound _ => true - |Abs (_,_,p) => has_boundh p 0 - |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) - |_ =>false -end; - -(*has_sub_abs checks if in a given Formula there are subformulas which are quantifed -too. Is no used no more.*) - -fun has_sub_abs fm = case fm of - Abs (_,_,_) => true - |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) - |_ =>false ; - -(*update_bounds called with i=0 udates the numeration of bounded variables because the -formula will not be quantified any more.*) - -fun update_bounds fm i = case fm of - Bound n => if n >= i then Bound (n-1) else fm - |Abs (x,T,p) => Abs(x,T,(update_bounds p (i+1))) - |t1 $ t2 => (update_bounds t1 i) $ (update_bounds t2 i) - |_ => fm ; - -(*psimpl : Simplification of propositions (general purpose)*) -fun psimpl1 fm = case fm of - Const("Not",_) $ Const ("False",_) => HOLogic.true_const - | Const("Not",_) $ Const ("True",_) => HOLogic.false_const - | Const("op &",_) $ Const ("False",_) $ q => HOLogic.false_const - | Const("op &",_) $ p $ Const ("False",_) => HOLogic.false_const - | Const("op &",_) $ Const ("True",_) $ q => q - | Const("op &",_) $ p $ Const ("True",_) => p - | Const("op |",_) $ Const ("False",_) $ q => q - | Const("op |",_) $ p $ Const ("False",_) => p - | Const("op |",_) $ Const ("True",_) $ q => HOLogic.true_const - | Const("op |",_) $ p $ Const ("True",_) => HOLogic.true_const - | Const("op -->",_) $ Const ("False",_) $ q => HOLogic.true_const - | Const("op -->",_) $ Const ("True",_) $ q => q - | Const("op -->",_) $ p $ Const ("True",_) => HOLogic.true_const - | Const("op -->",_) $ p $ Const ("False",_) => HOLogic.Not $ p - | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("True",_) $ q => q - | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("True",_) => p - | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("False",_) $ q => HOLogic.Not $ q - | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("False",_) => HOLogic.Not $ p - | _ => fm; - -fun psimpl fm = case fm of - Const ("Not",_) $ p => psimpl1 (HOLogic.Not $ (psimpl p)) - | Const("op &",_) $ p $ q => psimpl1 (HOLogic.mk_conj (psimpl p,psimpl q)) - | Const("op |",_) $ p $ q => psimpl1 (HOLogic.mk_disj (psimpl p,psimpl q)) - | Const("op -->",_) $ p $ q => psimpl1 (HOLogic.mk_imp(psimpl p,psimpl q)) - | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q => psimpl1 (HOLogic.mk_eq(psimpl p,psimpl q)) - | _ => fm; - - -(*simpl : Simplification of Terms involving quantifiers too. - This function is able to drop out some quantified expressions where there are no - bound varaibles.*) - -fun simpl1 fm = - case fm of - Const("All",_) $Abs(x,_,p) => if (has_bound fm ) then fm - else (update_bounds p 0) - | Const("Ex",_) $ Abs (x,_,p) => if has_bound fm then fm - else (update_bounds p 0) - | _ => psimpl fm; - -fun simpl fm = case fm of - Const ("Not",_) $ p => simpl1 (HOLogic.Not $(simpl p)) - | Const ("op &",_) $ p $ q => simpl1 (HOLogic.mk_conj (simpl p ,simpl q)) - | Const ("op |",_) $ p $ q => simpl1 (HOLogic.mk_disj (simpl p ,simpl q )) - | Const ("op -->",_) $ p $ q => simpl1 (HOLogic.mk_imp(simpl p ,simpl q )) - | Const("op =", Type ("fun",[Type ("bool", []),_]))$ p $ q => simpl1 - (HOLogic.mk_eq(simpl p ,simpl q )) -(* | Const ("All",Ta) $ Abs(Vn,VT,p) => simpl1(Const("All",Ta) $ - Abs(Vn,VT,simpl p )) - | Const ("Ex",Ta) $ Abs(Vn,VT,p) => simpl1(Const("Ex",Ta) $ - Abs(Vn,VT,simpl p )) -*) - | _ => fm; - -(* ------------------------------------------------------------------------- *) - -(* Puts fm into NNF*) - -fun nnf fm = if (is_arith_rel fm) then fm -else (case fm of - ( Const ("op &",_) $ p $ q) => HOLogic.conj $ (nnf p) $(nnf q) - | (Const("op |",_) $ p $q) => HOLogic.disj $ (nnf p)$(nnf q) - | (Const ("op -->",_) $ p $ q) => HOLogic.disj $ (nnf (HOLogic.Not $ p)) $ (nnf q) - | ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) =>(HOLogic.disj $ (HOLogic.conj $ (nnf p) $ (nnf q)) $ (HOLogic.conj $ (nnf (HOLogic.Not $ p) ) $ (nnf(HOLogic.Not $ q)))) - | (Const ("Not",_)) $ ((Const ("Not",_)) $ p) => (nnf p) - | (Const ("Not",_)) $ (( Const ("op &",_)) $ p $ q) =>HOLogic.disj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $q)) - | (Const ("Not",_)) $ (( Const ("op |",_)) $ p $ q) =>HOLogic.conj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $ q)) - | (Const ("Not",_)) $ (( Const ("op -->",_)) $ p $ q ) =>HOLogic.conj $ (nnf p) $(nnf(HOLogic.Not $ q)) - | (Const ("Not",_)) $ ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q ) =>(HOLogic.disj $ (HOLogic.conj $(nnf p) $ (nnf(HOLogic.Not $ q))) $ (HOLogic.conj $(nnf(HOLogic.Not $ p)) $ (nnf q))) - | _ => fm); - - -(* Function remred to remove redundancy in a list while keeping the order of appearance of the -elements. but VERY INEFFICIENT!! *) - -fun remred1 el [] = [] - |remred1 el (h::t) = if el=h then (remred1 el t) else h::(remred1 el t); - -fun remred [] = [] - |remred (x::l) = x::(remred1 x (remred l)); - -(*Makes sure that all free Variables are of the type integer but this function is only -used temporarily, this job must be done by the parser later on.*) - -fun mk_uni_vars T (node $ rest) = (case node of - Free (name,_) => Free (name,T) $ (mk_uni_vars T rest) - |_=> (mk_uni_vars T node) $ (mk_uni_vars T rest ) ) - |mk_uni_vars T (Free (v,_)) = Free (v,T) - |mk_uni_vars T tm = tm; - -fun mk_uni_int T (Const (@{const_name HOL.zero},T2)) = if T = T2 then (mk_number 0) else (Const (@{const_name HOL.zero},T2)) - |mk_uni_int T (Const (@{const_name HOL.one},T2)) = if T = T2 then (mk_number 1) else (Const (@{const_name HOL.one},T2)) - |mk_uni_int T (node $ rest) = (mk_uni_int T node) $ (mk_uni_int T rest ) - |mk_uni_int T (Abs(AV,AT,p)) = Abs(AV,AT,mk_uni_int T p) - |mk_uni_int T tm = tm; - - -(* Minusinfinity Version*) -fun myupto (m:IntInf.int) n = if m > n then [] else m::(myupto (m+1) n) - -fun coopermi vars1 fm = - case fm of - Const ("Ex",_) $ Abs(x0,T,p0) => - let - val (xn,p1) = Syntax.variant_abs (x0,T,p0) - val x = Free (xn,T) - val vars = (xn::vars1) - val p = unitycoeff x (posineq (simpl p1)) - val p_inf = simpl (minusinf x p) - val bset = bset x p - val js = myupto 1 (divlcm x p) - fun p_element j b = linrep vars x (linear_add vars b (mk_number j)) p - fun stage j = list_disj (linrep vars x (mk_number j) p_inf :: map (p_element j) bset) - in (list_disj (map stage js)) - end - | _ => error "cooper: not an existential formula"; - - - -(* The plusinfinity version of cooper*) -fun cooperpi vars1 fm = - case fm of - Const ("Ex",_) $ Abs(x0,T,p0) => let - val (xn,p1) = Syntax.variant_abs (x0,T,p0) - val x = Free (xn,T) - val vars = (xn::vars1) - val p = unitycoeff x (posineq (simpl p1)) - val p_inf = simpl (plusinf x p) - val aset = aset x p - val js = myupto 1 (divlcm x p) - fun p_element j a = linrep vars x (linear_sub vars a (mk_number j)) p - fun stage j = list_disj (linrep vars x (mk_number j) p_inf :: map (p_element j) aset) - in (list_disj (map stage js)) - end - | _ => error "cooper: not an existential formula"; - - -(* Try to find a withness for the formula *) - -fun inf_w mi d vars x p = - let val f = if mi then minusinf else plusinf in - case (simpl (minusinf x p)) of - Const("True",_) => (SOME (mk_number 1), HOLogic.true_const) - |Const("False",_) => (NONE,HOLogic.false_const) - |F => - let - fun h n = - case ((simpl o evalc) (linrep vars x (mk_number n) F)) of - Const("True",_) => (SOME (mk_number n),HOLogic.true_const) - |F' => if n=1 then (NONE,F') - else let val (rw,rf) = h (n-1) in - (rw,HOLogic.mk_disj(F',rf)) - end - - in (h d) - end - end; - -fun set_w d b st vars x p = let - fun h ns = case ns of - [] => (NONE,HOLogic.false_const) - |n::nl => ( case ((simpl o evalc) (linrep vars x n p)) of - Const("True",_) => (SOME n,HOLogic.true_const) - |F' => let val (rw,rf) = h nl - in (rw,HOLogic.mk_disj(F',rf)) - end) - val f = if b then linear_add else linear_sub - val p_elements = fold_rev (fn i => fn l => l union (map (fn e => f [] e (mk_number i)) st)) (myupto 1 d) [] - in h p_elements - end; - -fun withness d b st vars x p = case (inf_w b d vars x p) of - (SOME n,_) => (SOME n,HOLogic.true_const) - |(NONE,Pinf) => (case (set_w d b st vars x p) of - (SOME n,_) => (SOME n,HOLogic.true_const) - |(_,Pst) => (NONE,HOLogic.mk_disj(Pinf,Pst))); - - - - -(*Cooper main procedure*) - -exception STAGE_TRUE; - - -fun cooper vars1 fm = - case fm of - Const ("Ex",_) $ Abs(x0,T,p0) => let - val (xn,p1) = Syntax.variant_abs (x0,T,p0) - val x = Free (xn,T) - val vars = (xn::vars1) -(* val p = unitycoeff x (posineq (simpl p1)) *) - val p = unitycoeff x p1 - val ast = aset x p - val bst = bset x p - val js = myupto 1 (divlcm x p) - val (p_inf,f,S ) = - if (length bst) <= (length ast) - then (simpl (minusinf x p),linear_add,bst) - else (simpl (plusinf x p), linear_sub,ast) - fun p_element j a = linrep vars x (f vars a (mk_number j)) p - fun stage j = list_disj (linrep vars x (mk_number j) p_inf :: map (p_element j) S) - fun stageh n = ((if n = 0 then [] - else - let - val nth_stage = simpl (evalc (stage n)) - in - if (nth_stage = HOLogic.true_const) - then raise STAGE_TRUE - else if (nth_stage = HOLogic.false_const) then stageh (n-1) - else nth_stage::(stageh (n-1)) - end ) - handle STAGE_TRUE => [HOLogic.true_const]) - val slist = stageh (divlcm x p) - in (list_disj slist) - end - | _ => error "cooper: not an existential formula"; - - -(* A Version of cooper that returns a withness *) -fun cooper_w vars1 fm = - case fm of - Const ("Ex",_) $ Abs(x0,T,p0) => let - val (xn,p1) = Syntax.variant_abs (x0,T,p0) - val x = Free (xn,T) - val vars = (xn::vars1) -(* val p = unitycoeff x (posineq (simpl p1)) *) - val p = unitycoeff x p1 - val ast = aset x p - val bst = bset x p - val d = divlcm x p - val (p_inf,S ) = - if (length bst) <= (length ast) - then (true,bst) - else (false,ast) - in withness d p_inf S vars x p -(* fun p_element j a = linrep vars x (f vars a (mk_number j)) p - fun stage j = list_disj (linrep vars x (mk_number j) p_inf :: map (p_element j) S) - in (list_disj (map stage js)) -*) - end - | _ => error "cooper: not an existential formula"; - - -(* ------------------------------------------------------------------------- *) -(* Free variables in terms and formulas. *) -(* ------------------------------------------------------------------------- *) - -fun fvt tml = case tml of - [] => [] - | Free(x,_)::r => x::(fvt r) - -fun fv fm = fvt (term_frees fm); - - -(* ========================================================================= *) -(* Quantifier elimination. *) -(* ========================================================================= *) -(*conj[/disj]uncts lists iterated conj[disj]unctions*) - -fun disjuncts fm = case fm of - Const ("op |",_) $ p $ q => (disjuncts p) @ (disjuncts q) - | _ => [fm]; - -fun conjuncts fm = case fm of - Const ("op &",_) $p $ q => (conjuncts p) @ (conjuncts q) - | _ => [fm]; - - - -(* ------------------------------------------------------------------------- *) -(* Lift procedure given literal modifier, formula normalizer & basic quelim. *) -(* ------------------------------------------------------------------------- *) - -fun lift_qelim afn nfn qfn isat = -let -fun qelift vars fm = if (isat fm) then afn vars fm -else -case fm of - Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) - | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) - | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) - | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) - | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) - | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) - | (e as Const ("Ex",_)) $ Abs (x,T,p) => qfn vars (e$Abs (x,T,(nfn(qelift (x::vars) p)))) - | _ => fm - -in (fn fm => qelift (fv fm) fm) -end; - - -(* -fun lift_qelim afn nfn qfn isat = - let fun qelim x vars p = - let val cjs = conjuncts p - val (ycjs,ncjs) = List.partition (has_bound) cjs in - (if ycjs = [] then p else - let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT - ) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in - (fold_rev conj_help ncjs q) - end) - end - - fun qelift vars fm = if (isat fm) then afn vars fm - else - case fm of - Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) - | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) - | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) - | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) - | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) - | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) - | Const ("Ex",_) $ Abs (x,T,p) => let val djs = disjuncts(nfn(qelift (x::vars) p)) in - list_disj(map (qelim x vars) djs) end - | _ => fm - - in (fn fm => simpl(qelift (fv fm) fm)) - end; -*) - -(* ------------------------------------------------------------------------- *) -(* Cleverer (proposisional) NNF with conditional and literal modification. *) -(* ------------------------------------------------------------------------- *) - -(*Function Negate used by cnnf, negates a formula p*) - -fun negate (Const ("Not",_) $ p) = p - |negate p = (HOLogic.Not $ p); - -fun cnnf lfn = - let fun cnnfh fm = case fm of - (Const ("op &",_) $ p $ q) => HOLogic.mk_conj(cnnfh p,cnnfh q) - | (Const ("op |",_) $ p $ q) => HOLogic.mk_disj(cnnfh p,cnnfh q) - | (Const ("op -->",_) $ p $q) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh q) - | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => HOLogic.mk_disj( - HOLogic.mk_conj(cnnfh p,cnnfh q), - HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $q))) - - | (Const ("Not",_) $ (Const("Not",_) $ p)) => cnnfh p - | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) - | (Const ("Not",_) $(Const ("op |",_) $ (Const ("op &",_) $ p $ q) $ - (Const ("op &",_) $ p1 $ r))) => if p1 = negate p then - HOLogic.mk_disj( - cnnfh (HOLogic.mk_conj(p,cnnfh(HOLogic.Not $ q))), - cnnfh (HOLogic.mk_conj(p1,cnnfh(HOLogic.Not $ r)))) - else HOLogic.mk_conj( - cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p),cnnfh(HOLogic.Not $ q))), - cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p1),cnnfh(HOLogic.Not $ r))) - ) - | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) - | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) => HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)) - | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q)) => HOLogic.mk_disj(HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)),HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh q)) - | _ => lfn fm -in cnnfh - end; - -(*End- function the quantifierelimination an decion procedure of presburger formulas.*) - -(* -val integer_qelim = simpl o evalc o (lift_qelim linform (simpl o (cnnf posineq o evalc)) cooper is_arith_rel) ; -*) - - -val integer_qelim = simpl o evalc o (lift_qelim linform (cnnf posineq o evalc) cooper is_arith_rel) ; - -end;
--- a/src/HOL/Integ/cooper_proof.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,988 +0,0 @@ -(* Title: HOL/Integ/cooper_proof.ML - ID: $Id$ - Author: Amine Chaieb and Tobias Nipkow, TU Muenchen - -File containing the implementation of the proof -generation for Cooper Algorithm -*) - - -signature COOPER_PROOF = -sig - val qe_Not : thm - val qe_conjI : thm - val qe_disjI : thm - val qe_impI : thm - val qe_eqI : thm - val qe_exI : thm - val list_to_set : typ -> term list -> term - val qe_get_terms : thm -> term * term - val cooper_prv : theory -> term -> term -> thm - val proof_of_evalc : theory -> term -> thm - val proof_of_cnnf : theory -> term -> (term -> thm) -> thm - val proof_of_linform : theory -> string list -> term -> thm - val proof_of_adjustcoeffeq : theory -> term -> IntInf.int -> term -> thm - val prove_elementar : theory -> string -> term -> thm - val thm_of : theory -> (term -> (term list * (thm list -> thm))) -> term -> thm -end; - -structure CooperProof : COOPER_PROOF = -struct -open CooperDec; - -val presburger_ss = simpset () - addsimps [diff_int_def] delsimps [thm "diff_int_def_symmetric"]; - -val cboolT = ctyp_of HOL.thy HOLogic.boolT; - -(*Theorems that will be used later for the proofgeneration*) - -val zdvd_iff_zmod_eq_0 = thm "zdvd_iff_zmod_eq_0"; -val unity_coeff_ex = thm "unity_coeff_ex"; - -(* Theorems for proving the adjustment of the coefficients*) - -val ac_lt_eq = thm "ac_lt_eq"; -val ac_eq_eq = thm "ac_eq_eq"; -val ac_dvd_eq = thm "ac_dvd_eq"; -val ac_pi_eq = thm "ac_pi_eq"; - -(* The logical compination of the sythetised properties*) -val qe_Not = thm "qe_Not"; -val qe_conjI = thm "qe_conjI"; -val qe_disjI = thm "qe_disjI"; -val qe_impI = thm "qe_impI"; -val qe_eqI = thm "qe_eqI"; -val qe_exI = thm "qe_exI"; -val qe_ALLI = thm "qe_ALLI"; - -(*Modulo D property for Pminusinf an Plusinf *) -val fm_modd_minf = thm "fm_modd_minf"; -val not_dvd_modd_minf = thm "not_dvd_modd_minf"; -val dvd_modd_minf = thm "dvd_modd_minf"; - -val fm_modd_pinf = thm "fm_modd_pinf"; -val not_dvd_modd_pinf = thm "not_dvd_modd_pinf"; -val dvd_modd_pinf = thm "dvd_modd_pinf"; - -(* the minusinfinity proprty*) - -val fm_eq_minf = thm "fm_eq_minf"; -val neq_eq_minf = thm "neq_eq_minf"; -val eq_eq_minf = thm "eq_eq_minf"; -val le_eq_minf = thm "le_eq_minf"; -val len_eq_minf = thm "len_eq_minf"; -val not_dvd_eq_minf = thm "not_dvd_eq_minf"; -val dvd_eq_minf = thm "dvd_eq_minf"; - -(* the Plusinfinity proprty*) - -val fm_eq_pinf = thm "fm_eq_pinf"; -val neq_eq_pinf = thm "neq_eq_pinf"; -val eq_eq_pinf = thm "eq_eq_pinf"; -val le_eq_pinf = thm "le_eq_pinf"; -val len_eq_pinf = thm "len_eq_pinf"; -val not_dvd_eq_pinf = thm "not_dvd_eq_pinf"; -val dvd_eq_pinf = thm "dvd_eq_pinf"; - -(*Logical construction of the Property*) -val eq_minf_conjI = thm "eq_minf_conjI"; -val eq_minf_disjI = thm "eq_minf_disjI"; -val modd_minf_disjI = thm "modd_minf_disjI"; -val modd_minf_conjI = thm "modd_minf_conjI"; - -val eq_pinf_conjI = thm "eq_pinf_conjI"; -val eq_pinf_disjI = thm "eq_pinf_disjI"; -val modd_pinf_disjI = thm "modd_pinf_disjI"; -val modd_pinf_conjI = thm "modd_pinf_conjI"; - -(*Cooper Backwards...*) -(*Bset*) -val not_bst_p_fm = thm "not_bst_p_fm"; -val not_bst_p_ne = thm "not_bst_p_ne"; -val not_bst_p_eq = thm "not_bst_p_eq"; -val not_bst_p_gt = thm "not_bst_p_gt"; -val not_bst_p_lt = thm "not_bst_p_lt"; -val not_bst_p_ndvd = thm "not_bst_p_ndvd"; -val not_bst_p_dvd = thm "not_bst_p_dvd"; - -(*Aset*) -val not_ast_p_fm = thm "not_ast_p_fm"; -val not_ast_p_ne = thm "not_ast_p_ne"; -val not_ast_p_eq = thm "not_ast_p_eq"; -val not_ast_p_gt = thm "not_ast_p_gt"; -val not_ast_p_lt = thm "not_ast_p_lt"; -val not_ast_p_ndvd = thm "not_ast_p_ndvd"; -val not_ast_p_dvd = thm "not_ast_p_dvd"; - -(*Logical construction of the prop*) -(*Bset*) -val not_bst_p_conjI = thm "not_bst_p_conjI"; -val not_bst_p_disjI = thm "not_bst_p_disjI"; -val not_bst_p_Q_elim = thm "not_bst_p_Q_elim"; - -(*Aset*) -val not_ast_p_conjI = thm "not_ast_p_conjI"; -val not_ast_p_disjI = thm "not_ast_p_disjI"; -val not_ast_p_Q_elim = thm "not_ast_p_Q_elim"; - -(*Cooper*) -val cppi_eq = thm "cppi_eq"; -val cpmi_eq = thm "cpmi_eq"; - -(*Others*) -val simp_from_to = thm "simp_from_to"; -val P_eqtrue = thm "P_eqtrue"; -val P_eqfalse = thm "P_eqfalse"; - -(*For Proving NNF*) - -val nnf_nn = thm "nnf_nn"; -val nnf_im = thm "nnf_im"; -val nnf_eq = thm "nnf_eq"; -val nnf_sdj = thm "nnf_sdj"; -val nnf_ncj = thm "nnf_ncj"; -val nnf_nim = thm "nnf_nim"; -val nnf_neq = thm "nnf_neq"; -val nnf_ndj = thm "nnf_ndj"; - -(*For Proving term linearizition*) -val linearize_dvd = thm "linearize_dvd"; -val lf_lt = thm "lf_lt"; -val lf_eq = thm "lf_eq"; -val lf_dvd = thm "lf_dvd"; - - -(* ------------------------------------------------------------------------- *) -(*This function norm_zero_one replaces the occurences of Numeral1 and Numeral0*) -(*Respectively by their abstract representation Const(@{const_name HOL.one},..) and Const(@{const_name HOL.zero},..)*) -(*this is necessary because the theorems use this representation.*) -(* This function should be elminated in next versions...*) -(* ------------------------------------------------------------------------- *) - -fun norm_zero_one fm = case fm of - (Const (@{const_name HOL.times},_) $ c $ t) => - if c = one then (norm_zero_one t) - else if (dest_number c = ~1) - then (Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t)) - else (HOLogic.mk_binop @{const_name HOL.times} (norm_zero_one c,norm_zero_one t)) - |(node $ rest) => ((norm_zero_one node)$(norm_zero_one rest)) - |(Abs(x,T,p)) => (Abs(x,T,(norm_zero_one p))) - |_ => fm; - -(* ------------------------------------------------------------------------- *) -(*function list to Set, constructs a set containing all elements of a given list.*) -(* ------------------------------------------------------------------------- *) -fun list_to_set T1 l = let val T = (HOLogic.mk_setT T1) in - case l of - [] => Const ("{}",T) - |(h::t) => Const("insert", T1 --> (T --> T)) $ h $(list_to_set T1 t) - end; - -(* ------------------------------------------------------------------------- *) -(* Returns both sides of an equvalence in the theorem*) -(* ------------------------------------------------------------------------- *) -fun qe_get_terms th = let val (_$(Const("op =",Type ("fun",[Type ("bool", []),_])) $ A $ B )) = prop_of th in (A,B) end; - -(* ------------------------------------------------------------------------- *) -(*This function proove elementar will be used to generate proofs at - runtime*) (*It is thought to prove properties such as a dvd b - (essentially) that are only to make at runtime.*) -(* ------------------------------------------------------------------------- *) -fun prove_elementar thy s fm2 = - Goal.prove (ProofContext.init thy) [] [] (HOLogic.mk_Trueprop fm2) (fn _ => EVERY - (case s of - (*"ss" like simplification with simpset*) - "ss" => - let val ss = presburger_ss addsimps [zdvd_iff_zmod_eq_0,unity_coeff_ex] - in [simp_tac ss 1, TRY (simple_arith_tac 1)] end - - (*"bl" like blast tactic*) - (* Is only used in the harrisons like proof procedure *) - | "bl" => [blast_tac HOL_cs 1] - - (*"ed" like Existence disjunctions ...*) - (* Is only used in the harrisons like proof procedure *) - | "ed" => - let - val ex_disj_tacs = - let - val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1] - val tac2 = EVERY[etac exE 1, rtac exI 1, - REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1] - in [rtac iffI 1, - etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1, - REPEAT(EVERY[etac disjE 1, tac2]), tac2] - end - in ex_disj_tacs end - - | "fa" => [simple_arith_tac 1] - - | "sa" => - let val ss = presburger_ss addsimps zadd_ac - in [simp_tac ss 1, TRY (simple_arith_tac 1)] end - - (* like Existance Conjunction *) - | "ec" => - let val ss = presburger_ss addsimps zadd_ac - in [simp_tac ss 1, TRY (blast_tac HOL_cs 1)] end - - | "ac" => - let val ss = HOL_basic_ss addsimps zadd_ac - in [simp_tac ss 1] end - - | "lf" => - let val ss = presburger_ss addsimps zadd_ac - in [simp_tac ss 1, TRY (simple_arith_tac 1)] end)); - -(*=============================================================*) -(*-------------------------------------------------------------*) -(* The new compact model *) -(*-------------------------------------------------------------*) -(*=============================================================*) - -fun thm_of sg decomp t = - let val (ts,recomb) = decomp t - in recomb (map (thm_of sg decomp) ts) - end; - -(*==================================================*) -(* Compact Version for adjustcoeffeq *) -(*==================================================*) - -fun decomp_adjustcoeffeq sg x l fm = case fm of - (Const("Not",_)$(Const(@{const_name Orderings.less},_) $ zero $(rt as (Const (@{const_name HOL.plus}, _)$(Const (@{const_name HOL.times},_) $ c $ y ) $z )))) => - let - val m = l div (dest_number c) - val n = if (x = y) then abs (m) else 1 - val xtm = (HOLogic.mk_binop @{const_name HOL.times} ((mk_number ((m div n)*l) ), x)) - val rs = if (x = y) - then (HOLogic.mk_binrel @{const_name Orderings.less} (zero,linear_sub [] (mk_number n) (HOLogic.mk_binop @{const_name HOL.plus} ( xtm ,( linear_cmul n z) )))) - else HOLogic.mk_binrel @{const_name Orderings.less} (zero,linear_sub [] one rt ) - val ck = cterm_of sg (mk_number n) - val cc = cterm_of sg c - val ct = cterm_of sg z - val cx = cterm_of sg y - val pre = prove_elementar sg "lf" - (HOLogic.mk_binrel @{const_name Orderings.less} (zero, mk_number n)) - val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq))) - in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) - end - - |(Const(p,_) $a $( Const (@{const_name HOL.plus}, _)$(Const (@{const_name HOL.times},_) $ - c $ y ) $t )) => - if (is_arith_rel fm) andalso (x = y) - then - let val m = l div (dest_number c) - val k = (if p = @{const_name Orderings.less} then abs(m) else m) - val xtm = (HOLogic.mk_binop @{const_name HOL.times} ((mk_number ((m div k)*l) ), x)) - val rs = (HOLogic.mk_binrel p ((linear_cmul k a),(HOLogic.mk_binop @{const_name HOL.plus} ( xtm ,( linear_cmul k t) )))) - - val ck = cterm_of sg (mk_number k) - val cc = cterm_of sg c - val ct = cterm_of sg t - val cx = cterm_of sg x - val ca = cterm_of sg a - - in - case p of - @{const_name Orderings.less} => - let val pre = prove_elementar sg "lf" - (HOLogic.mk_binrel @{const_name Orderings.less} (zero, mk_number k)) - val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq))) - in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) - end - - |"op =" => - let val pre = prove_elementar sg "lf" - (HOLogic.Not $ (HOLogic.mk_binrel "op =" (zero, mk_number k))) - val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq))) - in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) - end - - |"Divides.dvd" => - let val pre = prove_elementar sg "lf" - (HOLogic.Not $ (HOLogic.mk_binrel "op =" (zero, mk_number k))) - val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq)) - in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) - - end - end - else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl) - - |( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not) - |( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI) - |( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI) - - |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl); - -fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l); - - - -(*==================================================*) -(* Finding rho for modd_minusinfinity *) -(*==================================================*) -fun rho_for_modd_minf x dlcm sg fm1 = -let - (*Some certified Terms*) - - val ctrue = cterm_of sg HOLogic.true_const - val cfalse = cterm_of sg HOLogic.false_const - val fm = norm_zero_one fm1 - in case fm1 of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) - then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if (y=x) andalso (c1 = zero) then - if (pm1 = one) then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) else - (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) - - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cz = cterm_of sg (norm_zero_one z) - val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero) - in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf))) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) - |(Const("Divides.dvd",_)$ d $ (db as (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ - c $ y ) $ z))) => - if y=x then let val cz = cterm_of sg (norm_zero_one z) - val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero) - in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf))) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) - - - |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf) - end; -(*=========================================================================*) -(*=========================================================================*) -fun rho_for_eq_minf x dlcm sg fm1 = - let - val fm = norm_zero_one fm1 - in case fm1 of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if (x=y) andalso (c1=zero) andalso (c2=one) - then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one)) - then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if (y=x) andalso (c1 =zero) then - if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else - (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_minf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cd = cterm_of sg (norm_zero_one d) - val cz = cterm_of sg (norm_zero_one z) - in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_minf)) - end - - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - - |(Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cd = cterm_of sg (norm_zero_one d) - val cz = cterm_of sg (norm_zero_one z) - in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_minf)) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - - - |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) - end; - -(*=====================================================*) -(*=====================================================*) -(*=========== minf proofs with the compact version==========*) -fun decomp_minf_eq x dlcm sg t = case t of - Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_conjI) - |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_disjI) - |_ => ([],fn [] => rho_for_eq_minf x dlcm sg t); - -fun decomp_minf_modd x dlcm sg t = case t of - Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_conjI) - |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_disjI) - |_ => ([],fn [] => rho_for_modd_minf x dlcm sg t); - -(* -------------------------------------------------------------*) -(* Finding rho for pinf_modd *) -(* -------------------------------------------------------------*) -fun rho_for_modd_pinf x dlcm sg fm1 = -let - (*Some certified Terms*) - - val ctrue = cterm_of sg HOLogic.true_const - val cfalse = cterm_of sg HOLogic.false_const - val fm = norm_zero_one fm1 - in case fm1 of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if ((x=y) andalso (c1= zero) andalso (c2= one)) - then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero) andalso (c2 = one)) - then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if ((y=x) andalso (c1 = zero)) then - if (pm1 = one) - then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf)) - else (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) - - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cz = cterm_of sg (norm_zero_one z) - val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero) - in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf))) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) - |(Const("Divides.dvd",_)$ d $ (db as (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ - c $ y ) $ z))) => - if y=x then let val cz = cterm_of sg (norm_zero_one z) - val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero) - in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf))) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) - - - |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf) - end; -(* -------------------------------------------------------------*) -(* Finding rho for pinf_eq *) -(* -------------------------------------------------------------*) -fun rho_for_eq_pinf x dlcm sg fm1 = - let - val fm = norm_zero_one fm1 - in case fm1 of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if (x=y) andalso (c1=zero) andalso (c2=one) - then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one)) - then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if (y=x) andalso (c1 =zero) then - if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else - (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_pinf)) - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cd = cterm_of sg (norm_zero_one d) - val cz = cterm_of sg (norm_zero_one z) - in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_pinf)) - end - - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - - |(Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then let val cd = cterm_of sg (norm_zero_one d) - val cz = cterm_of sg (norm_zero_one z) - in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_pinf)) - end - else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - - - |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) - end; - - - -fun minf_proof_of_c sg x dlcm t = - let val minf_eqth = thm_of sg (decomp_minf_eq x dlcm sg) t - val minf_moddth = thm_of sg (decomp_minf_modd x dlcm sg) t - in (minf_eqth, minf_moddth) -end; - -(*=========== pinf proofs with the compact version==========*) -fun decomp_pinf_eq x dlcm sg t = case t of - Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_conjI) - |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_disjI) - |_ =>([],fn [] => rho_for_eq_pinf x dlcm sg t) ; - -fun decomp_pinf_modd x dlcm sg t = case t of - Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_conjI) - |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_disjI) - |_ => ([],fn [] => rho_for_modd_pinf x dlcm sg t); - -fun pinf_proof_of_c sg x dlcm t = - let val pinf_eqth = thm_of sg (decomp_pinf_eq x dlcm sg) t - val pinf_moddth = thm_of sg (decomp_pinf_modd x dlcm sg) t - in (pinf_eqth,pinf_moddth) -end; - - -(* ------------------------------------------------------------------------- *) -(* Here we generate the theorem for the Bset Property in the simple direction*) -(* It is just an instantiation*) -(* ------------------------------------------------------------------------- *) -(* -fun bsetproof_of sg (x as Free(xn,xT)) fm bs dlcm = - let - val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) - val cdlcm = cterm_of sg dlcm - val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs)) - in instantiate' [] [SOME cdlcm,SOME cB, SOME cp] (bst_thm) -end; - -fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm = - let - val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) - val cdlcm = cterm_of sg dlcm - val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast)) - in instantiate' [] [SOME cdlcm,SOME cA, SOME cp] (ast_thm) -end; -*) - -(* For the generation of atomic Theorems*) -(* Prove the premisses on runtime and then make RS*) -(* ------------------------------------------------------------------------- *) - -(*========= this is rho ============*) -fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = - let - val cdlcm = cterm_of sg dlcm - val cB = cterm_of sg B - val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) - val cat = cterm_of sg (norm_zero_one at) - in - case at of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if (x=y) andalso (c1=zero) andalso (c2=one) - then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $(norm_zero_one z))) - val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero ,dlcm)) - in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_bst_p_ne))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, T) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel at) andalso (x=y) - then let - val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_number 1))) - in - let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const(@{const_name HOL.minus},T) $ (Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $ norm_zero_one z) $ HOLogic.mk_number HOLogic.intT 1)) - val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero, dlcm)) - in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_bst_p_eq))) - end - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if (y=x) andalso (c1 =zero) then - if pm1 = one then - let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $(norm_zero_one z))) - in (instantiate' [] [SOME cfma, SOME cdlcm]([th1,th2] MRS (not_bst_p_gt))) - end - else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero, dlcm)) - in (instantiate' [] [SOME cfma, SOME cB,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then - let val cz = cterm_of sg (norm_zero_one z) - val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) - in (instantiate' [] [SOME cfma, SOME cB,SOME cz] (th1 RS (not_bst_p_ndvd))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - |(Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then - let val cz = cterm_of sg (norm_zero_one z) - val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) - in (instantiate' [] [SOME cfma,SOME cB,SOME cz] (th1 RS (not_bst_p_dvd))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) - - end; - - -(* ------------------------------------------------------------------------- *) -(* Main interpretation function for this backwards dirction*) -(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*) -(*Help Function*) -(* ------------------------------------------------------------------------- *) - -(*==================== Proof with the compact version *) - -fun decomp_nbstp sg x dlcm B fm t = case t of - Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_conjI ) - |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_disjI) - |_ => ([], fn [] => generate_atomic_not_bst_p sg x fm dlcm B t); - -fun not_bst_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm B t = - let - val th = thm_of sg (decomp_nbstp sg x dlcm (list_to_set xT (map norm_zero_one B)) fm) t - val fma = absfree (xn,xT, norm_zero_one fm) - in let val th1 = prove_elementar sg "ss" (HOLogic.mk_eq (fma,fma)) - in [th,th1] MRS (not_bst_p_Q_elim) - end - end; - - -(* ------------------------------------------------------------------------- *) -(* Protokol interpretation function for the backwards direction for cooper's Theorem*) - -(* For the generation of atomic Theorems*) -(* Prove the premisses on runtime and then make RS*) -(* ------------------------------------------------------------------------- *) -(*========= this is rho ============*) -fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = - let - val cdlcm = cterm_of sg dlcm - val cA = cterm_of sg A - val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) - val cat = cterm_of sg (norm_zero_one at) - in - case at of - (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z))) => - if (x=y) andalso (c1=zero) andalso (c2=one) - then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $(norm_zero_one z))) - val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero, dlcm)) - in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_ast_p_ne))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const (@{const_name HOL.plus}, T) $(Const (@{const_name HOL.times},_) $ c2 $ y) $z)) => - if (is_arith_rel at) andalso (x=y) - then let val ast_z = norm_zero_one (linear_sub [] one z ) - val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const(@{const_name HOL.plus},T) $ (Const(@{const_name HOL.uminus},HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ one)) - val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero, dlcm)) - in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_ast_p_eq))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - |(Const(@{const_name Orderings.less},_) $ c1 $(Const (@{const_name HOL.plus}, _) $(Const (@{const_name HOL.times},_) $ pm1 $ y ) $ z )) => - if (y=x) andalso (c1 =zero) then - if pm1 = (mk_number ~1) then - let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A) - val th2 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero,dlcm)) - in (instantiate' [] [SOME cfma]([th2,th1] MRS (not_ast_p_lt))) - end - else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero, dlcm)) - in (instantiate' [] [SOME cfma, SOME cA,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - |Const ("Not",_) $ (Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then - let val cz = cterm_of sg (norm_zero_one z) - val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) - in (instantiate' [] [SOME cfma, SOME cA,SOME cz] (th1 RS (not_ast_p_ndvd))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - |(Const("Divides.dvd",_)$ d $ (Const (@{const_name HOL.plus},_) $ (Const (@{const_name HOL.times},_) $ c $ y ) $ z)) => - if y=x then - let val cz = cterm_of sg (norm_zero_one z) - val th1 = (prove_elementar sg "ss" (HOLogic.mk_binrel "op =" (HOLogic.mk_binop @{const_name Divides.mod} (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) - in (instantiate' [] [SOME cfma,SOME cA,SOME cz] (th1 RS (not_ast_p_dvd))) - end - else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) - - end; - -(* ------------------------------------------------------------------------ *) -(* Main interpretation function for this backwards dirction*) -(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*) -(*Help Function*) -(* ------------------------------------------------------------------------- *) - -fun decomp_nastp sg x dlcm A fm t = case t of - Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_conjI ) - |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_disjI) - |_ => ([], fn [] => generate_atomic_not_ast_p sg x fm dlcm A t); - -fun not_ast_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm A t = - let - val th = thm_of sg (decomp_nastp sg x dlcm (list_to_set xT (map norm_zero_one A)) fm) t - val fma = absfree (xn,xT, norm_zero_one fm) - in let val th1 = prove_elementar sg "ss" (HOLogic.mk_eq (fma,fma)) - in [th,th1] MRS (not_ast_p_Q_elim) - end - end; - - -(* -------------------------------*) -(* Finding rho and beta for evalc *) -(* -------------------------------*) - -fun rho_for_evalc sg at = case at of - (Const (p,_) $ s $ t) =>( - case AList.lookup (op =) operations p of - SOME f => - ((if (f ((dest_number s),(dest_number t))) - then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)) - else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))) - handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) - | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) - |Const("Not",_)$(Const (p,_) $ s $ t) =>( - case AList.lookup (op =) operations p of - SOME f => - ((if (f ((dest_number s),(dest_number t))) - then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)) - else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const))) - handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) - | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) - | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl; - - -(*=========================================================*) -fun decomp_evalc sg t = case t of - (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI) - |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI) - |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI) - |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI) - |_ => ([], fn [] => rho_for_evalc sg t); - - -fun proof_of_evalc sg fm = thm_of sg (decomp_evalc sg) fm; - -(*==================================================*) -(* Proof of linform with the compact model *) -(*==================================================*) - - -fun decomp_linform sg vars t = case t of - (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI) - |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI) - |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI) - |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI) - |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not) - |(Const("Divides.dvd",_)$d$r) => - if is_number d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd))) - else (warning "Nonlinear Term --- Non numeral leftside at dvd"; - raise COOPER) - |_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t))); - -fun proof_of_linform sg vars f = thm_of sg (decomp_linform sg vars) f; - -(* ------------------------------------------------------------------------- *) -(* Interpretaion of Protocols of the cooper procedure : minusinfinity version*) -(* ------------------------------------------------------------------------- *) -fun coopermi_proof_of sg (x as Free(xn,xT)) fm B dlcm = - (* Get the Bset thm*) - let val (minf_eqth, minf_moddth) = minf_proof_of_c sg x dlcm fm - val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero,dlcm)); - val nbstpthm = not_bst_p_proof_of_c sg x fm dlcm B fm - in (dpos,minf_eqth,nbstpthm,minf_moddth) -end; - -(* ------------------------------------------------------------------------- *) -(* Interpretaion of Protocols of the cooper procedure : plusinfinity version *) -(* ------------------------------------------------------------------------- *) -fun cooperpi_proof_of sg (x as Free(xn,xT)) fm A dlcm = - let val (pinf_eqth,pinf_moddth) = pinf_proof_of_c sg x dlcm fm - val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel @{const_name Orderings.less} (zero,dlcm)); - val nastpthm = not_ast_p_proof_of_c sg x fm dlcm A fm - in (dpos,pinf_eqth,nastpthm,pinf_moddth) -end; - -(* ------------------------------------------------------------------------- *) -(* Interpretaion of Protocols of the cooper procedure : full version*) -(* ------------------------------------------------------------------------- *) -fun cooper_thm sg s (x as Free(xn,xT)) cfm dlcm ast bst= case s of - "pi" => let val (dpsthm,pinf_eqth,nbpth,pinf_moddth) = cooperpi_proof_of sg x cfm ast dlcm - in [dpsthm,pinf_eqth,nbpth,pinf_moddth] MRS (cppi_eq) - end - |"mi" => let val (dpsthm,minf_eqth,nbpth,minf_moddth) = coopermi_proof_of sg x cfm bst dlcm - in [dpsthm,minf_eqth,nbpth,minf_moddth] MRS (cpmi_eq) - end - |_ => error "parameter error"; - -(* ------------------------------------------------------------------------- *) -(* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*) -(* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*) -(* ------------------------------------------------------------------------- *) - -(* val (timef:(unit->thm) -> thm,prtime,time_reset) = gen_timer();*) -(* val (timef2:(unit->thm) -> thm,prtime2,time_reset2) = gen_timer(); *) - -fun cooper_prv sg (x as Free(xn,xT)) efm = let - (* lfm_thm : efm = linearized form of efm*) - val lfm_thm = proof_of_linform sg [xn] efm - (*efm2 is the linearized form of efm *) - val efm2 = snd(qe_get_terms lfm_thm) - (* l is the lcm of all coefficients of x *) - val l = formlcm x efm2 - (*ac_thm: efm = efm2 with adjusted coefficients of x *) - val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans - (* fm is efm2 with adjusted coefficients of x *) - val fm = snd (qe_get_terms ac_thm) - (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*) - val cfm = unitycoeff x fm - (*afm is fm where c*x is replaced by 1*x or -1*x *) - val afm = adjustcoeff x l fm - (* P = %x.afm*) - val P = absfree(xn,xT,afm) - (* This simpset allows the elimination of the sets in bex {1..d} *) - val ss = presburger_ss addsimps - [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff] - (* uth : EX x.P(l*x) = EX x. l dvd x & P x*) - val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_number l))] (unity_coeff_ex) - (* e_ac_thm : Ex x. efm = EX x. fm*) - val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI) - (* A and B set of the formula*) - val A = aset x cfm - val B = bset x cfm - (* the divlcm (delta) of the formula*) - val dlcm = mk_number (divlcm x cfm) - (* Which set is smaller to generate the (hoepfully) shorter proof*) - val cms = if ((length A) < (length B )) then "pi" else "mi" -(* val _ = if cms = "pi" then writeln "Plusinfinity" else writeln "Minusinfinity"*) - (* synthesize the proof of cooper's theorem*) - (* cp_thm: EX x. cfm = Q*) - val cp_thm = cooper_thm sg cms x cfm dlcm A B - (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*) - (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*) -(* - val _ = prth cp_thm - val _ = writeln "Expanding the bounded EX..." -*) - val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans))) -(* - val _ = writeln "Expanded" *) - (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*) - val (lsuth,rsuth) = qe_get_terms (uth) - (* lseacth = EX x. efm; rseacth = EX x. fm*) - val (lseacth,rseacth) = qe_get_terms(e_ac_thm) - (* lscth = EX x. cfm; rscth = Q' *) - val (lscth,rscth) = qe_get_terms (exp_cp_thm) - (* u_c_thm: EX x. P(l*x) = Q'*) - val u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans - (* result: EX x. efm = Q'*) - in ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans) - end -|cooper_prv _ _ _ = error "Parameters format"; - -(* **************************************** *) -(* An Other Version of cooper proving *) -(* by giving a withness for EX *) -(* **************************************** *) - - - -fun cooper_prv_w sg (x as Free(xn,xT)) efm = let - (* lfm_thm : efm = linearized form of efm*) - val lfm_thm = proof_of_linform sg [xn] efm - (*efm2 is the linearized form of efm *) - val efm2 = snd(qe_get_terms lfm_thm) - (* l is the lcm of all coefficients of x *) - val l = formlcm x efm2 - (*ac_thm: efm = efm2 with adjusted coefficients of x *) - val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans - (* fm is efm2 with adjusted coefficients of x *) - val fm = snd (qe_get_terms ac_thm) - (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*) - val cfm = unitycoeff x fm - (*afm is fm where c*x is replaced by 1*x or -1*x *) - val afm = adjustcoeff x l fm - (* P = %x.afm*) - val P = absfree(xn,xT,afm) - (* This simpset allows the elimination of the sets in bex {1..d} *) - val ss = presburger_ss addsimps - [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff] - (* uth : EX x.P(l*x) = EX x. l dvd x & P x*) - val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_number l))] (unity_coeff_ex) - (* e_ac_thm : Ex x. efm = EX x. fm*) - val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI) - (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*) - val (lsuth,rsuth) = qe_get_terms (uth) - (* lseacth = EX x. efm; rseacth = EX x. fm*) - val (lseacth,rseacth) = qe_get_terms(e_ac_thm) - - val (w,rs) = cooper_w [] cfm - val exp_cp_thm = case w of - (* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*) - SOME n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*) - |_ => let - (* A and B set of the formula*) - val A = aset x cfm - val B = bset x cfm - (* the divlcm (delta) of the formula*) - val dlcm = mk_number (divlcm x cfm) - (* Which set is smaller to generate the (hoepfully) shorter proof*) - val cms = if ((length A) < (length B )) then "pi" else "mi" - (* synthesize the proof of cooper's theorem*) - (* cp_thm: EX x. cfm = Q*) - val cp_thm = cooper_thm sg cms x cfm dlcm A B - (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*) - (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*) - in refl RS (simplify ss (cp_thm RSN (2,trans))) - end - (* lscth = EX x. cfm; rscth = Q' *) - val (lscth,rscth) = qe_get_terms (exp_cp_thm) - (* u_c_thm: EX x. P(l*x) = Q'*) - val u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans - (* result: EX x. efm = Q'*) - in ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans) - end -|cooper_prv_w _ _ _ = error "Parameters format"; - - - -fun decomp_cnnf sg lfnp P = case P of - Const ("op &",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS qe_conjI ) - |Const ("op |",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS qe_disjI) - |Const ("Not",_) $ (Const("Not",_) $ p) => ([p], fn [th] => th RS nnf_nn) - |Const("Not",_) $ (Const(opn,T) $ p $ q) => - if opn = "op |" - then case (p,q) of - (A as (Const ("op &",_) $ r $ s),B as (Const ("op &",_) $ r1 $ t)) => - if r1 = negate r - then ([r,HOLogic.Not$s,r1,HOLogic.Not$t],fn [th1_1,th1_2,th2_1,th2_2] => [[th1_1,th1_1] MRS qe_conjI,[th2_1,th2_2] MRS qe_conjI] MRS nnf_sdj) - - else ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj) - |(_,_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj) - else ( - case (opn,T) of - ("op &",_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_ncj ) - |("op -->",_) => ([p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_nim ) - |("op =",Type ("fun",[Type ("bool", []),_])) => - ([HOLogic.conj $ p $ (HOLogic.Not $ q),HOLogic.conj $ (HOLogic.Not $ p) $ q], fn [th1,th2] => [th1,th2] MRS nnf_neq) - |(_,_) => ([], fn [] => lfnp P) -) - - |(Const ("op -->",_) $ p $ q) => ([HOLogic.Not$p,q], fn [th1,th2] => [th1,th2] MRS nnf_im) - - |(Const ("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q) => - ([HOLogic.conj $ p $ q,HOLogic.conj $ (HOLogic.Not $ p) $ (HOLogic.Not $ q) ], fn [th1,th2] =>[th1,th2] MRS nnf_eq ) - |_ => ([], fn [] => lfnp P); - - - - -fun proof_of_cnnf sg p lfnp = - let val th1 = thm_of sg (decomp_cnnf sg lfnp) p - val rs = snd(qe_get_terms th1) - val th2 = prove_elementar sg "ss" (HOLogic.mk_eq(rs,simpl rs)) - in [th1,th2] MRS trans - end; - -end; -
--- a/src/HOL/Integ/presburger.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,372 +0,0 @@ -(* Title: HOL/Integ/presburger.ML - ID: $Id$ - Author: Amine Chaieb and Stefan Berghofer, TU Muenchen - -Tactic for solving arithmetical Goals in Presburger Arithmetic. - -This version of presburger deals with occurences of functional symbols -in the subgoal and abstract over them to try to prove the more general -formula. It then resolves with the subgoal. To enable this feature -call the procedure with the parameter abs. -*) - -signature PRESBURGER = -sig - val presburger_tac : bool -> bool -> int -> tactic - val setup : theory -> theory - val trace : bool ref -end; - -structure Presburger: PRESBURGER = -struct - -val trace = ref false; -fun trace_msg s = if !trace then tracing s else (); - -(*-----------------------------------------------------------------*) -(*cooper_pp: provefunction for the one-existance quantifier elimination*) -(* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) -(*-----------------------------------------------------------------*) - - -val presburger_ss = simpset (); -val binarith = map thm - ["Pls_0_eq", "Min_1_eq", - "pred_Pls","pred_Min","pred_1","pred_0", - "succ_Pls", "succ_Min", "succ_1", "succ_0", - "add_Pls", "add_Min", "add_BIT_0", "add_BIT_10", - "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", - "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", - "add_Pls_right", "add_Min_right"]; - val intarithrel = - (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", - "int_le_number_of_eq","int_iszero_number_of_0", - "int_less_number_of_eq_neg"]) @ - (map (fn s => thm s RS thm "lift_bool") - ["int_iszero_number_of_Pls","int_iszero_number_of_1", - "int_neg_number_of_Min"])@ - (map (fn s => thm s RS thm "nlift_bool") - ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); - -val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym", - "int_number_of_diff_sym", "int_number_of_mult_sym"]; -val natarith = map thm ["add_nat_number_of", "diff_nat_number_of", - "mult_nat_number_of", "eq_nat_number_of", - "less_nat_number_of"] -val powerarith = - (map thm ["nat_number_of", "zpower_number_of_even", - "zpower_Pls", "zpower_Min"]) @ - [(MetaSimplifier.simplify true [thm "zero_eq_Numeral0_nring", - thm "one_eq_Numeral1_nring"] - (thm "zpower_number_of_odd"))] - -val comp_arith = binarith @ intarith @ intarithrel @ natarith - @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; - -fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = - let val (xn1,p1) = Syntax.variant_abs (xn,xT,p) - in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; - -fun mnnf_pp sg fm = CooperProof.proof_of_cnnf sg fm - (CooperProof.proof_of_evalc sg); - -fun tmproof_of_int_qelim sg fm = - Qelim.tproof_of_mlift_qelim sg CooperDec.is_arith_rel - (CooperProof.proof_of_linform sg) (mnnf_pp sg) (cooper_pp sg) fm; - - -(* Theorems to be used in this tactic*) - -val zdvd_int = thm "zdvd_int"; -val zdiff_int_split = thm "zdiff_int_split"; -val all_nat = thm "all_nat"; -val ex_nat = thm "ex_nat"; -val number_of1 = thm "number_of1"; -val number_of2 = thm "number_of2"; -val split_zdiv = thm "split_zdiv"; -val split_zmod = thm "split_zmod"; -val mod_div_equality' = thm "mod_div_equality'"; -val split_div' = thm "split_div'"; -val Suc_plus1 = thm "Suc_plus1"; -val imp_le_cong = thm "imp_le_cong"; -val conj_le_cong = thm "conj_le_cong"; -val nat_mod_add_eq = mod_add1_eq RS sym; -val nat_mod_add_left_eq = mod_add_left_eq RS sym; -val nat_mod_add_right_eq = mod_add_right_eq RS sym; -val int_mod_add_eq = @{thm zmod_zadd1_eq} RS sym; -val int_mod_add_left_eq = @{thm zmod_zadd_left_eq} RS sym; -val int_mod_add_right_eq = @{thm zmod_zadd_right_eq} RS sym; -val nat_div_add_eq = @{thm div_add1_eq} RS sym; -val int_div_add_eq = @{thm zdiv_zadd1_eq} RS sym; -val ZDIVISION_BY_ZERO_MOD = @{thm DIVISION_BY_ZERO} RS conjunct2; -val ZDIVISION_BY_ZERO_DIV = @{thm DIVISION_BY_ZERO} RS conjunct1; - - -(* extract all the constants in a term*) -fun add_term_typed_consts (Const (c, T), cs) = insert (op =) (c, T) cs - | add_term_typed_consts (t $ u, cs) = - add_term_typed_consts (t, add_term_typed_consts (u, cs)) - | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs) - | add_term_typed_consts (_, cs) = cs; - -fun term_typed_consts t = add_term_typed_consts(t,[]); - -(* Some Types*) -val bT = HOLogic.boolT; -val bitT = HOLogic.bitT; -val iT = HOLogic.intT; -val nT = HOLogic.natT; - -(* Allowed Consts in formulae for presburger tactic*) - -val allowed_consts = - [("All", (iT --> bT) --> bT), - ("Ex", (iT --> bT) --> bT), - ("All", (nT --> bT) --> bT), - ("Ex", (nT --> bT) --> bT), - - ("op &", bT --> bT --> bT), - ("op |", bT --> bT --> bT), - ("op -->", bT --> bT --> bT), - ("op =", bT --> bT --> bT), - ("Not", bT --> bT), - - (@{const_name Orderings.less_eq}, iT --> iT --> bT), - ("op =", iT --> iT --> bT), - (@{const_name Orderings.less}, iT --> iT --> bT), - (@{const_name Divides.dvd}, iT --> iT --> bT), - (@{const_name Divides.div}, iT --> iT --> iT), - (@{const_name Divides.mod}, iT --> iT --> iT), - (@{const_name HOL.plus}, iT --> iT --> iT), - (@{const_name HOL.minus}, iT --> iT --> iT), - (@{const_name HOL.times}, iT --> iT --> iT), - (@{const_name HOL.abs}, iT --> iT), - (@{const_name HOL.uminus}, iT --> iT), - (@{const_name Orderings.max}, iT --> iT --> iT), - (@{const_name Orderings.min}, iT --> iT --> iT), - - (@{const_name Orderings.less_eq}, nT --> nT --> bT), - ("op =", nT --> nT --> bT), - (@{const_name Orderings.less}, nT --> nT --> bT), - (@{const_name Divides.dvd}, nT --> nT --> bT), - (@{const_name Divides.div}, nT --> nT --> nT), - (@{const_name Divides.mod}, nT --> nT --> nT), - (@{const_name HOL.plus}, nT --> nT --> nT), - (@{const_name HOL.minus}, nT --> nT --> nT), - (@{const_name HOL.times}, nT --> nT --> nT), - (@{const_name Suc}, nT --> nT), - (@{const_name Orderings.max}, nT --> nT --> nT), - (@{const_name Orderings.min}, nT --> nT --> nT), - - (@{const_name Numeral.bit.B0}, bitT), - (@{const_name Numeral.bit.B1}, bitT), - (@{const_name Numeral.Bit}, iT --> bitT --> iT), - (@{const_name Numeral.Pls}, iT), - (@{const_name Numeral.Min}, iT), - (@{const_name Numeral.number_of}, iT --> iT), - (@{const_name Numeral.number_of}, iT --> nT), - (@{const_name HOL.zero}, nT), - (@{const_name HOL.zero}, iT), - (@{const_name HOL.one}, nT), - (@{const_name HOL.one}, iT), - (@{const_name False}, bT), - (@{const_name True}, bT)]; - -(* Preparation of the formula to be sent to the Integer quantifier *) -(* elimination procedure *) -(* Transforms meta implications and meta quantifiers to object *) -(* implications and object quantifiers *) - - -(*==================================*) -(* Abstracting on subterms ========*) -(*==================================*) -(* Returns occurences of terms that are function application of type int or nat*) - -fun getfuncs fm = case strip_comb fm of - (Free (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] - andalso not (ts = []) andalso forall (null o loose_bnos) ts - then [fm] - else Library.foldl op union ([], map getfuncs ts) - | (Var (_, T), ts as _ :: _) => - if body_type T mem [iT, nT] - andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] - else Library.foldl op union ([], map getfuncs ts) - | (Const (s, T), ts) => - if (s, T) mem allowed_consts orelse not (body_type T mem [iT, nT]) - then Library.foldl op union ([], map getfuncs ts) - else [fm] - | (Abs (s, T, t), _) => getfuncs t - | _ => []; - - -fun abstract_pres sg fm = - foldr (fn (t, u) => - let val T = fastype_of t - in all T $ Abs ("x", T, abstract_over (t, u)) end) - fm (getfuncs fm); - - - -(* hasfuncs_on_bounds dont care of the type of the functions applied! - It returns true if there is a subterm coresponding to the application of - a function on a bounded variable. - - Function applications are allowed only for well predefined functions a - consts*) - -fun has_free_funcs fm = case strip_comb fm of - (Free (_, T), ts as _ :: _) => - if (body_type T mem [iT,nT]) andalso (not (T mem [iT,nT])) - then true - else exists (fn x => x) (map has_free_funcs ts) - | (Var (_, T), ts as _ :: _) => - if (body_type T mem [iT,nT]) andalso not (T mem [iT,nT]) - then true - else exists (fn x => x) (map has_free_funcs ts) - | (Const (s, T), ts) => exists (fn x => x) (map has_free_funcs ts) - | (Abs (s, T, t), _) => has_free_funcs t - |_ => false; - - -(*returns true if the formula is relevant for presburger arithmetic tactic -The constants occuring in term t should be a subset of the allowed_consts - There also should be no occurences of application of functions on bounded - variables. Whenever this function will be used, it will be ensured that t - will not contain subterms with function symbols that could have been - abstracted over.*) - -fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso - map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @ - map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t) - subset [iT, nT] - andalso not (has_free_funcs t); - - -fun prepare_for_presburger sg q fm = - let - val ps = Logic.strip_params fm - val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) - val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm) - val _ = if relevant (rev ps) c then () - else (trace_msg ("Conclusion is not a presburger term:\n" ^ - Sign.string_of_term sg c); raise CooperDec.COOPER) - fun mk_all ((s, T), (P,n)) = - if 0 mem loose_bnos P then - (HOLogic.all_const T $ Abs (s, T, P), n) - else (incr_boundvars ~1 P, n-1) - fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; - val (rhs,irhs) = List.partition (relevant (rev ps)) hs - val np = length ps - val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) - (foldr HOLogic.mk_imp c rhs, np) ps - val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) - (term_frees fm' @ term_vars fm'); - val fm2 = foldr mk_all2 fm' vs - in (fm2, np + length vs, length rhs) end; - -(*Object quantifier to meta --*) -fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; - -(* object implication to meta---*) -fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; - -(* the presburger tactic*) - -(* Parameters : q = flag for quantify ofer free variables ; - a = flag for abstracting over function occurences - i = subgoal *) - -fun presburger_tac q a i = ObjectLogic.atomize_tac i THEN (fn st => - let - val g = List.nth (prems_of st, i - 1) - val sg = Thm.theory_of_thm st - (* The Abstraction step *) - val g' = if a then abstract_pres sg g else g - (* Transform the term*) - val (t,np,nh) = prepare_for_presburger sg q g' - (* Some simpsets for dealing with mod div abs and nat*) - val mod_div_simpset = Simplifier.theory_context sg HOL_basic_ss - addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, - nat_mod_add_right_eq, int_mod_add_eq, - int_mod_add_right_eq, int_mod_add_left_eq, - nat_div_add_eq, int_div_add_eq, - mod_self, @{thm zmod_self}, - DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, - ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, - @{thm zdiv_zero}, @{thm zmod_zero}, div_0,mod_0, - @{thm zdiv_1}, @{thm zmod_1}, @{thm div_1}, @{thm mod_1}, - Suc_plus1] - addsimps add_ac - addsimprocs [cancel_div_mod_proc] - val simpset0 = HOL_basic_ss - addsimps [@{thm mod_div_equality'}, @{thm Suc_plus1}] - addsimps comp_arith - addsplits [split_zdiv, split_zmod, split_div', @{thm split_min}, @{thm split_max}] - (* Simp rules for changing (n::int) to int n *) - val simpset1 = HOL_basic_ss - addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym) - [int_int_eq, zle_int, zless_int, zadd_int, zmult_int] - addsplits [zdiff_int_split] - (*simp rules for elimination of int n*) - - val simpset2 = HOL_basic_ss - addsimps [nat_0_le, all_nat, ex_nat, number_of1, number_of2, int_0, int_1] - addcongs [conj_le_cong, imp_le_cong] - (* simp rules for elimination of abs *) - val simpset3 = HOL_basic_ss addsplits [abs_split] - val ct = cterm_of sg (HOLogic.mk_Trueprop t) - (* Theorem for the nat --> int transformation *) - val pre_thm = Seq.hd (EVERY - [simp_tac mod_div_simpset 1, simp_tac simpset0 1, - TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), - TRY (simp_tac simpset3 1), TRY (simp_tac presburger_ss 1)] - (trivial ct)) - fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) - (* The result of the quantifier elimination *) - val (th, tac) = case (prop_of pre_thm) of - Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => - let val pth = - (* If quick_and_dirty then run without proof generation as oracle*) - if !quick_and_dirty - then presburger_oracle sg (Pattern.eta_long [] t1) -(* -assume (cterm_of sg - (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) -*) - else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) - in - (trace_msg ("calling procedure with term:\n" ^ - Sign.string_of_term sg t1); - ((pth RS iffD2) RS pre_thm, - assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) - end - | _ => (pre_thm, assm_tac i) - in (rtac (((mp_step nh) o (spec_step np)) th) i - THEN tac) st - end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st); - -val presburger_meth = - let val parse_flag = - Args.$$$ "no_quantify" >> K (apfst (K false)) - || Args.$$$ "no_abs" >> K (apsnd (K false)); - in - Method.simple_args - (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >> - curry (Library.foldl op |>) (true, true)) - (fn (q,a) => K (Method.SIMPLE_METHOD' (presburger_tac q a))) - end; - -val presburger_arith_tac = mk_arith_tactic "presburger" (fn i => fn st => - (warning "Trying full Presburger arithmetic ..."; - presburger_tac true true i st)); - -val setup = - Method.add_method ("presburger", presburger_meth, - "decision procedure for Presburger arithmetic") #> - arith_tactic_add presburger_arith_tac; - -end; - -val presburger_tac = Presburger.presburger_tac true true;
--- a/src/HOL/Integ/qelim.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,110 +0,0 @@ -(* Title: HOL/Integ/qelim.ML - ID: $Id$ - Author: Amine Chaieb and Tobias Nipkow, TU Muenchen - -File containing the implementation of the proof protocoling -and proof generation for multiple quantified formulae. -*) - -signature QELIM = - sig - val tproof_of_mlift_qelim: theory -> (term -> bool) -> - (string list -> term -> thm) -> (term -> thm) -> - (term -> thm) -> term -> thm - val standard_qelim_conv: (cterm list -> cterm -> thm) -> - (cterm list -> Conv.conv) -> (cterm list -> cterm -> thm) -> cterm -> thm - val gen_qelim_conv: Conv.conv -> Conv.conv -> Conv.conv -> - (cterm -> 'a -> 'a) -> 'a -> ('a -> cterm -> thm) -> - ('a -> Conv.conv) -> ('a -> cterm -> thm) -> Conv.conv - -end; - -structure Qelim : QELIM = -struct -open CooperDec; -open CooperProof; -open Conv; - -val cboolT = ctyp_of HOL.thy HOLogic.boolT; - -(* List of the theorems to be used in the following*) - -val qe_ex_conj = thm "qe_ex_conj"; -val qe_ex_nconj = thm "qe_ex_nconj"; -val qe_ALL = thm "qe_ALL"; - - -(*============= Compact version =====*) - - -fun decomp_qe is_at afnp nfnp qfnp sg P = - if is_at P then ([], fn [] => afnp [] P) else - case P of - (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI) - |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI) - |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI) - |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI) - |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not) - |(Const("Ex",_)$Abs(xn,xT,p)) => - let val (xn1,p1) = Syntax.variant_abs(xn,xT,p) - in ([p1], - fn [th1_1] => - let val th2 = [th1_1,nfnp (snd (qe_get_terms th1_1))] MRS trans - val eth1 = (forall_intr (cterm_of sg (Free(xn1,xT))) th2) COMP qe_exI - val th3 = qfnp (snd(qe_get_terms eth1)) - in [eth1,th3] MRS trans - end ) - end - - |(Const("All",_)$Abs(xn,xT,p)) => ([(HOLogic.exists_const xT)$Abs(xn,xT,HOLogic.Not $ p)], fn [th] => th RS qe_ALL) - | _ => ([],fn [] => instantiate' [SOME (ctyp_of sg (type_of P))] [SOME (cterm_of sg P)] refl); - - -fun tproof_of_mlift_qelim sg isat afnp nfnp qfnp p = - let val th1 = thm_of sg (decomp_qe isat afnp nfnp qfnp sg) p - val th2 = nfnp (snd (qe_get_terms th1)) - in [th1,th2] MRS trans - end; - -val is_refl = op aconv o Logic.dest_equals o Thm.prop_of; - -fun gen_qelim_conv precv postcv simpex_conv ins env atcv ncv qcv = - let fun conv p = - case (term_of p) of - Const(s,T)$_$_ => if domain_type T = HOLogic.boolT - andalso s mem ["op &","op |","op -->","op ="] - then binop_conv conv p else atcv env p - | Const("Not",_)$_ => arg_conv conv p - | Const("Ex",_)$Abs(s,_,_) => - let - val (e,p0) = Thm.dest_comb p - val (x,p') = Thm.dest_abs (SOME s) p0 - val th = Thm.abstract_rule s x - (((gen_qelim_conv precv postcv simpex_conv ins (ins x env) atcv ncv qcv) - then_conv (ncv (ins x env))) p') - |> Drule.arg_cong_rule e - val th' = simpex_conv (Thm.rhs_of th) - val (l,r) = Thm.dest_equals (cprop_of th') - in if is_refl th' then Thm.transitive th (qcv env (Thm.rhs_of th)) - else Thm.transitive (Thm.transitive th th') (conv r) end - | _ => atcv env p - in precv then_conv conv then_conv postcv end; - -fun cterm_frees ct = - let fun h acc t = - case (term_of t) of - _$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) - | Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc)) - | Free _ => insert (op aconvc) t acc - | _ => acc - in h [] ct end; - -val standard_qelim_conv = - let val pcv = Simplifier.rewrite - (HOL_basic_ss addsimps (simp_thms @ (List.take(ex_simps,4)) - @ [not_all,ex_disj_distrib])) - in fn atcv => fn ncv => fn qcv => fn p => - gen_qelim_conv pcv pcv pcv (curry (op ::)) (cterm_frees p) atcv ncv qcv p - end; - -end;
--- a/src/HOL/Integ/reflected_cooper.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,119 +0,0 @@ -(* $Id$ *) -(* The oracle for Presburger arithmetic based on the verified Code *) - (* in HOL/ex/Reflected_Presburger.thy *) - -structure ReflectedCooper = -struct - -open Generated; -(* pseudo reification : term -> intterm *) - -fun i_of_term vs t = case t of - Free(xn,xT) => (case AList.lookup (op =) vs t of - NONE => error "Variable not found in the list!!" - | SOME n => Var n) - | Const(@{const_name HOL.zero},iT) => Cst 0 - | Const(@{const_name HOL.one},iT) => Cst 1 - | Bound i => Var (nat (IntInf.fromInt i)) - | Const(@{const_name HOL.uminus},_)$t' => Neg (i_of_term vs t') - | Const (@{const_name HOL.plus},_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2) - | Const (@{const_name HOL.minus},_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2) - | Const (@{const_name HOL.times},_)$t1$t2 => Mult (i_of_term vs t1,i_of_term vs t2) - | Const (@{const_name Numeral.number_of},_)$t' => Cst (HOLogic.dest_numeral t') - | _ => error "i_of_term: unknown term"; - -(* pseudo reification : term -> QF *) -fun qf_of_term vs t = case t of - Const("True",_) => T - | Const("False",_) => F - | Const(@{const_name Orderings.less},_)$t1$t2 => Lt (i_of_term vs t1,i_of_term vs t2) - | Const(@{const_name Orderings.less_eq},_)$t1$t2 => Le (i_of_term vs t1,i_of_term vs t2) - | Const ("Divides.dvd",_)$t1$t2 => - Divides(i_of_term vs t1,i_of_term vs t2) - | Const("op =",eqT)$t1$t2 => - if (domain_type eqT = HOLogic.intT) - then let val i1 = i_of_term vs t1 - val i2 = i_of_term vs t2 - in Eq (i1,i2) - end - else Equ(qf_of_term vs t1,qf_of_term vs t2) - | Const("op &",_)$t1$t2 => And(qf_of_term vs t1,qf_of_term vs t2) - | Const("op |",_)$t1$t2 => Or(qf_of_term vs t1,qf_of_term vs t2) - | Const("op -->",_)$t1$t2 => Imp(qf_of_term vs t1,qf_of_term vs t2) - | Const("Not",_)$t' => NOT(qf_of_term vs t') - | Const("Ex",_)$Abs(xn,xT,p) => - QEx(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) - | Const("All",_)$Abs(xn,xT,p) => - QAll(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) - | _ => error "qf_of_term : unknown term!"; - -(* -fun parse thy s = term_of (Thm.read_cterm thy (s, HOLogic.boolT)); - -val t = parse "ALL (i::int) (j::int). i < 8* j --> (i - 1 = j + 3 + 2*j) & (j <= -i + k ) | 4 = i | 5 dvd i"; -*) -fun zip [] [] = [] - | zip (x::xs) (y::ys) = (x,y)::(zip xs ys); - - -fun start_vs t = - let val fs = term_frees t - in zip fs (map (nat o IntInf.fromInt) (0 upto (length fs - 1))) - end ; - -(* transform intterm and QF back to terms *) -val iT = HOLogic.intT; -val bT = HOLogic.boolT; -fun myassoc2 l v = - case l of - [] => NONE - | (x,v')::xs => if v = v' then SOME x - else myassoc2 xs v; - -fun term_of_i vs t = - case t of - Cst i => CooperDec.mk_number i - | Var n => valOf (myassoc2 vs n) - | Neg t' => Const(@{const_name HOL.uminus},iT --> iT)$(term_of_i vs t') - | Add(t1,t2) => Const(@{const_name HOL.plus},[iT,iT] ---> iT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | Sub(t1,t2) => Const(@{const_name HOL.minus},[iT,iT] ---> iT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | Mult(t1,t2) => Const(@{const_name HOL.times},[iT,iT] ---> iT)$ - (term_of_i vs t1)$(term_of_i vs t2); - -fun term_of_qf vs t = - case t of - T => HOLogic.true_const - | F => HOLogic.false_const - | Lt(t1,t2) => Const(@{const_name Orderings.less},[iT,iT] ---> bT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | Le(t1,t2) => Const(@{const_name Orderings.less_eq},[iT,iT] ---> bT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | Gt(t1,t2) => Const(@{const_name Orderings.less},[iT,iT] ---> bT)$ - (term_of_i vs t2)$(term_of_i vs t1) - | Ge(t1,t2) => Const(@{const_name Orderings.less_eq},[iT,iT] ---> bT)$ - (term_of_i vs t2)$(term_of_i vs t1) - | Eq(t1,t2) => Const("op =",[iT,iT] ---> bT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | Divides(t1,t2) => Const("Divides.dvd",[iT,iT] ---> bT)$ - (term_of_i vs t1)$(term_of_i vs t2) - | NOT t' => HOLogic.Not$(term_of_qf vs t') - | And(t1,t2) => HOLogic.conj$(term_of_qf vs t1)$(term_of_qf vs t2) - | Or(t1,t2) => HOLogic.disj$(term_of_qf vs t1)$(term_of_qf vs t2) - | Imp(t1,t2) => HOLogic.imp$(term_of_qf vs t1)$(term_of_qf vs t2) - | Equ(t1,t2) => (HOLogic.eq_const bT)$(term_of_qf vs t1)$ - (term_of_qf vs t2) - | _ => error "If this is raised, Isabelle/HOL or generate_code is inconsistent!"; - -(* The oracle *) -fun presburger_oracle thy t = - let val vs = start_vs t - val result = lift_un (term_of_qf vs) (pa (qf_of_term vs t)) - in - case result of - None => raise CooperDec.COOPER - | Some t' => HOLogic.mk_Trueprop (HOLogic.mk_eq(t,t')) - end ; - -end;
--- a/src/HOL/Integ/reflected_presburger.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2172 +0,0 @@ -(* $Id$ *) - - (* Caution: This file should not be modified. *) - (* It is autmatically generated from HOL/ex/Reflected_Presburger.thy *) -fun nat (i:IntInf.int) = if i < 0 then 0 else i : IntInf.int; -structure Generated = -struct - -datatype intterm = Cst of IntInf.int | Var of IntInf.int | Neg of intterm - | Add of intterm * intterm | Sub of intterm * intterm - | Mult of intterm * intterm; - -datatype QF = Lt of intterm * intterm | Gt of intterm * intterm - | Le of intterm * intterm | Ge of intterm * intterm | Eq of intterm * intterm - | Divides of intterm * intterm | T | F | NOT of QF | And of QF * QF - | Or of QF * QF | Imp of QF * QF | Equ of QF * QF | QAll of QF | QEx of QF; - -datatype 'a option = None | Some of 'a; - -fun lift_un c None = None - | lift_un c (Some p) = Some (c p); - -fun lift_bin (c, (Some a, Some b)) = Some (c a b) - | lift_bin (c, (None, y)) = None - | lift_bin (c, (Some y, None)) = None; - -fun lift_qe qe None = None - | lift_qe qe (Some p) = qe p; - -fun qelim (qe, QAll p) = lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe, p)))) - | qelim (qe, QEx p) = lift_qe qe (qelim (qe, p)) - | qelim (qe, And (p, q)) = - lift_bin ((fn x => fn xa => And (x, xa)), (qelim (qe, p), qelim (qe, q))) - | qelim (qe, Or (p, q)) = - lift_bin ((fn x => fn xa => Or (x, xa)), (qelim (qe, p), qelim (qe, q))) - | qelim (qe, Imp (p, q)) = - lift_bin ((fn x => fn xa => Imp (x, xa)), (qelim (qe, p), qelim (qe, q))) - | qelim (qe, Equ (p, q)) = - lift_bin ((fn x => fn xa => Equ (x, xa)), (qelim (qe, p), qelim (qe, q))) - | qelim (qe, NOT p) = lift_un NOT (qelim (qe, p)) - | qelim (qe, Lt (w, x)) = Some (Lt (w, x)) - | qelim (qe, Gt (y, z)) = Some (Gt (y, z)) - | qelim (qe, Le (aa, ab)) = Some (Le (aa, ab)) - | qelim (qe, Ge (ac, ad)) = Some (Ge (ac, ad)) - | qelim (qe, Eq (ae, af)) = Some (Eq (ae, af)) - | qelim (qe, Divides (ag, ah)) = Some (Divides (ag, ah)) - | qelim (qe, T) = Some T - | qelim (qe, F) = Some F; - -fun lin_mul (c, Cst i) = Cst (c * i) - | lin_mul (c, Add (Mult (Cst c', Var n), r)) = - (if (c = 0) then Cst 0 - else Add (Mult (Cst (c * c'), Var n), lin_mul (c, r))); - -fun op_60_def0 m n = IntInf.< (m,n); - -fun op_60_61_def0 m n = not (op_60_def0 n m); - -fun lin_add (Add (Mult (Cst c1, Var n1), r1), Add (Mult (Cst c2, Var n2), r2)) = - (if (n1 = n2) - then let val c = Cst (c1 + c2) - in (if ((c1 + c2) = 0) then lin_add (r1, r2) - else Add (Mult (c, Var n1), lin_add (r1, r2))) - end - else (if op_60_61_def0 n1 n2 - then Add (Mult (Cst c1, Var n1), - lin_add (r1, Add (Mult (Cst c2, Var n2), r2))) - else Add (Mult (Cst c2, Var n2), - lin_add (Add (Mult (Cst c1, Var n1), r1), r2)))) - | lin_add (Add (Mult (Cst c1, Var n1), r1), Cst b) = - Add (Mult (Cst c1, Var n1), lin_add (r1, Cst b)) - | lin_add (Cst x, Add (Mult (Cst c2, Var n2), r2)) = - Add (Mult (Cst c2, Var n2), lin_add (Cst x, r2)) - | lin_add (Cst b1, Cst b2) = Cst (b1 + b2); - -fun lin_neg i = lin_mul (~1, i); - -fun linearize (Cst b) = Some (Cst b) - | linearize (Var n) = Some (Add (Mult (Cst 1, Var n), Cst 0)) - | linearize (Neg i) = lift_un lin_neg (linearize i) - | linearize (Add (i, j)) = - lift_bin ((fn x => fn y => lin_add (x, y)), (linearize i, linearize j)) - | linearize (Sub (i, j)) = - lift_bin - ((fn x => fn y => lin_add (x, lin_neg y)), (linearize i, linearize j)) - | linearize (Mult (i, j)) = - (case linearize i of None => None - | Some x => - (case x of - Cst xa => - (case linearize j of None => None - | Some x => Some (lin_mul (xa, x))) - | Var xa => - (case linearize j of None => None - | Some xa => - (case xa of Cst xa => Some (lin_mul (xa, x)) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)) - | Neg xa => - (case linearize j of None => None - | Some xa => - (case xa of Cst xa => Some (lin_mul (xa, x)) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)) - | Add (xa, xb) => - (case linearize j of None => None - | Some xa => - (case xa of Cst xa => Some (lin_mul (xa, x)) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)) - | Sub (xa, xb) => - (case linearize j of None => None - | Some xa => - (case xa of Cst xa => Some (lin_mul (xa, x)) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)) - | Mult (xa, xb) => - (case linearize j of None => None - | Some xa => - (case xa of Cst xa => Some (lin_mul (xa, x)) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)))); - -fun linform (Le (it1, it2)) = - lift_bin - ((fn x => fn y => Le (lin_add (x, lin_neg y), Cst 0)), - (linearize it1, linearize it2)) - | linform (Eq (it1, it2)) = - lift_bin - ((fn x => fn y => Eq (lin_add (x, lin_neg y), Cst 0)), - (linearize it1, linearize it2)) - | linform (Divides (d, t)) = - (case linearize d of None => None - | Some x => - (case x of - Cst xa => - (if (xa = 0) then None - else (case linearize t of None => None - | Some xa => Some (Divides (x, xa)))) - | Var xa => None | Neg xa => None | Add (xa, xb) => None - | Sub (xa, xb) => None | Mult (xa, xb) => None)) - | linform T = Some T - | linform F = Some F - | linform (NOT p) = lift_un NOT (linform p) - | linform (And (p, q)) = - lift_bin ((fn f => fn g => And (f, g)), (linform p, linform q)) - | linform (Or (p, q)) = - lift_bin ((fn f => fn g => Or (f, g)), (linform p, linform q)); - -fun nnf (Lt (it1, it2)) = Le (Sub (it1, it2), Cst (~ 1)) - | nnf (Gt (it1, it2)) = Le (Sub (it2, it1), Cst (~ 1)) - | nnf (Le (it1, it2)) = Le (it1, it2) - | nnf (Ge (it1, it2)) = Le (it2, it1) - | nnf (Eq (it1, it2)) = Eq (it2, it1) - | nnf (Divides (d, t)) = Divides (d, t) - | nnf T = T - | nnf F = F - | nnf (And (p, q)) = And (nnf p, nnf q) - | nnf (Or (p, q)) = Or (nnf p, nnf q) - | nnf (Imp (p, q)) = Or (nnf (NOT p), nnf q) - | nnf (Equ (p, q)) = Or (And (nnf p, nnf q), And (nnf (NOT p), nnf (NOT q))) - | nnf (NOT (Lt (it1, it2))) = Le (it2, it1) - | nnf (NOT (Gt (it1, it2))) = Le (it1, it2) - | nnf (NOT (Le (it1, it2))) = Le (Sub (it2, it1), Cst (~ 1)) - | nnf (NOT (Ge (it1, it2))) = Le (Sub (it1, it2), Cst (~ 1)) - | nnf (NOT (Eq (it1, it2))) = NOT (Eq (it1, it2)) - | nnf (NOT (Divides (d, t))) = NOT (Divides (d, t)) - | nnf (NOT T) = F - | nnf (NOT F) = T - | nnf (NOT (NOT p)) = nnf p - | nnf (NOT (And (p, q))) = Or (nnf (NOT p), nnf (NOT q)) - | nnf (NOT (Or (p, q))) = And (nnf (NOT p), nnf (NOT q)) - | nnf (NOT (Imp (p, q))) = And (nnf p, nnf (NOT q)) - | nnf (NOT (Equ (p, q))) = - Or (And (nnf p, nnf (NOT q)), And (nnf (NOT p), nnf q)); - -fun op_45_def2 z w = IntInf.+ (z,~ w); - -fun op_45_def0 m n = nat (op_45_def2 (m) (n)); - -val id_1_def0 : IntInf.int = (0 + 1); - -fun decrvarsI (Cst i) = Cst i - | decrvarsI (Var n) = Var (op_45_def0 n id_1_def0) - | decrvarsI (Neg a) = Neg (decrvarsI a) - | decrvarsI (Add (a, b)) = Add (decrvarsI a, decrvarsI b) - | decrvarsI (Sub (a, b)) = Sub (decrvarsI a, decrvarsI b) - | decrvarsI (Mult (a, b)) = Mult (decrvarsI a, decrvarsI b); - -fun decrvars (Lt (a, b)) = Lt (decrvarsI a, decrvarsI b) - | decrvars (Gt (a, b)) = Gt (decrvarsI a, decrvarsI b) - | decrvars (Le (a, b)) = Le (decrvarsI a, decrvarsI b) - | decrvars (Ge (a, b)) = Ge (decrvarsI a, decrvarsI b) - | decrvars (Eq (a, b)) = Eq (decrvarsI a, decrvarsI b) - | decrvars (Divides (a, b)) = Divides (decrvarsI a, decrvarsI b) - | decrvars T = T - | decrvars F = F - | decrvars (NOT p) = NOT (decrvars p) - | decrvars (And (p, q)) = And (decrvars p, decrvars q) - | decrvars (Or (p, q)) = Or (decrvars p, decrvars q) - | decrvars (Imp (p, q)) = Imp (decrvars p, decrvars q) - | decrvars (Equ (p, q)) = Equ (decrvars p, decrvars q); - -fun op_64 [] ys = ys - | op_64 (x :: xs) ys = (x :: op_64 xs ys); - -fun map f [] = [] - | map f (x :: xs) = (f x :: map f xs); - -fun iupto (i:IntInf.int, j:IntInf.int) = (if (j < i) then [] else (i :: iupto ((i + 1), j))); - -fun all_sums (j:IntInf.int, []) = [] - | all_sums (j, (i :: is)) = - op_64 (map (fn x => lin_add (i, Cst x)) (iupto (1, j))) (all_sums (j, is)); - -fun split x = (fn p => x (fst p) (snd p)); - -fun negateSnd x = split (fn q => fn r => (q, IntInf.~ r)) x; - -fun adjust b = - (fn (q:IntInf.int, r:IntInf.int) => - (if (0 <= op_45_def2 r b) then ((((2:IntInf.int) * q) + (1:IntInf.int)), op_45_def2 r b) - else (((2:IntInf.int) * q), r))); - -fun negDivAlg (a:IntInf.int, b:IntInf.int) = - (if ((0 <= (a + b)) orelse (b <= 0)) then (~1, (a + b)) - else adjust b (negDivAlg (a, (2 * b)))); - -fun posDivAlg (a:IntInf.int, b:IntInf.int) = - (if ((a < b) orelse (b <= 0)) then (0, a) - else adjust b (posDivAlg (a, (2 * b)))); - -fun divAlg x = - split (fn a:IntInf.int => fn b:IntInf.int => - (if (0 <= a) - then (if (0 <= b) then posDivAlg (a, b) - else (if (a = 0) then (0, 0) - else negateSnd (negDivAlg (~ a, ~ b)))) - else (if (0 < b) then negDivAlg (a, b) - else negateSnd (posDivAlg (~ a, ~ b))))) - x; - -fun op_mod_def1 a b = snd (divAlg (a, b)); - -fun op_dvd m n = (op_mod_def1 n m = 0); - -fun psimpl (Le (l, r)) = - (case lift_bin - ((fn x => fn y => lin_add (x, lin_neg y)), - (linearize l, linearize r)) of - None => Le (l, r) - | Some x => - (case x of Cst xa => (if (xa <= 0) then T else F) - | Var xa => Le (x, Cst 0) | Neg xa => Le (x, Cst 0) - | Add (xa, xb) => Le (x, Cst 0) | Sub (xa, xb) => Le (x, Cst 0) - | Mult (xa, xb) => Le (x, Cst 0))) - | psimpl (Eq (l, r)) = - (case lift_bin - ((fn x => fn y => lin_add (x, lin_neg y)), - (linearize l, linearize r)) of - None => Eq (l, r) - | Some x => - (case x of Cst xa => (if (xa = 0) then T else F) - | Var xa => Eq (x, Cst 0) | Neg xa => Eq (x, Cst 0) - | Add (xa, xb) => Eq (x, Cst 0) | Sub (xa, xb) => Eq (x, Cst 0) - | Mult (xa, xb) => Eq (x, Cst 0))) - | psimpl (Divides (Cst d, t)) = - (case linearize t of None => Divides (Cst d, t) - | Some x => - (case x of Cst xa => (if op_dvd d xa then T else F) - | Var xa => Divides (Cst d, x) | Neg xa => Divides (Cst d, x) - | Add (xa, xb) => Divides (Cst d, x) - | Sub (xa, xb) => Divides (Cst d, x) - | Mult (xa, xb) => Divides (Cst d, x))) - | psimpl (Equ (p, q)) = - let val p' = psimpl p; val q' = psimpl q - in (case p' of - Lt (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Gt (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Le (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Ge (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Eq (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Divides (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | T => q' - | F => (case q' of Lt (x, xa) => NOT q' | Gt (x, xa) => NOT q' - | Le (x, xa) => NOT q' | Ge (x, xa) => NOT q' - | Eq (x, xa) => NOT q' | Divides (x, xa) => NOT q' | T => F - | F => T | NOT x => x | And (x, xa) => NOT q' - | Or (x, xa) => NOT q' | Imp (x, xa) => NOT q' - | Equ (x, xa) => NOT q' | QAll x => NOT q' | QEx x => NOT q') - | NOT x => - (case q' of Lt (xa, xb) => Equ (p', q') - | Gt (xa, xb) => Equ (p', q') | Le (xa, xb) => Equ (p', q') - | Ge (xa, xb) => Equ (p', q') | Eq (xa, xb) => Equ (p', q') - | Divides (xa, xb) => Equ (p', q') | T => p' | F => x - | NOT xa => Equ (x, xa) | And (xa, xb) => Equ (p', q') - | Or (xa, xb) => Equ (p', q') | Imp (xa, xb) => Equ (p', q') - | Equ (xa, xb) => Equ (p', q') | QAll xa => Equ (p', q') - | QEx xa => Equ (p', q')) - | And (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Or (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Imp (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | Equ (x, xa) => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | QAll x => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q')) - | QEx x => - (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') - | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') - | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') - | T => p' | F => NOT p' | NOT x => Equ (p', q') - | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') - | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') - | QAll x => Equ (p', q') | QEx x => Equ (p', q'))) - end - | psimpl (NOT p) = - let val p' = psimpl p - in (case p' of Lt (x, xa) => NOT p' | Gt (x, xa) => NOT p' - | Le (x, xa) => NOT p' | Ge (x, xa) => NOT p' | Eq (x, xa) => NOT p' - | Divides (x, xa) => NOT p' | T => F | F => T | NOT x => x - | And (x, xa) => NOT p' | Or (x, xa) => NOT p' | Imp (x, xa) => NOT p' - | Equ (x, xa) => NOT p' | QAll x => NOT p' | QEx x => NOT p') - end - | psimpl (Lt (u, v)) = Lt (u, v) - | psimpl (Gt (w, x)) = Gt (w, x) - | psimpl (Ge (aa, ab)) = Ge (aa, ab) - | psimpl (Divides (Var bp, af)) = Divides (Var bp, af) - | psimpl (Divides (Neg bq, af)) = Divides (Neg bq, af) - | psimpl (Divides (Add (br, bs), af)) = Divides (Add (br, bs), af) - | psimpl (Divides (Sub (bt, bu), af)) = Divides (Sub (bt, bu), af) - | psimpl (Divides (Mult (bv, bw), af)) = Divides (Mult (bv, bw), af) - | psimpl T = T - | psimpl F = F - | psimpl (QAll ap) = QAll ap - | psimpl (QEx aq) = QEx aq - | psimpl (And (p, q)) = - let val p' = psimpl p - in (case p' of - Lt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Gt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Le (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Ge (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Eq (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Divides (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | T => psimpl q | F => F - | NOT x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | And (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Or (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Imp (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | Equ (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | QAll x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end - | QEx x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => And (p', q') - | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') - | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') - | Divides (x, xa) => And (p', q') | T => p' | F => F - | NOT x => And (p', q') | And (x, xa) => And (p', q') - | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') - | Equ (x, xa) => And (p', q') | QAll x => And (p', q') - | QEx x => And (p', q')) - end) - end - | psimpl (Or (p, q)) = - let val p' = psimpl p - in (case p' of - Lt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') | Gt (x, xa) => Or (p', q') - | Le (x, xa) => Or (p', q') | Ge (x, xa) => Or (p', q') - | Eq (x, xa) => Or (p', q') | Divides (x, xa) => Or (p', q') - | T => T | F => p' | NOT x => Or (p', q') - | And (x, xa) => Or (p', q') | Or (x, xa) => Or (p', q') - | Imp (x, xa) => Or (p', q') | Equ (x, xa) => Or (p', q') - | QAll x => Or (p', q') | QEx x => Or (p', q')) - end - | Gt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Le (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Ge (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Eq (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Divides (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | T => T | F => psimpl q - | NOT x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | And (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Or (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Imp (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | Equ (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | QAll x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end - | QEx x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Or (p', q') - | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') - | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') - | Divides (x, xa) => Or (p', q') | T => T | F => p' - | NOT x => Or (p', q') | And (x, xa) => Or (p', q') - | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') - | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') - | QEx x => Or (p', q')) - end) - end - | psimpl (Imp (p, q)) = - let val p' = psimpl p - in (case p' of - Lt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Gt (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Le (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Ge (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Eq (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Divides (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | T => psimpl q | F => T - | NOT x => - let val q' = psimpl q - in (case q' of Lt (xa, xb) => Or (x, q') - | Gt (xa, xb) => Or (x, q') | Le (xa, xb) => Or (x, q') - | Ge (xa, xb) => Or (x, q') | Eq (xa, xb) => Or (x, q') - | Divides (xa, xb) => Or (x, q') | T => T | F => x - | NOT xa => Or (x, q') | And (xa, xb) => Or (x, q') - | Or (xa, xb) => Or (x, q') | Imp (xa, xb) => Or (x, q') - | Equ (xa, xb) => Or (x, q') | QAll xa => Or (x, q') - | QEx xa => Or (x, q')) - end - | And (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Or (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Imp (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | Equ (x, xa) => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | QAll x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end - | QEx x => - let val q' = psimpl q - in (case q' of Lt (x, xa) => Imp (p', q') - | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') - | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') - | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' - | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') - | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') - | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') - | QEx x => Imp (p', q')) - end) - end; - -fun subst_it i (Cst b) = Cst b - | subst_it i (Var n) = (if (n = 0) then i else Var n) - | subst_it i (Neg it) = Neg (subst_it i it) - | subst_it i (Add (it1, it2)) = Add (subst_it i it1, subst_it i it2) - | subst_it i (Sub (it1, it2)) = Sub (subst_it i it1, subst_it i it2) - | subst_it i (Mult (it1, it2)) = Mult (subst_it i it1, subst_it i it2); - -fun subst_p i (Le (it1, it2)) = Le (subst_it i it1, subst_it i it2) - | subst_p i (Lt (it1, it2)) = Lt (subst_it i it1, subst_it i it2) - | subst_p i (Ge (it1, it2)) = Ge (subst_it i it1, subst_it i it2) - | subst_p i (Gt (it1, it2)) = Gt (subst_it i it1, subst_it i it2) - | subst_p i (Eq (it1, it2)) = Eq (subst_it i it1, subst_it i it2) - | subst_p i (Divides (d, t)) = Divides (subst_it i d, subst_it i t) - | subst_p i T = T - | subst_p i F = F - | subst_p i (And (p, q)) = And (subst_p i p, subst_p i q) - | subst_p i (Or (p, q)) = Or (subst_p i p, subst_p i q) - | subst_p i (Imp (p, q)) = Imp (subst_p i p, subst_p i q) - | subst_p i (Equ (p, q)) = Equ (subst_p i p, subst_p i q) - | subst_p i (NOT p) = NOT (subst_p i p); - -fun explode_disj ([], p) = F - | explode_disj ((i :: is), p) = - let val pi = psimpl (subst_p i p) - in (case pi of - Lt (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Gt (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Le (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Ge (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Eq (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Divides (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | T => T | F => explode_disj (is, p) - | NOT x => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | And (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Or (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Imp (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | Equ (x, xa) => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | QAll x => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end - | QEx x => - let val r = explode_disj (is, p) - in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) - | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) - | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) - | T => T | F => pi | NOT x => Or (pi, r) - | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) - | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) - | QAll x => Or (pi, r) | QEx x => Or (pi, r)) - end) - end; - -fun minusinf (And (p, q)) = And (minusinf p, minusinf q) - | minusinf (Or (p, q)) = Or (minusinf p, minusinf q) - | minusinf (Lt (u, v)) = Lt (u, v) - | minusinf (Gt (w, x)) = Gt (w, x) - | minusinf (Le (Cst bo, z)) = Le (Cst bo, z) - | minusinf (Le (Var bp, z)) = Le (Var bp, z) - | minusinf (Le (Neg bq, z)) = Le (Neg bq, z) - | minusinf (Le (Add (Cst cg, bs), z)) = Le (Add (Cst cg, bs), z) - | minusinf (Le (Add (Var ch, bs), z)) = Le (Add (Var ch, bs), z) - | minusinf (Le (Add (Neg ci, bs), z)) = Le (Add (Neg ci, bs), z) - | minusinf (Le (Add (Add (cj, ck), bs), z)) = Le (Add (Add (cj, ck), bs), z) - | minusinf (Le (Add (Sub (cl, cm), bs), z)) = Le (Add (Sub (cl, cm), bs), z) - | minusinf (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = - Le (Add (Mult (Cst cy, Cst dq), bs), z) - | minusinf (Le (Add (Mult (Cst cy, Var ei), bs), z)) = - (if (ei = 0) then (if (cy < 0) then F else T) - else Le (Add (Mult (Cst cy, Var (op_45_def0 ei id_1_def0 + 1)), bs), z)) - | minusinf (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = - Le (Add (Mult (Cst cy, Neg ds), bs), z) - | minusinf (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = - Le (Add (Mult (Cst cy, Add (dt, du)), bs), z) - | minusinf (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = - Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z) - | minusinf (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = - Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z) - | minusinf (Le (Add (Mult (Var cz, co), bs), z)) = - Le (Add (Mult (Var cz, co), bs), z) - | minusinf (Le (Add (Mult (Neg da, co), bs), z)) = - Le (Add (Mult (Neg da, co), bs), z) - | minusinf (Le (Add (Mult (Add (db, dc), co), bs), z)) = - Le (Add (Mult (Add (db, dc), co), bs), z) - | minusinf (Le (Add (Mult (Sub (dd, de), co), bs), z)) = - Le (Add (Mult (Sub (dd, de), co), bs), z) - | minusinf (Le (Add (Mult (Mult (df, dg), co), bs), z)) = - Le (Add (Mult (Mult (df, dg), co), bs), z) - | minusinf (Le (Sub (bt, bu), z)) = Le (Sub (bt, bu), z) - | minusinf (Le (Mult (bv, bw), z)) = Le (Mult (bv, bw), z) - | minusinf (Ge (aa, ab)) = Ge (aa, ab) - | minusinf (Eq (Cst ek, ad)) = Eq (Cst ek, ad) - | minusinf (Eq (Var el, ad)) = Eq (Var el, ad) - | minusinf (Eq (Neg em, ad)) = Eq (Neg em, ad) - | minusinf (Eq (Add (Cst fc, eo), ad)) = Eq (Add (Cst fc, eo), ad) - | minusinf (Eq (Add (Var fd, eo), ad)) = Eq (Add (Var fd, eo), ad) - | minusinf (Eq (Add (Neg fe, eo), ad)) = Eq (Add (Neg fe, eo), ad) - | minusinf (Eq (Add (Add (ff, fg), eo), ad)) = Eq (Add (Add (ff, fg), eo), ad) - | minusinf (Eq (Add (Sub (fh, fi), eo), ad)) = Eq (Add (Sub (fh, fi), eo), ad) - | minusinf (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = - Eq (Add (Mult (Cst fu, Cst gm), eo), ad) - | minusinf (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = - (if (he = 0) then F - else Eq (Add (Mult (Cst fu, Var (op_45_def0 he id_1_def0 + 1)), eo), ad)) - | minusinf (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = - Eq (Add (Mult (Cst fu, Neg go), eo), ad) - | minusinf (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = - Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad) - | minusinf (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = - Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad) - | minusinf (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = - Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad) - | minusinf (Eq (Add (Mult (Var fv, fk), eo), ad)) = - Eq (Add (Mult (Var fv, fk), eo), ad) - | minusinf (Eq (Add (Mult (Neg fw, fk), eo), ad)) = - Eq (Add (Mult (Neg fw, fk), eo), ad) - | minusinf (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = - Eq (Add (Mult (Add (fx, fy), fk), eo), ad) - | minusinf (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = - Eq (Add (Mult (Sub (fz, ga), fk), eo), ad) - | minusinf (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = - Eq (Add (Mult (Mult (gb, gc), fk), eo), ad) - | minusinf (Eq (Sub (ep, eq), ad)) = Eq (Sub (ep, eq), ad) - | minusinf (Eq (Mult (er, es), ad)) = Eq (Mult (er, es), ad) - | minusinf (Divides (ae, af)) = Divides (ae, af) - | minusinf T = T - | minusinf F = F - | minusinf (NOT (Lt (hg, hh))) = NOT (Lt (hg, hh)) - | minusinf (NOT (Gt (hi, hj))) = NOT (Gt (hi, hj)) - | minusinf (NOT (Le (hk, hl))) = NOT (Le (hk, hl)) - | minusinf (NOT (Ge (hm, hn))) = NOT (Ge (hm, hn)) - | minusinf (NOT (Eq (Cst ja, hp))) = NOT (Eq (Cst ja, hp)) - | minusinf (NOT (Eq (Var jb, hp))) = NOT (Eq (Var jb, hp)) - | minusinf (NOT (Eq (Neg jc, hp))) = NOT (Eq (Neg jc, hp)) - | minusinf (NOT (Eq (Add (Cst js, je), hp))) = NOT (Eq (Add (Cst js, je), hp)) - | minusinf (NOT (Eq (Add (Var jt, je), hp))) = NOT (Eq (Add (Var jt, je), hp)) - | minusinf (NOT (Eq (Add (Neg ju, je), hp))) = NOT (Eq (Add (Neg ju, je), hp)) - | minusinf (NOT (Eq (Add (Add (jv, jw), je), hp))) = - NOT (Eq (Add (Add (jv, jw), je), hp)) - | minusinf (NOT (Eq (Add (Sub (jx, jy), je), hp))) = - NOT (Eq (Add (Sub (jx, jy), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = - NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = - (if (lu = 0) then T - else NOT (Eq (Add (Mult (Cst kk, Var (op_45_def0 lu id_1_def0 + 1)), je), - hp))) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = - NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = - NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = - NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = - NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = - NOT (Eq (Add (Mult (Var kl, ka), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = - NOT (Eq (Add (Mult (Neg km, ka), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = - NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = - NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp)) - | minusinf (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = - NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp)) - | minusinf (NOT (Eq (Sub (jf, jg), hp))) = NOT (Eq (Sub (jf, jg), hp)) - | minusinf (NOT (Eq (Mult (jh, ji), hp))) = NOT (Eq (Mult (jh, ji), hp)) - | minusinf (NOT (Divides (hq, hr))) = NOT (Divides (hq, hr)) - | minusinf (NOT T) = NOT T - | minusinf (NOT F) = NOT F - | minusinf (NOT (NOT hs)) = NOT (NOT hs) - | minusinf (NOT (And (ht, hu))) = NOT (And (ht, hu)) - | minusinf (NOT (Or (hv, hw))) = NOT (Or (hv, hw)) - | minusinf (NOT (Imp (hx, hy))) = NOT (Imp (hx, hy)) - | minusinf (NOT (Equ (hz, ia))) = NOT (Equ (hz, ia)) - | minusinf (NOT (QAll ib)) = NOT (QAll ib) - | minusinf (NOT (QEx ic)) = NOT (QEx ic) - | minusinf (Imp (al, am)) = Imp (al, am) - | minusinf (Equ (an, ao)) = Equ (an, ao) - | minusinf (QAll ap) = QAll ap - | minusinf (QEx aq) = QEx aq; - -fun abs (i:IntInf.int) = (if (i < 0) then IntInf.~ i else i); - -fun op_div_def1 a b = fst (divAlg (a, b)); - -fun op_mod_def0 m n = nat (op_mod_def1 (m) (n)); - -fun ngcd (m:IntInf.int, n:IntInf.int) = (if (n = 0) then m else ngcd (n, op_mod_def0 m n)); - -fun igcd x = split (fn a => fn b => (ngcd (nat (abs a), nat (abs b)))) x; - -fun ilcm (a:IntInf.int) (b:IntInf.int) = op_div_def1 (a * b) (igcd (a, b)); - -fun divlcm (NOT p) = divlcm p - | divlcm (And (p, q)) = ilcm (divlcm p) (divlcm q) - | divlcm (Or (p, q)) = ilcm (divlcm p) (divlcm q) - | divlcm (Lt (u, v)) = 1 - | divlcm (Gt (w, x)) = 1 - | divlcm (Le (y, z)) = 1 - | divlcm (Ge (aa, ab)) = 1 - | divlcm (Eq (ac, ad)) = 1 - | divlcm (Divides (Cst bo, Cst cg)) = 1 - | divlcm (Divides (Cst bo, Var ch)) = 1 - | divlcm (Divides (Cst bo, Neg ci)) = 1 - | divlcm (Divides (Cst bo, Add (Cst cy, ck))) = 1 - | divlcm (Divides (Cst bo, Add (Var cz, ck))) = 1 - | divlcm (Divides (Cst bo, Add (Neg da, ck))) = 1 - | divlcm (Divides (Cst bo, Add (Add (db, dc), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Sub (dd, de), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Cst ei), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Var fa), ck))) = - (if (fa = 0) then abs bo else 1) - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Neg ek), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Add (el, em)), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Sub (en, eo)), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Mult (ep, eq)), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Var dr, dg), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Neg ds, dg), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Add (dt, du), dg), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Sub (dv, dw), dg), ck))) = 1 - | divlcm (Divides (Cst bo, Add (Mult (Mult (dx, dy), dg), ck))) = 1 - | divlcm (Divides (Cst bo, Sub (cl, cm))) = 1 - | divlcm (Divides (Cst bo, Mult (cn, co))) = 1 - | divlcm (Divides (Var bp, af)) = 1 - | divlcm (Divides (Neg bq, af)) = 1 - | divlcm (Divides (Add (br, bs), af)) = 1 - | divlcm (Divides (Sub (bt, bu), af)) = 1 - | divlcm (Divides (Mult (bv, bw), af)) = 1 - | divlcm T = 1 - | divlcm F = 1 - | divlcm (Imp (al, am)) = 1 - | divlcm (Equ (an, ao)) = 1 - | divlcm (QAll ap) = 1 - | divlcm (QEx aq) = 1; - -fun explode_minf (q, B) = - let val d = divlcm q; val pm = minusinf q; - val dj1 = explode_disj (map (fn x => Cst x) (iupto (1, d)), pm) - in (case dj1 of - Lt (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Gt (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Le (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Ge (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Eq (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Divides (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | T => T | F => explode_disj (all_sums (d, B), q) - | NOT x => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | And (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Or (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Imp (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | Equ (x, xa) => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | QAll x => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end - | QEx x => - let val dj2 = explode_disj (all_sums (d, B), q) - in (case dj2 of Lt (x, xa) => Or (dj1, dj2) - | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) - | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) - | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 - | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) - | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) - | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) - | QEx x => Or (dj1, dj2)) - end) - end; - -fun mirror (And (p, q)) = And (mirror p, mirror q) - | mirror (Or (p, q)) = Or (mirror p, mirror q) - | mirror (Lt (u, v)) = Lt (u, v) - | mirror (Gt (w, x)) = Gt (w, x) - | mirror (Le (Cst bp, aa)) = Le (Cst bp, aa) - | mirror (Le (Var bq, aa)) = Le (Var bq, aa) - | mirror (Le (Neg br, aa)) = Le (Neg br, aa) - | mirror (Le (Add (Cst ch, bt), aa)) = Le (Add (Cst ch, bt), aa) - | mirror (Le (Add (Var ci, bt), aa)) = Le (Add (Var ci, bt), aa) - | mirror (Le (Add (Neg cj, bt), aa)) = Le (Add (Neg cj, bt), aa) - | mirror (Le (Add (Add (ck, cl), bt), aa)) = Le (Add (Add (ck, cl), bt), aa) - | mirror (Le (Add (Sub (cm, cn), bt), aa)) = Le (Add (Sub (cm, cn), bt), aa) - | mirror (Le (Add (Mult (Cst cz, Cst dr), bt), aa)) = - Le (Add (Mult (Cst cz, Cst dr), bt), aa) - | mirror (Le (Add (Mult (Cst cz, Var ej), bt), aa)) = - (if (ej = 0) then Le (Add (Mult (Cst (~ cz), Var 0), bt), aa) - else Le (Add (Mult (Cst cz, Var (op_45_def0 ej id_1_def0 + 1)), bt), aa)) - | mirror (Le (Add (Mult (Cst cz, Neg dt), bt), aa)) = - Le (Add (Mult (Cst cz, Neg dt), bt), aa) - | mirror (Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa)) = - Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa) - | mirror (Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa)) = - Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa) - | mirror (Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa)) = - Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa) - | mirror (Le (Add (Mult (Var da, cp), bt), aa)) = - Le (Add (Mult (Var da, cp), bt), aa) - | mirror (Le (Add (Mult (Neg db, cp), bt), aa)) = - Le (Add (Mult (Neg db, cp), bt), aa) - | mirror (Le (Add (Mult (Add (dc, dd), cp), bt), aa)) = - Le (Add (Mult (Add (dc, dd), cp), bt), aa) - | mirror (Le (Add (Mult (Sub (de, df), cp), bt), aa)) = - Le (Add (Mult (Sub (de, df), cp), bt), aa) - | mirror (Le (Add (Mult (Mult (dg, dh), cp), bt), aa)) = - Le (Add (Mult (Mult (dg, dh), cp), bt), aa) - | mirror (Le (Sub (bu, bv), aa)) = Le (Sub (bu, bv), aa) - | mirror (Le (Mult (bw, bx), aa)) = Le (Mult (bw, bx), aa) - | mirror (Ge (ab, ac)) = Ge (ab, ac) - | mirror (Eq (Cst el, ae)) = Eq (Cst el, ae) - | mirror (Eq (Var em, ae)) = Eq (Var em, ae) - | mirror (Eq (Neg en, ae)) = Eq (Neg en, ae) - | mirror (Eq (Add (Cst fd, ep), ae)) = Eq (Add (Cst fd, ep), ae) - | mirror (Eq (Add (Var fe, ep), ae)) = Eq (Add (Var fe, ep), ae) - | mirror (Eq (Add (Neg ff, ep), ae)) = Eq (Add (Neg ff, ep), ae) - | mirror (Eq (Add (Add (fg, fh), ep), ae)) = Eq (Add (Add (fg, fh), ep), ae) - | mirror (Eq (Add (Sub (fi, fj), ep), ae)) = Eq (Add (Sub (fi, fj), ep), ae) - | mirror (Eq (Add (Mult (Cst fv, Cst gn), ep), ae)) = - Eq (Add (Mult (Cst fv, Cst gn), ep), ae) - | mirror (Eq (Add (Mult (Cst fv, Var hf), ep), ae)) = - (if (hf = 0) then Eq (Add (Mult (Cst (~ fv), Var 0), ep), ae) - else Eq (Add (Mult (Cst fv, Var (op_45_def0 hf id_1_def0 + 1)), ep), ae)) - | mirror (Eq (Add (Mult (Cst fv, Neg gp), ep), ae)) = - Eq (Add (Mult (Cst fv, Neg gp), ep), ae) - | mirror (Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae)) = - Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae) - | mirror (Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae)) = - Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae) - | mirror (Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae)) = - Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae) - | mirror (Eq (Add (Mult (Var fw, fl), ep), ae)) = - Eq (Add (Mult (Var fw, fl), ep), ae) - | mirror (Eq (Add (Mult (Neg fx, fl), ep), ae)) = - Eq (Add (Mult (Neg fx, fl), ep), ae) - | mirror (Eq (Add (Mult (Add (fy, fz), fl), ep), ae)) = - Eq (Add (Mult (Add (fy, fz), fl), ep), ae) - | mirror (Eq (Add (Mult (Sub (ga, gb), fl), ep), ae)) = - Eq (Add (Mult (Sub (ga, gb), fl), ep), ae) - | mirror (Eq (Add (Mult (Mult (gc, gd), fl), ep), ae)) = - Eq (Add (Mult (Mult (gc, gd), fl), ep), ae) - | mirror (Eq (Sub (eq, er), ae)) = Eq (Sub (eq, er), ae) - | mirror (Eq (Mult (es, et), ae)) = Eq (Mult (es, et), ae) - | mirror (Divides (Cst hh, Cst hz)) = Divides (Cst hh, Cst hz) - | mirror (Divides (Cst hh, Var ia)) = Divides (Cst hh, Var ia) - | mirror (Divides (Cst hh, Neg ib)) = Divides (Cst hh, Neg ib) - | mirror (Divides (Cst hh, Add (Cst ir, id))) = - Divides (Cst hh, Add (Cst ir, id)) - | mirror (Divides (Cst hh, Add (Var is, id))) = - Divides (Cst hh, Add (Var is, id)) - | mirror (Divides (Cst hh, Add (Neg it, id))) = - Divides (Cst hh, Add (Neg it, id)) - | mirror (Divides (Cst hh, Add (Add (iu, iv), id))) = - Divides (Cst hh, Add (Add (iu, iv), id)) - | mirror (Divides (Cst hh, Add (Sub (iw, ix), id))) = - Divides (Cst hh, Add (Sub (iw, ix), id)) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id))) = - Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id)) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Var kt), id))) = - (if (kt = 0) then Divides (Cst hh, Add (Mult (Cst (~ jj), Var 0), id)) - else Divides - (Cst hh, - Add (Mult (Cst jj, Var (op_45_def0 kt id_1_def0 + 1)), id))) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id))) = - Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id)) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id))) = - Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id)) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id))) = - Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id)) - | mirror (Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id))) = - Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id)) - | mirror (Divides (Cst hh, Add (Mult (Var jk, iz), id))) = - Divides (Cst hh, Add (Mult (Var jk, iz), id)) - | mirror (Divides (Cst hh, Add (Mult (Neg jl, iz), id))) = - Divides (Cst hh, Add (Mult (Neg jl, iz), id)) - | mirror (Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id))) = - Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id)) - | mirror (Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id))) = - Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id)) - | mirror (Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id))) = - Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id)) - | mirror (Divides (Cst hh, Sub (ie, if'))) = Divides (Cst hh, Sub (ie, if')) - | mirror (Divides (Cst hh, Mult (ig, ih))) = Divides (Cst hh, Mult (ig, ih)) - | mirror (Divides (Var hi, ag)) = Divides (Var hi, ag) - | mirror (Divides (Neg hj, ag)) = Divides (Neg hj, ag) - | mirror (Divides (Add (hk, hl), ag)) = Divides (Add (hk, hl), ag) - | mirror (Divides (Sub (hm, hn), ag)) = Divides (Sub (hm, hn), ag) - | mirror (Divides (Mult (ho, hp), ag)) = Divides (Mult (ho, hp), ag) - | mirror T = T - | mirror F = F - | mirror (NOT (Lt (kv, kw))) = NOT (Lt (kv, kw)) - | mirror (NOT (Gt (kx, ky))) = NOT (Gt (kx, ky)) - | mirror (NOT (Le (kz, la))) = NOT (Le (kz, la)) - | mirror (NOT (Ge (lb, lc))) = NOT (Ge (lb, lc)) - | mirror (NOT (Eq (Cst mp, le))) = NOT (Eq (Cst mp, le)) - | mirror (NOT (Eq (Var mq, le))) = NOT (Eq (Var mq, le)) - | mirror (NOT (Eq (Neg mr, le))) = NOT (Eq (Neg mr, le)) - | mirror (NOT (Eq (Add (Cst nh, mt), le))) = NOT (Eq (Add (Cst nh, mt), le)) - | mirror (NOT (Eq (Add (Var ni, mt), le))) = NOT (Eq (Add (Var ni, mt), le)) - | mirror (NOT (Eq (Add (Neg nj, mt), le))) = NOT (Eq (Add (Neg nj, mt), le)) - | mirror (NOT (Eq (Add (Add (nk, nl), mt), le))) = - NOT (Eq (Add (Add (nk, nl), mt), le)) - | mirror (NOT (Eq (Add (Sub (nm, nn), mt), le))) = - NOT (Eq (Add (Sub (nm, nn), mt), le)) - | mirror (NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le))) = - NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le)) - | mirror (NOT (Eq (Add (Mult (Cst nz, Var pj), mt), le))) = - (if (pj = 0) then NOT (Eq (Add (Mult (Cst (~ nz), Var 0), mt), le)) - else NOT (Eq (Add (Mult (Cst nz, Var (op_45_def0 pj id_1_def0 + 1)), mt), - le))) - | mirror (NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le))) = - NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le)) - | mirror (NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le))) = - NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le)) - | mirror (NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le))) = - NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le)) - | mirror (NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le))) = - NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le)) - | mirror (NOT (Eq (Add (Mult (Var oa, np), mt), le))) = - NOT (Eq (Add (Mult (Var oa, np), mt), le)) - | mirror (NOT (Eq (Add (Mult (Neg ob, np), mt), le))) = - NOT (Eq (Add (Mult (Neg ob, np), mt), le)) - | mirror (NOT (Eq (Add (Mult (Add (oc, od), np), mt), le))) = - NOT (Eq (Add (Mult (Add (oc, od), np), mt), le)) - | mirror (NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le))) = - NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le)) - | mirror (NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le))) = - NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le)) - | mirror (NOT (Eq (Sub (mu, mv), le))) = NOT (Eq (Sub (mu, mv), le)) - | mirror (NOT (Eq (Mult (mw, mx), le))) = NOT (Eq (Mult (mw, mx), le)) - | mirror (NOT (Divides (Cst pl, Cst qd))) = NOT (Divides (Cst pl, Cst qd)) - | mirror (NOT (Divides (Cst pl, Var qe))) = NOT (Divides (Cst pl, Var qe)) - | mirror (NOT (Divides (Cst pl, Neg qf))) = NOT (Divides (Cst pl, Neg qf)) - | mirror (NOT (Divides (Cst pl, Add (Cst qv, qh)))) = - NOT (Divides (Cst pl, Add (Cst qv, qh))) - | mirror (NOT (Divides (Cst pl, Add (Var qw, qh)))) = - NOT (Divides (Cst pl, Add (Var qw, qh))) - | mirror (NOT (Divides (Cst pl, Add (Neg qx, qh)))) = - NOT (Divides (Cst pl, Add (Neg qx, qh))) - | mirror (NOT (Divides (Cst pl, Add (Add (qy, qz), qh)))) = - NOT (Divides (Cst pl, Add (Add (qy, qz), qh))) - | mirror (NOT (Divides (Cst pl, Add (Sub (ra, rb), qh)))) = - NOT (Divides (Cst pl, Add (Sub (ra, rb), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Var sx), qh)))) = - (if (sx = 0) - then NOT (Divides (Cst pl, Add (Mult (Cst (~ rn), Var 0), qh))) - else NOT (Divides - (Cst pl, - Add (Mult (Cst rn, Var (op_45_def0 sx id_1_def0 + 1)), - qh)))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh))) - | mirror (NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh)))) = - NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh))) - | mirror (NOT (Divides (Cst pl, Sub (qi, qj)))) = - NOT (Divides (Cst pl, Sub (qi, qj))) - | mirror (NOT (Divides (Cst pl, Mult (qk, ql)))) = - NOT (Divides (Cst pl, Mult (qk, ql))) - | mirror (NOT (Divides (Var pm, lg))) = NOT (Divides (Var pm, lg)) - | mirror (NOT (Divides (Neg pn, lg))) = NOT (Divides (Neg pn, lg)) - | mirror (NOT (Divides (Add (po, pp), lg))) = NOT (Divides (Add (po, pp), lg)) - | mirror (NOT (Divides (Sub (pq, pr), lg))) = NOT (Divides (Sub (pq, pr), lg)) - | mirror (NOT (Divides (Mult (ps, pt), lg))) = - NOT (Divides (Mult (ps, pt), lg)) - | mirror (NOT T) = NOT T - | mirror (NOT F) = NOT F - | mirror (NOT (NOT lh)) = NOT (NOT lh) - | mirror (NOT (And (li, lj))) = NOT (And (li, lj)) - | mirror (NOT (Or (lk, ll))) = NOT (Or (lk, ll)) - | mirror (NOT (Imp (lm, ln))) = NOT (Imp (lm, ln)) - | mirror (NOT (Equ (lo, lp))) = NOT (Equ (lo, lp)) - | mirror (NOT (QAll lq)) = NOT (QAll lq) - | mirror (NOT (QEx lr)) = NOT (QEx lr) - | mirror (Imp (am, an)) = Imp (am, an) - | mirror (Equ (ao, ap)) = Equ (ao, ap) - | mirror (QAll aq) = QAll aq - | mirror (QEx ar) = QEx ar; - -fun op_43_def0 m n = nat ((m) + (n)); - -fun size_def1 [] = (0:IntInf.int) - | size_def1 (a :: list) = op_43_def0 (size_def1 list) (0 + 1); - -fun aset (And (p, q)) = op_64 (aset p) (aset q) - | aset (Or (p, q)) = op_64 (aset p) (aset q) - | aset (Lt (u, v)) = [] - | aset (Gt (w, x)) = [] - | aset (Le (Cst bo, z)) = [] - | aset (Le (Var bp, z)) = [] - | aset (Le (Neg bq, z)) = [] - | aset (Le (Add (Cst cg, bs), z)) = [] - | aset (Le (Add (Var ch, bs), z)) = [] - | aset (Le (Add (Neg ci, bs), z)) = [] - | aset (Le (Add (Add (cj, ck), bs), z)) = [] - | aset (Le (Add (Sub (cl, cm), bs), z)) = [] - | aset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] - | aset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = - (if (ei = 0) - then (if (cy < 0) then [lin_add (bs, Cst 1)] - else [lin_neg bs, lin_add (lin_neg bs, Cst 1)]) - else []) - | aset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] - | aset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] - | aset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] - | aset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] - | aset (Le (Add (Mult (Var cz, co), bs), z)) = [] - | aset (Le (Add (Mult (Neg da, co), bs), z)) = [] - | aset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] - | aset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] - | aset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] - | aset (Le (Sub (bt, bu), z)) = [] - | aset (Le (Mult (bv, bw), z)) = [] - | aset (Ge (aa, ab)) = [] - | aset (Eq (Cst ek, ad)) = [] - | aset (Eq (Var el, ad)) = [] - | aset (Eq (Neg em, ad)) = [] - | aset (Eq (Add (Cst fc, eo), ad)) = [] - | aset (Eq (Add (Var fd, eo), ad)) = [] - | aset (Eq (Add (Neg fe, eo), ad)) = [] - | aset (Eq (Add (Add (ff, fg), eo), ad)) = [] - | aset (Eq (Add (Sub (fh, fi), eo), ad)) = [] - | aset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] - | aset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = - (if (he = 0) - then (if (fu < 0) then [lin_add (eo, Cst 1)] - else [lin_add (lin_neg eo, Cst 1)]) - else []) - | aset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] - | aset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] - | aset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] - | aset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] - | aset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] - | aset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] - | aset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] - | aset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] - | aset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] - | aset (Eq (Sub (ep, eq), ad)) = [] - | aset (Eq (Mult (er, es), ad)) = [] - | aset (Divides (ae, af)) = [] - | aset T = [] - | aset F = [] - | aset (NOT (Lt (hg, hh))) = [] - | aset (NOT (Gt (hi, hj))) = [] - | aset (NOT (Le (hk, hl))) = [] - | aset (NOT (Ge (hm, hn))) = [] - | aset (NOT (Eq (Cst ja, hp))) = [] - | aset (NOT (Eq (Var jb, hp))) = [] - | aset (NOT (Eq (Neg jc, hp))) = [] - | aset (NOT (Eq (Add (Cst js, je), hp))) = [] - | aset (NOT (Eq (Add (Var jt, je), hp))) = [] - | aset (NOT (Eq (Add (Neg ju, je), hp))) = [] - | aset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] - | aset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = - (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) - | aset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] - | aset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] - | aset (NOT (Eq (Sub (jf, jg), hp))) = [] - | aset (NOT (Eq (Mult (jh, ji), hp))) = [] - | aset (NOT (Divides (hq, hr))) = [] - | aset (NOT T) = [] - | aset (NOT F) = [] - | aset (NOT (NOT hs)) = [] - | aset (NOT (And (ht, hu))) = [] - | aset (NOT (Or (hv, hw))) = [] - | aset (NOT (Imp (hx, hy))) = [] - | aset (NOT (Equ (hz, ia))) = [] - | aset (NOT (QAll ib)) = [] - | aset (NOT (QEx ic)) = [] - | aset (Imp (al, am)) = [] - | aset (Equ (an, ao)) = [] - | aset (QAll ap) = [] - | aset (QEx aq) = []; - -fun op_mem x [] = false - | op_mem x (y :: ys) = (if (y = x) then true else op_mem x ys); - -fun list_insert x xs = (if op_mem x xs then xs else (x :: xs)); - -fun list_set [] = [] - | list_set (x :: xs) = list_insert x (list_set xs); - -fun bset (And (p, q)) = op_64 (bset p) (bset q) - | bset (Or (p, q)) = op_64 (bset p) (bset q) - | bset (Lt (u, v)) = [] - | bset (Gt (w, x)) = [] - | bset (Le (Cst bo, z)) = [] - | bset (Le (Var bp, z)) = [] - | bset (Le (Neg bq, z)) = [] - | bset (Le (Add (Cst cg, bs), z)) = [] - | bset (Le (Add (Var ch, bs), z)) = [] - | bset (Le (Add (Neg ci, bs), z)) = [] - | bset (Le (Add (Add (cj, ck), bs), z)) = [] - | bset (Le (Add (Sub (cl, cm), bs), z)) = [] - | bset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] - | bset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = - (if (ei = 0) - then (if (cy < 0) then [lin_add (bs, Cst ~1), bs] - else [lin_add (lin_neg bs, Cst ~1)]) - else []) - | bset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] - | bset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] - | bset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] - | bset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] - | bset (Le (Add (Mult (Var cz, co), bs), z)) = [] - | bset (Le (Add (Mult (Neg da, co), bs), z)) = [] - | bset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] - | bset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] - | bset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] - | bset (Le (Sub (bt, bu), z)) = [] - | bset (Le (Mult (bv, bw), z)) = [] - | bset (Ge (aa, ab)) = [] - | bset (Eq (Cst ek, ad)) = [] - | bset (Eq (Var el, ad)) = [] - | bset (Eq (Neg em, ad)) = [] - | bset (Eq (Add (Cst fc, eo), ad)) = [] - | bset (Eq (Add (Var fd, eo), ad)) = [] - | bset (Eq (Add (Neg fe, eo), ad)) = [] - | bset (Eq (Add (Add (ff, fg), eo), ad)) = [] - | bset (Eq (Add (Sub (fh, fi), eo), ad)) = [] - | bset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] - | bset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = - (if (he = 0) - then (if (fu < 0) then [lin_add (eo, Cst ~1)] - else [lin_add (lin_neg eo, Cst ~1)]) - else []) - | bset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] - | bset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] - | bset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] - | bset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] - | bset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] - | bset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] - | bset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] - | bset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] - | bset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] - | bset (Eq (Sub (ep, eq), ad)) = [] - | bset (Eq (Mult (er, es), ad)) = [] - | bset (Divides (ae, af)) = [] - | bset T = [] - | bset F = [] - | bset (NOT (Lt (hg, hh))) = [] - | bset (NOT (Gt (hi, hj))) = [] - | bset (NOT (Le (hk, hl))) = [] - | bset (NOT (Ge (hm, hn))) = [] - | bset (NOT (Eq (Cst ja, hp))) = [] - | bset (NOT (Eq (Var jb, hp))) = [] - | bset (NOT (Eq (Neg jc, hp))) = [] - | bset (NOT (Eq (Add (Cst js, je), hp))) = [] - | bset (NOT (Eq (Add (Var jt, je), hp))) = [] - | bset (NOT (Eq (Add (Neg ju, je), hp))) = [] - | bset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] - | bset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = - (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) - | bset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] - | bset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] - | bset (NOT (Eq (Sub (jf, jg), hp))) = [] - | bset (NOT (Eq (Mult (jh, ji), hp))) = [] - | bset (NOT (Divides (hq, hr))) = [] - | bset (NOT T) = [] - | bset (NOT F) = [] - | bset (NOT (NOT hs)) = [] - | bset (NOT (And (ht, hu))) = [] - | bset (NOT (Or (hv, hw))) = [] - | bset (NOT (Imp (hx, hy))) = [] - | bset (NOT (Equ (hz, ia))) = [] - | bset (NOT (QAll ib)) = [] - | bset (NOT (QEx ic)) = [] - | bset (Imp (al, am)) = [] - | bset (Equ (an, ao)) = [] - | bset (QAll ap) = [] - | bset (QEx aq) = []; - -fun adjustcoeff (l:IntInf.int, Le (Add (Mult (Cst c, Var 0), r), Cst i)) = - (if (c <= 0) - then Le (Add (Mult (Cst ~1, Var 0), lin_mul (~ (op_div_def1 l c), r)), - Cst 0) - else Le (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) - | adjustcoeff (l, Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = - Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst c, Var 0), r), Cst i))) = - NOT (Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) - | adjustcoeff (l, And (p, q)) = And (adjustcoeff (l, p), adjustcoeff (l, q)) - | adjustcoeff (l, Or (p, q)) = Or (adjustcoeff (l, p), adjustcoeff (l, q)) - | adjustcoeff (l, Lt (w, x)) = Lt (w, x) - | adjustcoeff (l, Gt (y, z)) = Gt (y, z) - | adjustcoeff (l, Le (Cst bq, ab)) = Le (Cst bq, ab) - | adjustcoeff (l, Le (Var br, ab)) = Le (Var br, ab) - | adjustcoeff (l, Le (Neg bs, ab)) = Le (Neg bs, ab) - | adjustcoeff (l, Le (Add (Cst ci, bu), ab)) = Le (Add (Cst ci, bu), ab) - | adjustcoeff (l, Le (Add (Var cj, bu), ab)) = Le (Add (Var cj, bu), ab) - | adjustcoeff (l, Le (Add (Neg ck, bu), ab)) = Le (Add (Neg ck, bu), ab) - | adjustcoeff (l, Le (Add (Add (cl, cm), bu), ab)) = - Le (Add (Add (cl, cm), bu), ab) - | adjustcoeff (l, Le (Add (Sub (cn, co), bu), ab)) = - Le (Add (Sub (cn, co), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Cst ds), bu), ab)) = - Le (Add (Mult (Cst da, Cst ds), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Var en)) = - Le (Add (Mult (Cst da, Var 0), bu), Var en) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Neg eo)) = - Le (Add (Mult (Cst da, Var 0), bu), Neg eo) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq))) = - Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq)) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es))) = - Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es)) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu))) = - Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu)) - | adjustcoeff (l, Le (Add (Mult (Cst da, Var ek), bu), ab)) = - Le (Add (Mult (Cst da, Var ek), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Neg du), bu), ab)) = - Le (Add (Mult (Cst da, Neg du), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab)) = - Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab)) = - Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab)) = - Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Var db, cq), bu), ab)) = - Le (Add (Mult (Var db, cq), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Neg dc, cq), bu), ab)) = - Le (Add (Mult (Neg dc, cq), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Add (dd, de), cq), bu), ab)) = - Le (Add (Mult (Add (dd, de), cq), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Sub (df, dg), cq), bu), ab)) = - Le (Add (Mult (Sub (df, dg), cq), bu), ab) - | adjustcoeff (l, Le (Add (Mult (Mult (dh, di), cq), bu), ab)) = - Le (Add (Mult (Mult (dh, di), cq), bu), ab) - | adjustcoeff (l, Le (Sub (bv, bw), ab)) = Le (Sub (bv, bw), ab) - | adjustcoeff (l, Le (Mult (bx, by), ab)) = Le (Mult (bx, by), ab) - | adjustcoeff (l, Ge (ac, ad)) = Ge (ac, ad) - | adjustcoeff (l, Eq (Cst fe, af)) = Eq (Cst fe, af) - | adjustcoeff (l, Eq (Var ff, af)) = Eq (Var ff, af) - | adjustcoeff (l, Eq (Neg fg, af)) = Eq (Neg fg, af) - | adjustcoeff (l, Eq (Add (Cst fw, fi), af)) = Eq (Add (Cst fw, fi), af) - | adjustcoeff (l, Eq (Add (Var fx, fi), af)) = Eq (Add (Var fx, fi), af) - | adjustcoeff (l, Eq (Add (Neg fy, fi), af)) = Eq (Add (Neg fy, fi), af) - | adjustcoeff (l, Eq (Add (Add (fz, ga), fi), af)) = - Eq (Add (Add (fz, ga), fi), af) - | adjustcoeff (l, Eq (Add (Sub (gb, gc), fi), af)) = - Eq (Add (Sub (gb, gc), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Cst hg), fi), af)) = - Eq (Add (Mult (Cst go, Cst hg), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Var ib)) = - Eq (Add (Mult (Cst go, Var 0), fi), Var ib) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Neg ic)) = - Eq (Add (Mult (Cst go, Var 0), fi), Neg ic) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie))) = - Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie)) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig))) = - Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig)) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii))) = - Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii)) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Var hy), fi), af)) = - Eq (Add (Mult (Cst go, Var hy), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Neg hi), fi), af)) = - Eq (Add (Mult (Cst go, Neg hi), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af)) = - Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af)) = - Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af)) = - Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Var gp, ge), fi), af)) = - Eq (Add (Mult (Var gp, ge), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Neg gq, ge), fi), af)) = - Eq (Add (Mult (Neg gq, ge), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Add (gr, gs), ge), fi), af)) = - Eq (Add (Mult (Add (gr, gs), ge), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Sub (gt, gu), ge), fi), af)) = - Eq (Add (Mult (Sub (gt, gu), ge), fi), af) - | adjustcoeff (l, Eq (Add (Mult (Mult (gv, gw), ge), fi), af)) = - Eq (Add (Mult (Mult (gv, gw), ge), fi), af) - | adjustcoeff (l, Eq (Sub (fj, fk), af)) = Eq (Sub (fj, fk), af) - | adjustcoeff (l, Eq (Mult (fl, fm), af)) = Eq (Mult (fl, fm), af) - | adjustcoeff (l, Divides (Cst is, Cst jk)) = Divides (Cst is, Cst jk) - | adjustcoeff (l, Divides (Cst is, Var jl)) = Divides (Cst is, Var jl) - | adjustcoeff (l, Divides (Cst is, Neg jm)) = Divides (Cst is, Neg jm) - | adjustcoeff (l, Divides (Cst is, Add (Cst kc, jo))) = - Divides (Cst is, Add (Cst kc, jo)) - | adjustcoeff (l, Divides (Cst is, Add (Var kd, jo))) = - Divides (Cst is, Add (Var kd, jo)) - | adjustcoeff (l, Divides (Cst is, Add (Neg ke, jo))) = - Divides (Cst is, Add (Neg ke, jo)) - | adjustcoeff (l, Divides (Cst is, Add (Add (kf, kg), jo))) = - Divides (Cst is, Add (Add (kf, kg), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Sub (kh, ki), jo))) = - Divides (Cst is, Add (Sub (kh, ki), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo))) = - Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Var me), jo))) = - (if (me = 0) - then Divides - (Cst (op_div_def1 l ku * is), - Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l ku, jo))) - else Divides - (Cst is, - Add (Mult (Cst ku, Var (op_45_def0 me id_1_def0 + 1)), jo))) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo))) = - Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo))) = - Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo))) = - Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo))) = - Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Var kv, kk), jo))) = - Divides (Cst is, Add (Mult (Var kv, kk), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Neg kw, kk), jo))) = - Divides (Cst is, Add (Mult (Neg kw, kk), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo))) = - Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo))) = - Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo)) - | adjustcoeff (l, Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo))) = - Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo)) - | adjustcoeff (l, Divides (Cst is, Sub (jp, jq))) = - Divides (Cst is, Sub (jp, jq)) - | adjustcoeff (l, Divides (Cst is, Mult (jr, js))) = - Divides (Cst is, Mult (jr, js)) - | adjustcoeff (l, Divides (Var it, ah)) = Divides (Var it, ah) - | adjustcoeff (l, Divides (Neg iu, ah)) = Divides (Neg iu, ah) - | adjustcoeff (l, Divides (Add (iv, iw), ah)) = Divides (Add (iv, iw), ah) - | adjustcoeff (l, Divides (Sub (ix, iy), ah)) = Divides (Sub (ix, iy), ah) - | adjustcoeff (l, Divides (Mult (iz, ja), ah)) = Divides (Mult (iz, ja), ah) - | adjustcoeff (l, T) = T - | adjustcoeff (l, F) = F - | adjustcoeff (l, NOT (Lt (mg, mh))) = NOT (Lt (mg, mh)) - | adjustcoeff (l, NOT (Gt (mi, mj))) = NOT (Gt (mi, mj)) - | adjustcoeff (l, NOT (Le (mk, ml))) = NOT (Le (mk, ml)) - | adjustcoeff (l, NOT (Ge (mm, mn))) = NOT (Ge (mm, mn)) - | adjustcoeff (l, NOT (Eq (Cst oa, mp))) = NOT (Eq (Cst oa, mp)) - | adjustcoeff (l, NOT (Eq (Var ob, mp))) = NOT (Eq (Var ob, mp)) - | adjustcoeff (l, NOT (Eq (Neg oc, mp))) = NOT (Eq (Neg oc, mp)) - | adjustcoeff (l, NOT (Eq (Add (Cst os, oe), mp))) = - NOT (Eq (Add (Cst os, oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Var ot, oe), mp))) = - NOT (Eq (Add (Var ot, oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Neg ou, oe), mp))) = - NOT (Eq (Add (Neg ou, oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Add (ov, ow), oe), mp))) = - NOT (Eq (Add (Add (ov, ow), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Sub (ox, oy), oe), mp))) = - NOT (Eq (Add (Sub (ox, oy), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx))) = - NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy))) = - NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra)))) = - NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra))) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc)))) = - NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc))) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re)))) = - NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re))) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp))) = - NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Var pl, pa), oe), mp))) = - NOT (Eq (Add (Mult (Var pl, pa), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Neg pm, pa), oe), mp))) = - NOT (Eq (Add (Mult (Neg pm, pa), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp))) = - NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp))) = - NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp)) - | adjustcoeff (l, NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp))) = - NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp)) - | adjustcoeff (l, NOT (Eq (Sub (of', og), mp))) = NOT (Eq (Sub (of', og), mp)) - | adjustcoeff (l, NOT (Eq (Mult (oh, oi), mp))) = NOT (Eq (Mult (oh, oi), mp)) - | adjustcoeff (l, NOT (Divides (Cst ro, Cst sg))) = - NOT (Divides (Cst ro, Cst sg)) - | adjustcoeff (l, NOT (Divides (Cst ro, Var sh))) = - NOT (Divides (Cst ro, Var sh)) - | adjustcoeff (l, NOT (Divides (Cst ro, Neg si))) = - NOT (Divides (Cst ro, Neg si)) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Cst sy, sk)))) = - NOT (Divides (Cst ro, Add (Cst sy, sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Var sz, sk)))) = - NOT (Divides (Cst ro, Add (Var sz, sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Neg ta, sk)))) = - NOT (Divides (Cst ro, Add (Neg ta, sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Add (tb, tc), sk)))) = - NOT (Divides (Cst ro, Add (Add (tb, tc), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Sub (td, te), sk)))) = - NOT (Divides (Cst ro, Add (Sub (td, te), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Var va), sk)))) = - (if (va = 0) - then NOT (Divides - (Cst (op_div_def1 l tq * ro), - Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l tq, sk)))) - else NOT (Divides - (Cst ro, - Add (Mult (Cst tq, Var (op_45_def0 va id_1_def0 + 1)), - sk)))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk))) - | adjustcoeff - (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk))) - | adjustcoeff - (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk))) - | adjustcoeff - (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk))) - | adjustcoeff - (l, NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk)))) = - NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk))) - | adjustcoeff (l, NOT (Divides (Cst ro, Sub (sl, sm)))) = - NOT (Divides (Cst ro, Sub (sl, sm))) - | adjustcoeff (l, NOT (Divides (Cst ro, Mult (sn, so)))) = - NOT (Divides (Cst ro, Mult (sn, so))) - | adjustcoeff (l, NOT (Divides (Var rp, mr))) = NOT (Divides (Var rp, mr)) - | adjustcoeff (l, NOT (Divides (Neg rq, mr))) = NOT (Divides (Neg rq, mr)) - | adjustcoeff (l, NOT (Divides (Add (rr, rs), mr))) = - NOT (Divides (Add (rr, rs), mr)) - | adjustcoeff (l, NOT (Divides (Sub (rt, ru), mr))) = - NOT (Divides (Sub (rt, ru), mr)) - | adjustcoeff (l, NOT (Divides (Mult (rv, rw), mr))) = - NOT (Divides (Mult (rv, rw), mr)) - | adjustcoeff (l, NOT T) = NOT T - | adjustcoeff (l, NOT F) = NOT F - | adjustcoeff (l, NOT (NOT ms)) = NOT (NOT ms) - | adjustcoeff (l, NOT (And (mt, mu))) = NOT (And (mt, mu)) - | adjustcoeff (l, NOT (Or (mv, mw))) = NOT (Or (mv, mw)) - | adjustcoeff (l, NOT (Imp (mx, my))) = NOT (Imp (mx, my)) - | adjustcoeff (l, NOT (Equ (mz, na))) = NOT (Equ (mz, na)) - | adjustcoeff (l, NOT (QAll nb)) = NOT (QAll nb) - | adjustcoeff (l, NOT (QEx nc)) = NOT (QEx nc) - | adjustcoeff (l, Imp (an, ao)) = Imp (an, ao) - | adjustcoeff (l, Equ (ap, aq)) = Equ (ap, aq) - | adjustcoeff (l, QAll ar) = QAll ar - | adjustcoeff (l, QEx as') = QEx as'; - -fun formlcm (Le (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c - | formlcm (Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c - | formlcm (NOT p) = formlcm p - | formlcm (And (p, q)) = ilcm (formlcm p) (formlcm q) - | formlcm (Or (p, q)) = ilcm (formlcm p) (formlcm q) - | formlcm (Lt (u, v)) = 1 - | formlcm (Gt (w, x)) = 1 - | formlcm (Le (Cst bo, z)) = 1 - | formlcm (Le (Var bp, z)) = 1 - | formlcm (Le (Neg bq, z)) = 1 - | formlcm (Le (Add (Cst cg, bs), z)) = 1 - | formlcm (Le (Add (Var ch, bs), z)) = 1 - | formlcm (Le (Add (Neg ci, bs), z)) = 1 - | formlcm (Le (Add (Add (cj, ck), bs), z)) = 1 - | formlcm (Le (Add (Sub (cl, cm), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Var el)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Neg em)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Add (en, eo))) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Sub (ep, eq))) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Mult (er, es))) = 1 - | formlcm (Le (Add (Mult (Cst cy, Var ei ), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = 1 - | formlcm (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = 1 - | formlcm (Le (Add (Mult (Var cz, co), bs), z)) = 1 - | formlcm (Le (Add (Mult (Neg da, co), bs), z)) = 1 - | formlcm (Le (Add (Mult (Add (db, dc), co), bs), z)) = 1 - | formlcm (Le (Add (Mult (Sub (dd, de), co), bs), z)) = 1 - | formlcm (Le (Add (Mult (Mult (df, dg), co), bs), z)) = 1 - | formlcm (Le (Sub (bt, bu), z)) = 1 - | formlcm (Le (Mult (bv, bw), z)) = 1 - | formlcm (Ge (aa, ab)) = 1 - | formlcm (Eq (Cst fc, ad)) = 1 - | formlcm (Eq (Var fd, ad)) = 1 - | formlcm (Eq (Neg fe, ad)) = 1 - | formlcm (Eq (Add (Cst fu, fg), ad)) = 1 - | formlcm (Eq (Add (Var fv, fg), ad)) = 1 - | formlcm (Eq (Add (Neg fw, fg), ad)) = 1 - | formlcm (Eq (Add (Add (fx, fy), fg), ad)) = 1 - | formlcm (Eq (Add (Sub (fz, ga), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Cst he), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Var hz)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Neg ia)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Add (ib, ic))) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Sub (id, ie))) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Mult (if', ig))) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Var hw), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Neg hg), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Add (hh, hi)), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Sub (hj, hk)), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Cst gm, Mult (hl, hm)), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Var gn, gc), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Neg go, gc), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Add (gp, gq), gc), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Sub (gr, gs), gc), fg), ad)) = 1 - | formlcm (Eq (Add (Mult (Mult (gt, gu), gc), fg), ad)) = 1 - | formlcm (Eq (Sub (fh, fi), ad)) = 1 - | formlcm (Eq (Mult (fj, fk), ad)) = 1 - | formlcm (Divides (Cst iq, Cst ji)) = 1 - | formlcm (Divides (Cst iq, Var jj)) = 1 - | formlcm (Divides (Cst iq, Neg jk)) = 1 - | formlcm (Divides (Cst iq, Add (Cst ka, jm))) = 1 - | formlcm (Divides (Cst iq, Add (Var kb, jm))) = 1 - | formlcm (Divides (Cst iq, Add (Neg kc, jm))) = 1 - | formlcm (Divides (Cst iq, Add (Add (kd, ke), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Sub (kf, kg), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Cst lk), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Var mc), jm))) = - (if (mc = 0) then abs ks else 1) - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Neg lm), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Add (ln, lo)), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Sub (lp, lq)), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Mult (lr, ls)), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Var kt, ki), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Neg ku, ki), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Add (kv, kw), ki), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Sub (kx, ky), ki), jm))) = 1 - | formlcm (Divides (Cst iq, Add (Mult (Mult (kz, la), ki), jm))) = 1 - | formlcm (Divides (Cst iq, Sub (jn, jo))) = 1 - | formlcm (Divides (Cst iq, Mult (jp, jq))) = 1 - | formlcm (Divides (Var ir, af)) = 1 - | formlcm (Divides (Neg is, af)) = 1 - | formlcm (Divides (Add (it, iu), af)) = 1 - | formlcm (Divides (Sub (iv, iw), af)) = 1 - | formlcm (Divides (Mult (ix, iy), af)) = 1 - | formlcm T = 1 - | formlcm F = 1 - | formlcm (Imp (al, am)) = 1 - | formlcm (Equ (an, ao)) = 1 - | formlcm (QAll ap) = 1 - | formlcm (QEx aq) = 1; - -fun unitycoeff p = - let val l = formlcm p; val p' = adjustcoeff (l, p) - in (if (l = 1) then p' - else And (Divides (Cst l, Add (Mult (Cst 1, Var 0), Cst 0)), p')) - end; - -fun unify p = - let val q = unitycoeff p; val B = list_set (bset q); val A = list_set (aset q) - in (if op_60_61_def0 (size_def1 B) (size_def1 A) then (q, B) - else (mirror q, map lin_neg A)) - end; - -fun cooper p = - lift_un (fn q => decrvars (explode_minf (unify q))) (linform (nnf p)); - -fun pa p = lift_un psimpl (qelim (cooper, p)); - -val test = pa; - -end;
--- a/src/HOL/IsaMakefile Thu May 31 11:00:06 2007 +0200 +++ b/src/HOL/IsaMakefile Thu May 31 12:06:31 2007 +0200 @@ -62,8 +62,7 @@ Pure: @cd $(SRC)/Pure; $(ISATOOL) make Pure -$(OUT)/HOL: $(OUT)/Pure $(SRC)/Pure/General/int.ML $(SRC)/Pure/General/rat.ML \ - $(SRC)/Provers/Arith/abel_cancel.ML \ +$(OUT)/HOL: $(OUT)/Pure $(SRC)/Provers/Arith/abel_cancel.ML \ $(SRC)/Provers/Arith/assoc_fold.ML \ $(SRC)/Provers/Arith/cancel_div_mod.ML \ $(SRC)/Provers/Arith/cancel_numeral_factor.ML \ @@ -81,40 +80,42 @@ $(SRC)/Provers/induct_method.ML $(SRC)/Provers/order.ML \ $(SRC)/Provers/project_rule.ML $(SRC)/Provers/quantifier1.ML \ $(SRC)/Provers/quasi.ML $(SRC)/Provers/splitter.ML \ - $(SRC)/Provers/trancl.ML $(SRC)/TFL/casesplit.ML $(SRC)/TFL/dcterm.ML \ - $(SRC)/TFL/post.ML $(SRC)/TFL/rules.ML $(SRC)/TFL/tfl.ML \ - $(SRC)/TFL/thms.ML $(SRC)/TFL/thry.ML $(SRC)/TFL/usyntax.ML \ - $(SRC)/TFL/utils.ML ATP_Linkup.thy Accessible_Part.thy \ - Code_Generator.thy Datatype.thy Divides.thy Equiv_Relations.thy \ - Extraction.thy Finite_Set.thy FixedPoint.thy Fun.thy FunDef.thy \ - HOL.thy Hilbert_Choice.thy Inductive.thy Integ/IntArith.thy \ - Integ/IntDef.thy Integ/IntDiv.thy Integ/NatBin.thy \ - Integ/NatSimprocs.thy Integ/Numeral.thy Integ/Presburger.thy \ - Integ/cooper_dec.ML Integ/cooper_proof.ML Integ/int_arith1.ML \ - Integ/int_factor_simprocs.ML Integ/nat_simprocs.ML Integ/presburger.ML \ - Integ/qelim.ML Integ/reflected_cooper.ML Integ/reflected_presburger.ML \ - Lattices.thy List.thy Main.thy Map.thy Nat.ML Nat.thy \ - OrderedGroup.thy Orderings.thy Power.thy Predicate.thy PreList.thy \ + $(SRC)/Provers/trancl.ML $(SRC)/Pure/General/int.ML \ + $(SRC)/Pure/General/rat.ML $(SRC)/TFL/casesplit.ML \ + $(SRC)/TFL/dcterm.ML $(SRC)/TFL/post.ML $(SRC)/TFL/rules.ML \ + $(SRC)/TFL/tfl.ML $(SRC)/TFL/thms.ML $(SRC)/TFL/thry.ML \ + $(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML ATP_Linkup.thy \ + Accessible_Part.thy Code_Generator.thy Datatype.thy Divides.thy \ + Equiv_Relations.thy Extraction.thy Finite_Set.thy FixedPoint.thy \ + Fun.thy FunDef.thy HOL.thy Hilbert_Choice.thy Inductive.thy \ + Integ/IntArith.thy Integ/IntDef.thy Integ/IntDiv.thy Integ/NatBin.thy \ + Integ/NatSimprocs.thy Integ/Numeral.thy Integ/int_arith1.ML \ + Integ/int_factor_simprocs.ML Integ/nat_simprocs.ML Lattices.thy \ + List.thy Main.thy Map.thy Nat.ML Nat.thy OrderedGroup.thy \ + Orderings.thy Power.thy PreList.thy Predicate.thy Presburger.thy \ Product_Type.thy ROOT.ML Recdef.thy Record.thy Refute.thy Relation.thy \ Relation_Power.thy Ring_and_Field.thy SAT.thy Set.thy SetInterval.thy \ - Sum_Type.thy Tools/res_reconstruct.ML Tools/ATP/reduce_axiomsN.ML \ - Tools/ATP/watcher.ML Tools/cnf_funcs.ML Tools/datatype_abs_proofs.ML \ - Tools/datatype_aux.ML Tools/datatype_case.ML Tools/datatype_codegen.ML \ + Sum_Type.thy Tools/ATP/reduce_axiomsN.ML Tools/ATP/watcher.ML \ + Tools/Presburger/cooper_dec.ML Tools/Presburger/cooper_proof.ML \ + Tools/Presburger/presburger.ML Tools/Presburger/qelim.ML \ + Tools/Presburger/reflected_cooper.ML \ + Tools/Presburger/reflected_presburger.ML Tools/cnf_funcs.ML \ + Tools/datatype_abs_proofs.ML Tools/datatype_aux.ML \ + Tools/datatype_case.ML Tools/datatype_codegen.ML \ Tools/datatype_hooks.ML Tools/datatype_package.ML \ Tools/datatype_prop.ML Tools/datatype_realizer.ML \ Tools/datatype_rep_proofs.ML Tools/function_package/auto_term.ML \ Tools/function_package/context_tree.ML \ Tools/function_package/fundef_common.ML \ + Tools/function_package/fundef_core.ML \ Tools/function_package/fundef_datatype.ML \ Tools/function_package/fundef_lib.ML \ Tools/function_package/fundef_package.ML \ - Tools/function_package/fundef_core.ML \ Tools/function_package/inductive_wrap.ML \ Tools/function_package/lexicographic_order.ML \ Tools/function_package/mutual.ML \ Tools/function_package/pattern_split.ML \ - Tools/function_package/sum_tools.ML \ - Tools/inductive_codegen.ML \ + Tools/function_package/sum_tools.ML Tools/inductive_codegen.ML \ Tools/inductive_package.ML Tools/inductive_realizer.ML Tools/meson.ML \ Tools/numeral_syntax.ML Tools/old_inductive_package.ML \ Tools/polyhash.ML Tools/primrec_package.ML Tools/prop_logic.ML \ @@ -122,8 +123,9 @@ Tools/record_package.ML Tools/refute.ML Tools/refute_isar.ML \ Tools/res_atp.ML Tools/res_atp_methods.ML Tools/res_atp_provers.ML \ Tools/res_atpset.ML Tools/res_axioms.ML Tools/res_clause.ML \ - Tools/res_hol_clause.ML Tools/rewrite_hol_proof.ML Tools/sat_funcs.ML \ - Tools/sat_solver.ML Tools/specification_package.ML Tools/split_rule.ML \ + Tools/res_hol_clause.ML Tools/res_reconstruct.ML \ + Tools/rewrite_hol_proof.ML Tools/sat_funcs.ML Tools/sat_solver.ML \ + Tools/specification_package.ML Tools/split_rule.ML \ Tools/string_syntax.ML Tools/typecopy_package.ML \ Tools/typedef_codegen.ML Tools/typedef_package.ML \ Transitive_Closure.thy Typedef.thy Wellfounded_Recursion.thy \
--- a/src/HOL/Presburger.thy Thu May 31 11:00:06 2007 +0200 +++ b/src/HOL/Presburger.thy Thu May 31 12:06:31 2007 +0200 @@ -9,10 +9,14 @@ header {* Presburger Arithmetic: Cooper's Algorithm *} theory Presburger -imports NatSimprocs "../SetInterval" +imports "Integ/NatSimprocs" SetInterval uses - ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") - ("reflected_presburger.ML") ("reflected_cooper.ML") ("presburger.ML") + ("Tools/Presburger/cooper_dec.ML") + ("Tools/Presburger/cooper_proof.ML") + ("Tools/Presburger/qelim.ML") + ("Tools/Presburger/reflected_presburger.ML") + ("Tools/Presburger/reflected_cooper.ML") + ("Tools/Presburger/presburger.ML") begin text {* Theorem for unitifying the coeffitients of @{text x} in an existential formula*} @@ -1047,15 +1051,15 @@ show ?thesis by (simp add: 1) qed -use "cooper_dec.ML" -use "reflected_presburger.ML" -use "reflected_cooper.ML" +use "Tools/Presburger/cooper_dec.ML" +use "Tools/Presburger/reflected_presburger.ML" +use "Tools/Presburger/reflected_cooper.ML" oracle presburger_oracle ("term") = ReflectedCooper.presburger_oracle -use "cooper_proof.ML" -use "qelim.ML" -use "presburger.ML" +use "Tools/Presburger/cooper_proof.ML" +use "Tools/Presburger/qelim.ML" +use "Tools/Presburger/presburger.ML" setup "Presburger.setup"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Bin.thy Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,692 @@ +(* Title: ZF/Bin.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + + The sign Pls stands for an infinite string of leading 0's. + The sign Min stands for an infinite string of leading 1's. + +A number can have multiple representations, namely leading 0's with sign +Pls and leading 1's with sign Min. See twos-compl.ML/int_of_binary for +the numerical interpretation. + +The representation expects that (m mod 2) is 0 or 1, even if m is negative; +For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 +*) + +header{*Arithmetic on Binary Integers*} + +theory Bin +imports Int Datatype +uses "Tools/numeral_syntax.ML" +begin + +consts bin :: i +datatype + "bin" = Pls + | Min + | Bit ("w: bin", "b: bool") (infixl "BIT" 90) + +syntax + "_Int" :: "xnum => i" ("_") + +consts + integ_of :: "i=>i" + NCons :: "[i,i]=>i" + bin_succ :: "i=>i" + bin_pred :: "i=>i" + bin_minus :: "i=>i" + bin_adder :: "i=>i" + bin_mult :: "[i,i]=>i" + +primrec + integ_of_Pls: "integ_of (Pls) = $# 0" + integ_of_Min: "integ_of (Min) = $-($#1)" + integ_of_BIT: "integ_of (w BIT b) = $#b $+ integ_of(w) $+ integ_of(w)" + + (** recall that cond(1,b,c)=b and cond(0,b,c)=0 **) + +primrec (*NCons adds a bit, suppressing leading 0s and 1s*) + NCons_Pls: "NCons (Pls,b) = cond(b,Pls BIT b,Pls)" + NCons_Min: "NCons (Min,b) = cond(b,Min,Min BIT b)" + NCons_BIT: "NCons (w BIT c,b) = w BIT c BIT b" + +primrec (*successor. If a BIT, can change a 0 to a 1 without recursion.*) + bin_succ_Pls: "bin_succ (Pls) = Pls BIT 1" + bin_succ_Min: "bin_succ (Min) = Pls" + bin_succ_BIT: "bin_succ (w BIT b) = cond(b, bin_succ(w) BIT 0, NCons(w,1))" + +primrec (*predecessor*) + bin_pred_Pls: "bin_pred (Pls) = Min" + bin_pred_Min: "bin_pred (Min) = Min BIT 0" + bin_pred_BIT: "bin_pred (w BIT b) = cond(b, NCons(w,0), bin_pred(w) BIT 1)" + +primrec (*unary negation*) + bin_minus_Pls: + "bin_minus (Pls) = Pls" + bin_minus_Min: + "bin_minus (Min) = Pls BIT 1" + bin_minus_BIT: + "bin_minus (w BIT b) = cond(b, bin_pred(NCons(bin_minus(w),0)), + bin_minus(w) BIT 0)" + +primrec (*sum*) + bin_adder_Pls: + "bin_adder (Pls) = (lam w:bin. w)" + bin_adder_Min: + "bin_adder (Min) = (lam w:bin. bin_pred(w))" + bin_adder_BIT: + "bin_adder (v BIT x) = + (lam w:bin. + bin_case (v BIT x, bin_pred(v BIT x), + %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w), + x xor y), + w))" + +(*The bin_case above replaces the following mutually recursive function: +primrec + "adding (v,x,Pls) = v BIT x" + "adding (v,x,Min) = bin_pred(v BIT x)" + "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)), + x xor y)" +*) + +constdefs + bin_add :: "[i,i]=>i" + "bin_add(v,w) == bin_adder(v)`w" + + +primrec + bin_mult_Pls: + "bin_mult (Pls,w) = Pls" + bin_mult_Min: + "bin_mult (Min,w) = bin_minus(w)" + bin_mult_BIT: + "bin_mult (v BIT b,w) = cond(b, bin_add(NCons(bin_mult(v,w),0),w), + NCons(bin_mult(v,w),0))" + +setup NumeralSyntax.setup + + +declare bin.intros [simp,TC] + +lemma NCons_Pls_0: "NCons(Pls,0) = Pls" +by simp + +lemma NCons_Pls_1: "NCons(Pls,1) = Pls BIT 1" +by simp + +lemma NCons_Min_0: "NCons(Min,0) = Min BIT 0" +by simp + +lemma NCons_Min_1: "NCons(Min,1) = Min" +by simp + +lemma NCons_BIT: "NCons(w BIT x,b) = w BIT x BIT b" +by (simp add: bin.case_eqns) + +lemmas NCons_simps [simp] = + NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT + + + +(** Type checking **) + +lemma integ_of_type [TC]: "w: bin ==> integ_of(w) : int" +apply (induct_tac "w") +apply (simp_all add: bool_into_nat) +done + +lemma NCons_type [TC]: "[| w: bin; b: bool |] ==> NCons(w,b) : bin" +by (induct_tac "w", auto) + +lemma bin_succ_type [TC]: "w: bin ==> bin_succ(w) : bin" +by (induct_tac "w", auto) + +lemma bin_pred_type [TC]: "w: bin ==> bin_pred(w) : bin" +by (induct_tac "w", auto) + +lemma bin_minus_type [TC]: "w: bin ==> bin_minus(w) : bin" +by (induct_tac "w", auto) + +(*This proof is complicated by the mutual recursion*) +lemma bin_add_type [rule_format,TC]: + "v: bin ==> ALL w: bin. bin_add(v,w) : bin" +apply (unfold bin_add_def) +apply (induct_tac "v") +apply (rule_tac [3] ballI) +apply (rename_tac [3] "w'") +apply (induct_tac [3] "w'") +apply (simp_all add: NCons_type) +done + +lemma bin_mult_type [TC]: "[| v: bin; w: bin |] ==> bin_mult(v,w) : bin" +by (induct_tac "v", auto) + + +subsubsection{*The Carry and Borrow Functions, + @{term bin_succ} and @{term bin_pred}*} + +(*NCons preserves the integer value of its argument*) +lemma integ_of_NCons [simp]: + "[| w: bin; b: bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)" +apply (erule bin.cases) +apply (auto elim!: boolE) +done + +lemma integ_of_succ [simp]: + "w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)" +apply (erule bin.induct) +apply (auto simp add: zadd_ac elim!: boolE) +done + +lemma integ_of_pred [simp]: + "w: bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" +apply (erule bin.induct) +apply (auto simp add: zadd_ac elim!: boolE) +done + + +subsubsection{*@{term bin_minus}: Unary Negation of Binary Integers*} + +lemma integ_of_minus: "w: bin ==> integ_of(bin_minus(w)) = $- integ_of(w)" +apply (erule bin.induct) +apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE) +done + + +subsubsection{*@{term bin_add}: Binary Addition*} + +lemma bin_add_Pls [simp]: "w: bin ==> bin_add(Pls,w) = w" +by (unfold bin_add_def, simp) + +lemma bin_add_Pls_right: "w: bin ==> bin_add(w,Pls) = w" +apply (unfold bin_add_def) +apply (erule bin.induct, auto) +done + +lemma bin_add_Min [simp]: "w: bin ==> bin_add(Min,w) = bin_pred(w)" +by (unfold bin_add_def, simp) + +lemma bin_add_Min_right: "w: bin ==> bin_add(w,Min) = bin_pred(w)" +apply (unfold bin_add_def) +apply (erule bin.induct, auto) +done + +lemma bin_add_BIT_Pls [simp]: "bin_add(v BIT x,Pls) = v BIT x" +by (unfold bin_add_def, simp) + +lemma bin_add_BIT_Min [simp]: "bin_add(v BIT x,Min) = bin_pred(v BIT x)" +by (unfold bin_add_def, simp) + +lemma bin_add_BIT_BIT [simp]: + "[| w: bin; y: bool |] + ==> bin_add(v BIT x, w BIT y) = + NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)" +by (unfold bin_add_def, simp) + +lemma integ_of_add [rule_format]: + "v: bin ==> + ALL w: bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)" +apply (erule bin.induct, simp, simp) +apply (rule ballI) +apply (induct_tac "wa") +apply (auto simp add: zadd_ac elim!: boolE) +done + +(*Subtraction*) +lemma diff_integ_of_eq: + "[| v: bin; w: bin |] + ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" +apply (unfold zdiff_def) +apply (simp add: integ_of_add integ_of_minus) +done + + +subsubsection{*@{term bin_mult}: Binary Multiplication*} + +lemma integ_of_mult: + "[| v: bin; w: bin |] + ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)" +apply (induct_tac "v", simp) +apply (simp add: integ_of_minus) +apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib elim!: boolE) +done + + +subsection{*Computations*} + +(** extra rules for bin_succ, bin_pred **) + +lemma bin_succ_1: "bin_succ(w BIT 1) = bin_succ(w) BIT 0" +by simp + +lemma bin_succ_0: "bin_succ(w BIT 0) = NCons(w,1)" +by simp + +lemma bin_pred_1: "bin_pred(w BIT 1) = NCons(w,0)" +by simp + +lemma bin_pred_0: "bin_pred(w BIT 0) = bin_pred(w) BIT 1" +by simp + +(** extra rules for bin_minus **) + +lemma bin_minus_1: "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))" +by simp + +lemma bin_minus_0: "bin_minus(w BIT 0) = bin_minus(w) BIT 0" +by simp + +(** extra rules for bin_add **) + +lemma bin_add_BIT_11: "w: bin ==> bin_add(v BIT 1, w BIT 1) = + NCons(bin_add(v, bin_succ(w)), 0)" +by simp + +lemma bin_add_BIT_10: "w: bin ==> bin_add(v BIT 1, w BIT 0) = + NCons(bin_add(v,w), 1)" +by simp + +lemma bin_add_BIT_0: "[| w: bin; y: bool |] + ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)" +by simp + +(** extra rules for bin_mult **) + +lemma bin_mult_1: "bin_mult(v BIT 1, w) = bin_add(NCons(bin_mult(v,w),0), w)" +by simp + +lemma bin_mult_0: "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)" +by simp + + +(** Simplification rules with integer constants **) + +lemma int_of_0: "$#0 = #0" +by simp + +lemma int_of_succ: "$# succ(n) = #1 $+ $#n" +by (simp add: int_of_add [symmetric] natify_succ) + +lemma zminus_0 [simp]: "$- #0 = #0" +by simp + +lemma zadd_0_intify [simp]: "#0 $+ z = intify(z)" +by simp + +lemma zadd_0_right_intify [simp]: "z $+ #0 = intify(z)" +by simp + +lemma zmult_1_intify [simp]: "#1 $* z = intify(z)" +by simp + +lemma zmult_1_right_intify [simp]: "z $* #1 = intify(z)" +by (subst zmult_commute, simp) + +lemma zmult_0 [simp]: "#0 $* z = #0" +by simp + +lemma zmult_0_right [simp]: "z $* #0 = #0" +by (subst zmult_commute, simp) + +lemma zmult_minus1 [simp]: "#-1 $* z = $-z" +by (simp add: zcompare_rls) + +lemma zmult_minus1_right [simp]: "z $* #-1 = $-z" +apply (subst zmult_commute) +apply (rule zmult_minus1) +done + + +subsection{*Simplification Rules for Comparison of Binary Numbers*} +text{*Thanks to Norbert Voelker*} + +(** Equals (=) **) + +lemma eq_integ_of_eq: + "[| v: bin; w: bin |] + ==> ((integ_of(v)) = integ_of(w)) <-> + iszero (integ_of (bin_add (v, bin_minus(w))))" +apply (unfold iszero_def) +apply (simp add: zcompare_rls integ_of_add integ_of_minus) +done + +lemma iszero_integ_of_Pls: "iszero (integ_of(Pls))" +by (unfold iszero_def, simp) + + +lemma nonzero_integ_of_Min: "~ iszero (integ_of(Min))" +apply (unfold iszero_def) +apply (simp add: zminus_equation) +done + +lemma iszero_integ_of_BIT: + "[| w: bin; x: bool |] + ==> iszero (integ_of (w BIT x)) <-> (x=0 & iszero (integ_of(w)))" +apply (unfold iszero_def, simp) +apply (subgoal_tac "integ_of (w) : int") +apply typecheck +apply (drule int_cases) +apply (safe elim!: boolE) +apply (simp_all (asm_lr) add: zcompare_rls zminus_zadd_distrib [symmetric] + int_of_add [symmetric]) +done + +lemma iszero_integ_of_0: + "w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))" +by (simp only: iszero_integ_of_BIT, blast) + +lemma iszero_integ_of_1: "w: bin ==> ~ iszero (integ_of (w BIT 1))" +by (simp only: iszero_integ_of_BIT, blast) + + + +(** Less-than (<) **) + +lemma less_integ_of_eq_neg: + "[| v: bin; w: bin |] + ==> integ_of(v) $< integ_of(w) + <-> znegative (integ_of (bin_add (v, bin_minus(w))))" +apply (unfold zless_def zdiff_def) +apply (simp add: integ_of_minus integ_of_add) +done + +lemma not_neg_integ_of_Pls: "~ znegative (integ_of(Pls))" +by simp + +lemma neg_integ_of_Min: "znegative (integ_of(Min))" +by simp + +lemma neg_integ_of_BIT: + "[| w: bin; x: bool |] + ==> znegative (integ_of (w BIT x)) <-> znegative (integ_of(w))" +apply simp +apply (subgoal_tac "integ_of (w) : int") +apply typecheck +apply (drule int_cases) +apply (auto elim!: boolE simp add: int_of_add [symmetric] zcompare_rls) +apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def + int_of_add [symmetric]) +apply (subgoal_tac "$#1 $- $# succ (succ (n #+ n)) = $- $# succ (n #+ n) ") + apply (simp add: zdiff_def) +apply (simp add: equation_zminus int_of_diff [symmetric]) +done + +(** Less-than-or-equals (<=) **) + +lemma le_integ_of_eq_not_less: + "(integ_of(x) $<= (integ_of(w))) <-> ~ (integ_of(w) $< (integ_of(x)))" +by (simp add: not_zless_iff_zle [THEN iff_sym]) + + +(*Delete the original rewrites, with their clumsy conditional expressions*) +declare bin_succ_BIT [simp del] + bin_pred_BIT [simp del] + bin_minus_BIT [simp del] + NCons_Pls [simp del] + NCons_Min [simp del] + bin_adder_BIT [simp del] + bin_mult_BIT [simp del] + +(*Hide the binary representation of integer constants*) +declare integ_of_Pls [simp del] integ_of_Min [simp del] integ_of_BIT [simp del] + + +lemmas bin_arith_extra_simps = + integ_of_add [symmetric] + integ_of_minus [symmetric] + integ_of_mult [symmetric] + bin_succ_1 bin_succ_0 + bin_pred_1 bin_pred_0 + bin_minus_1 bin_minus_0 + bin_add_Pls_right bin_add_Min_right + bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11 + diff_integ_of_eq + bin_mult_1 bin_mult_0 NCons_simps + + +(*For making a minimal simpset, one must include these default simprules + of thy. Also include simp_thms, or at least (~False)=True*) +lemmas bin_arith_simps = + bin_pred_Pls bin_pred_Min + bin_succ_Pls bin_succ_Min + bin_add_Pls bin_add_Min + bin_minus_Pls bin_minus_Min + bin_mult_Pls bin_mult_Min + bin_arith_extra_simps + +(*Simplification of relational operations*) +lemmas bin_rel_simps = + eq_integ_of_eq iszero_integ_of_Pls nonzero_integ_of_Min + iszero_integ_of_0 iszero_integ_of_1 + less_integ_of_eq_neg + not_neg_integ_of_Pls neg_integ_of_Min neg_integ_of_BIT + le_integ_of_eq_not_less + +declare bin_arith_simps [simp] +declare bin_rel_simps [simp] + + +(** Simplification of arithmetic when nested to the right **) + +lemma add_integ_of_left [simp]: + "[| v: bin; w: bin |] + ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)" +by (simp add: zadd_assoc [symmetric]) + +lemma mult_integ_of_left [simp]: + "[| v: bin; w: bin |] + ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)" +by (simp add: zmult_assoc [symmetric]) + +lemma add_integ_of_diff1 [simp]: + "[| v: bin; w: bin |] + ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)" +apply (unfold zdiff_def) +apply (rule add_integ_of_left, auto) +done + +lemma add_integ_of_diff2 [simp]: + "[| v: bin; w: bin |] + ==> integ_of(v) $+ (c $- integ_of(w)) = + integ_of (bin_add (v, bin_minus(w))) $+ (c)" +apply (subst diff_integ_of_eq [symmetric]) +apply (simp_all add: zdiff_def zadd_ac) +done + + +(** More for integer constants **) + +declare int_of_0 [simp] int_of_succ [simp] + +lemma zdiff0 [simp]: "#0 $- x = $-x" +by (simp add: zdiff_def) + +lemma zdiff0_right [simp]: "x $- #0 = intify(x)" +by (simp add: zdiff_def) + +lemma zdiff_self [simp]: "x $- x = #0" +by (simp add: zdiff_def) + +lemma znegative_iff_zless_0: "k: int ==> znegative(k) <-> k $< #0" +by (simp add: zless_def) + +lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k: int|] ==> znegative($-k)" +by (simp add: zless_def) + +lemma zero_zle_int_of [simp]: "#0 $<= $# n" +by (simp add: not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym]) + +lemma nat_of_0 [simp]: "nat_of(#0) = 0" +by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of) + +lemma nat_le_int0_lemma: "[| z $<= $#0; z: int |] ==> nat_of(z) = 0" +by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of) + +lemma nat_le_int0: "z $<= $#0 ==> nat_of(z) = 0" +apply (subgoal_tac "nat_of (intify (z)) = 0") +apply (rule_tac [2] nat_le_int0_lemma, auto) +done + +lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 ==> natify(n) = 0" +by (rule not_znegative_imp_zero, auto) + +lemma nat_of_zminus_int_of: "nat_of($- $# n) = 0" +by (simp add: nat_of_def int_of_def raw_nat_of zminus image_intrel_int) + +lemma int_of_nat_of: "#0 $<= z ==> $# nat_of(z) = intify(z)" +apply (rule not_zneg_nat_of_intify) +apply (simp add: znegative_iff_zless_0 not_zless_iff_zle) +done + +declare int_of_nat_of [simp] nat_of_zminus_int_of [simp] + +lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $<= z then intify(z) else #0)" +by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless) + +lemma zless_nat_iff_int_zless: "[| m: nat; z: int |] ==> (m < nat_of(z)) <-> ($#m $< z)" +apply (case_tac "znegative (z) ") +apply (erule_tac [2] not_zneg_nat_of [THEN subst]) +apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans] + simp add: znegative_iff_zless_0) +done + + +(** nat_of and zless **) + +(*An alternative condition is $#0 <= w *) +lemma zless_nat_conj_lemma: "$#0 $< z ==> (nat_of(w) < nat_of(z)) <-> (w $< z)" +apply (rule iff_trans) +apply (rule zless_int_of [THEN iff_sym]) +apply (auto simp add: int_of_nat_of_if simp del: zless_int_of) +apply (auto elim: zless_asym simp add: not_zle_iff_zless) +apply (blast intro: zless_zle_trans) +done + +lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) <-> ($#0 $< z & w $< z)" +apply (case_tac "$#0 $< z") +apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle) +done + +(*This simprule cannot be added unless we can find a way to make eq_integ_of_eq + unconditional! + [The condition "True" is a hack to prevent looping. + Conditional rewrite rules are tried after unconditional ones, so a rule + like eq_nat_number_of will be tried first to eliminate #mm=#nn.] + lemma integ_of_reorient [simp]: + "True ==> (integ_of(w) = x) <-> (x = integ_of(w))" + by auto +*) + +lemma integ_of_minus_reorient [simp]: + "(integ_of(w) = $- x) <-> ($- x = integ_of(w))" +by auto + +lemma integ_of_add_reorient [simp]: + "(integ_of(w) = x $+ y) <-> (x $+ y = integ_of(w))" +by auto + +lemma integ_of_diff_reorient [simp]: + "(integ_of(w) = x $- y) <-> (x $- y = integ_of(w))" +by auto + +lemma integ_of_mult_reorient [simp]: + "(integ_of(w) = x $* y) <-> (x $* y = integ_of(w))" +by auto + +ML +{* +val bin_pred_Pls = thm "bin_pred_Pls"; +val bin_pred_Min = thm "bin_pred_Min"; +val bin_minus_Pls = thm "bin_minus_Pls"; +val bin_minus_Min = thm "bin_minus_Min"; + +val NCons_Pls_0 = thm "NCons_Pls_0"; +val NCons_Pls_1 = thm "NCons_Pls_1"; +val NCons_Min_0 = thm "NCons_Min_0"; +val NCons_Min_1 = thm "NCons_Min_1"; +val NCons_BIT = thm "NCons_BIT"; +val NCons_simps = thms "NCons_simps"; +val integ_of_type = thm "integ_of_type"; +val NCons_type = thm "NCons_type"; +val bin_succ_type = thm "bin_succ_type"; +val bin_pred_type = thm "bin_pred_type"; +val bin_minus_type = thm "bin_minus_type"; +val bin_add_type = thm "bin_add_type"; +val bin_mult_type = thm "bin_mult_type"; +val integ_of_NCons = thm "integ_of_NCons"; +val integ_of_succ = thm "integ_of_succ"; +val integ_of_pred = thm "integ_of_pred"; +val integ_of_minus = thm "integ_of_minus"; +val bin_add_Pls = thm "bin_add_Pls"; +val bin_add_Pls_right = thm "bin_add_Pls_right"; +val bin_add_Min = thm "bin_add_Min"; +val bin_add_Min_right = thm "bin_add_Min_right"; +val bin_add_BIT_Pls = thm "bin_add_BIT_Pls"; +val bin_add_BIT_Min = thm "bin_add_BIT_Min"; +val bin_add_BIT_BIT = thm "bin_add_BIT_BIT"; +val integ_of_add = thm "integ_of_add"; +val diff_integ_of_eq = thm "diff_integ_of_eq"; +val integ_of_mult = thm "integ_of_mult"; +val bin_succ_1 = thm "bin_succ_1"; +val bin_succ_0 = thm "bin_succ_0"; +val bin_pred_1 = thm "bin_pred_1"; +val bin_pred_0 = thm "bin_pred_0"; +val bin_minus_1 = thm "bin_minus_1"; +val bin_minus_0 = thm "bin_minus_0"; +val bin_add_BIT_11 = thm "bin_add_BIT_11"; +val bin_add_BIT_10 = thm "bin_add_BIT_10"; +val bin_add_BIT_0 = thm "bin_add_BIT_0"; +val bin_mult_1 = thm "bin_mult_1"; +val bin_mult_0 = thm "bin_mult_0"; +val int_of_0 = thm "int_of_0"; +val int_of_succ = thm "int_of_succ"; +val zminus_0 = thm "zminus_0"; +val zadd_0_intify = thm "zadd_0_intify"; +val zadd_0_right_intify = thm "zadd_0_right_intify"; +val zmult_1_intify = thm "zmult_1_intify"; +val zmult_1_right_intify = thm "zmult_1_right_intify"; +val zmult_0 = thm "zmult_0"; +val zmult_0_right = thm "zmult_0_right"; +val zmult_minus1 = thm "zmult_minus1"; +val zmult_minus1_right = thm "zmult_minus1_right"; +val eq_integ_of_eq = thm "eq_integ_of_eq"; +val iszero_integ_of_Pls = thm "iszero_integ_of_Pls"; +val nonzero_integ_of_Min = thm "nonzero_integ_of_Min"; +val iszero_integ_of_BIT = thm "iszero_integ_of_BIT"; +val iszero_integ_of_0 = thm "iszero_integ_of_0"; +val iszero_integ_of_1 = thm "iszero_integ_of_1"; +val less_integ_of_eq_neg = thm "less_integ_of_eq_neg"; +val not_neg_integ_of_Pls = thm "not_neg_integ_of_Pls"; +val neg_integ_of_Min = thm "neg_integ_of_Min"; +val neg_integ_of_BIT = thm "neg_integ_of_BIT"; +val le_integ_of_eq_not_less = thm "le_integ_of_eq_not_less"; +val bin_arith_extra_simps = thms "bin_arith_extra_simps"; +val bin_arith_simps = thms "bin_arith_simps"; +val bin_rel_simps = thms "bin_rel_simps"; +val add_integ_of_left = thm "add_integ_of_left"; +val mult_integ_of_left = thm "mult_integ_of_left"; +val add_integ_of_diff1 = thm "add_integ_of_diff1"; +val add_integ_of_diff2 = thm "add_integ_of_diff2"; +val zdiff0 = thm "zdiff0"; +val zdiff0_right = thm "zdiff0_right"; +val zdiff_self = thm "zdiff_self"; +val znegative_iff_zless_0 = thm "znegative_iff_zless_0"; +val zero_zless_imp_znegative_zminus = thm "zero_zless_imp_znegative_zminus"; +val zero_zle_int_of = thm "zero_zle_int_of"; +val nat_of_0 = thm "nat_of_0"; +val nat_le_int0 = thm "nat_le_int0"; +val int_of_eq_0_imp_natify_eq_0 = thm "int_of_eq_0_imp_natify_eq_0"; +val nat_of_zminus_int_of = thm "nat_of_zminus_int_of"; +val int_of_nat_of = thm "int_of_nat_of"; +val int_of_nat_of_if = thm "int_of_nat_of_if"; +val zless_nat_iff_int_zless = thm "zless_nat_iff_int_zless"; +val zless_nat_conj = thm "zless_nat_conj"; +val integ_of_minus_reorient = thm "integ_of_minus_reorient"; +val integ_of_add_reorient = thm "integ_of_add_reorient"; +val integ_of_diff_reorient = thm "integ_of_diff_reorient"; +val integ_of_mult_reorient = thm "integ_of_mult_reorient"; +*} + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/EquivClass.thy Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,265 @@ +(* Title: ZF/EquivClass.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +*) + +header{*Equivalence Relations*} + +theory EquivClass imports Trancl Perm begin + +constdefs + + quotient :: "[i,i]=>i" (infixl "'/'/" 90) (*set of equiv classes*) + "A//r == {r``{x} . x:A}" + + congruent :: "[i,i=>i]=>o" + "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)" + + congruent2 :: "[i,i,[i,i]=>i]=>o" + "congruent2(r1,r2,b) == ALL y1 z1 y2 z2. + <y1,z1>:r1 --> <y2,z2>:r2 --> b(y1,y2) = b(z1,z2)" + +syntax + RESPECTS ::"[i=>i, i] => o" (infixr "respects" 80) + RESPECTS2 ::"[i=>i, i] => o" (infixr "respects2 " 80) + --{*Abbreviation for the common case where the relations are identical*} + +translations + "f respects r" == "congruent(r,f)" + "f respects2 r" => "congruent2(r,r,f)" + +subsection{*Suppes, Theorem 70: + @{term r} is an equiv relation iff @{term "converse(r) O r = r"}*} + +(** first half: equiv(A,r) ==> converse(r) O r = r **) + +lemma sym_trans_comp_subset: + "[| sym(r); trans(r) |] ==> converse(r) O r <= r" +by (unfold trans_def sym_def, blast) + +lemma refl_comp_subset: + "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r" +by (unfold refl_def, blast) + +lemma equiv_comp_eq: + "equiv(A,r) ==> converse(r) O r = r" +apply (unfold equiv_def) +apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset) +done + +(*second half*) +lemma comp_equivI: + "[| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)" +apply (unfold equiv_def refl_def sym_def trans_def) +apply (erule equalityE) +apply (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r", blast+) +done + +(** Equivalence classes **) + +(*Lemma for the next result*) +lemma equiv_class_subset: + "[| sym(r); trans(r); <a,b>: r |] ==> r``{a} <= r``{b}" +by (unfold trans_def sym_def, blast) + +lemma equiv_class_eq: + "[| equiv(A,r); <a,b>: r |] ==> r``{a} = r``{b}" +apply (unfold equiv_def) +apply (safe del: subsetI intro!: equalityI equiv_class_subset) +apply (unfold sym_def, blast) +done + +lemma equiv_class_self: + "[| equiv(A,r); a: A |] ==> a: r``{a}" +by (unfold equiv_def refl_def, blast) + +(*Lemma for the next result*) +lemma subset_equiv_class: + "[| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> <a,b>: r" +by (unfold equiv_def refl_def, blast) + +lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> <a,b>: r" +by (assumption | rule equalityD2 subset_equiv_class)+ + +(*thus r``{a} = r``{b} as well*) +lemma equiv_class_nondisjoint: + "[| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> <a,b>: r" +by (unfold equiv_def trans_def sym_def, blast) + +lemma equiv_type: "equiv(A,r) ==> r <= A*A" +by (unfold equiv_def, blast) + +lemma equiv_class_eq_iff: + "equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A" +by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) + +lemma eq_equiv_class_iff: + "[| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> <x,y>: r" +by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) + +(*** Quotients ***) + +(** Introduction/elimination rules -- needed? **) + +lemma quotientI [TC]: "x:A ==> r``{x}: A//r" +apply (unfold quotient_def) +apply (erule RepFunI) +done + +lemma quotientE: + "[| X: A//r; !!x. [| X = r``{x}; x:A |] ==> P |] ==> P" +by (unfold quotient_def, blast) + +lemma Union_quotient: + "equiv(A,r) ==> Union(A//r) = A" +by (unfold equiv_def refl_def quotient_def, blast) + +lemma quotient_disj: + "[| equiv(A,r); X: A//r; Y: A//r |] ==> X=Y | (X Int Y <= 0)" +apply (unfold quotient_def) +apply (safe intro!: equiv_class_eq, assumption) +apply (unfold equiv_def trans_def sym_def, blast) +done + +subsection{*Defining Unary Operations upon Equivalence Classes*} + +(** Could have a locale with the premises equiv(A,r) and congruent(r,b) +**) + +(*Conversion rule*) +lemma UN_equiv_class: + "[| equiv(A,r); b respects r; a: A |] ==> (UN x:r``{a}. b(x)) = b(a)" +apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)") + apply simp + apply (blast intro: equiv_class_self) +apply (unfold equiv_def sym_def congruent_def, blast) +done + +(*type checking of UN x:r``{a}. b(x) *) +lemma UN_equiv_class_type: + "[| equiv(A,r); b respects r; X: A//r; !!x. x : A ==> b(x) : B |] + ==> (UN x:X. b(x)) : B" +apply (unfold quotient_def, safe) +apply (simp (no_asm_simp) add: UN_equiv_class) +done + +(*Sufficient conditions for injectiveness. Could weaken premises! + major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B +*) +lemma UN_equiv_class_inject: + "[| equiv(A,r); b respects r; + (UN x:X. b(x))=(UN y:Y. b(y)); X: A//r; Y: A//r; + !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |] + ==> X=Y" +apply (unfold quotient_def, safe) +apply (rule equiv_class_eq, assumption) +apply (simp add: UN_equiv_class [of A r b]) +done + + +subsection{*Defining Binary Operations upon Equivalence Classes*} + +lemma congruent2_implies_congruent: + "[| equiv(A,r1); congruent2(r1,r2,b); a: A |] ==> congruent(r2,b(a))" +by (unfold congruent_def congruent2_def equiv_def refl_def, blast) + +lemma congruent2_implies_congruent_UN: + "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a: A2 |] ==> + congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))" +apply (unfold congruent_def, safe) +apply (frule equiv_type [THEN subsetD], assumption) +apply clarify +apply (simp add: UN_equiv_class congruent2_implies_congruent) +apply (unfold congruent2_def equiv_def refl_def, blast) +done + +lemma UN_equiv_class2: + "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2 |] + ==> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)" +by (simp add: UN_equiv_class congruent2_implies_congruent + congruent2_implies_congruent_UN) + +(*type checking*) +lemma UN_equiv_class_type2: + "[| equiv(A,r); b respects2 r; + X1: A//r; X2: A//r; + !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B + |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B" +apply (unfold quotient_def, safe) +apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN + congruent2_implies_congruent quotientI) +done + + +(*Suggested by John Harrison -- the two subproofs may be MUCH simpler + than the direct proof*) +lemma congruent2I: + "[| equiv(A1,r1); equiv(A2,r2); + !! y z w. [| w \<in> A2; <y,z> \<in> r1 |] ==> b(y,w) = b(z,w); + !! y z w. [| w \<in> A1; <y,z> \<in> r2 |] ==> b(w,y) = b(w,z) + |] ==> congruent2(r1,r2,b)" +apply (unfold congruent2_def equiv_def refl_def, safe) +apply (blast intro: trans) +done + +lemma congruent2_commuteI: + assumes equivA: "equiv(A,r)" + and commute: "!! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y)" + and congt: "!! y z w. [| w: A; <y,z>: r |] ==> b(w,y) = b(w,z)" + shows "b respects2 r" +apply (insert equivA [THEN equiv_type, THEN subsetD]) +apply (rule congruent2I [OF equivA equivA]) +apply (rule commute [THEN trans]) +apply (rule_tac [3] commute [THEN trans, symmetric]) +apply (rule_tac [5] sym) +apply (blast intro: congt)+ +done + +(*Obsolete?*) +lemma congruent_commuteI: + "[| equiv(A,r); Z: A//r; + !!w. [| w: A |] ==> congruent(r, %z. b(w,z)); + !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) + |] ==> congruent(r, %w. UN z: Z. b(w,z))" +apply (simp (no_asm) add: congruent_def) +apply (safe elim!: quotientE) +apply (frule equiv_type [THEN subsetD], assumption) +apply (simp add: UN_equiv_class [of A r]) +apply (simp add: congruent_def) +done + +ML +{* +val sym_trans_comp_subset = thm "sym_trans_comp_subset"; +val refl_comp_subset = thm "refl_comp_subset"; +val equiv_comp_eq = thm "equiv_comp_eq"; +val comp_equivI = thm "comp_equivI"; +val equiv_class_subset = thm "equiv_class_subset"; +val equiv_class_eq = thm "equiv_class_eq"; +val equiv_class_self = thm "equiv_class_self"; +val subset_equiv_class = thm "subset_equiv_class"; +val eq_equiv_class = thm "eq_equiv_class"; +val equiv_class_nondisjoint = thm "equiv_class_nondisjoint"; +val equiv_type = thm "equiv_type"; +val equiv_class_eq_iff = thm "equiv_class_eq_iff"; +val eq_equiv_class_iff = thm "eq_equiv_class_iff"; +val quotientI = thm "quotientI"; +val quotientE = thm "quotientE"; +val Union_quotient = thm "Union_quotient"; +val quotient_disj = thm "quotient_disj"; +val UN_equiv_class = thm "UN_equiv_class"; +val UN_equiv_class_type = thm "UN_equiv_class_type"; +val UN_equiv_class_inject = thm "UN_equiv_class_inject"; +val congruent2_implies_congruent = thm "congruent2_implies_congruent"; +val congruent2_implies_congruent_UN = thm "congruent2_implies_congruent_UN"; +val congruent_commuteI = thm "congruent_commuteI"; +val UN_equiv_class2 = thm "UN_equiv_class2"; +val UN_equiv_class_type2 = thm "UN_equiv_class_type2"; +val congruent2I = thm "congruent2I"; +val congruent2_commuteI = thm "congruent2_commuteI"; +val congruent_commuteI = thm "congruent_commuteI"; +*} + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Int.thy Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,1057 @@ +(* Title: ZF/Int.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +*) + +header{*The Integers as Equivalence Classes Over Pairs of Natural Numbers*} + +theory Int imports EquivClass ArithSimp begin + +constdefs + intrel :: i + "intrel == {p : (nat*nat)*(nat*nat). + \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}" + + int :: i + "int == (nat*nat)//intrel" + + int_of :: "i=>i" --{*coercion from nat to int*} ("$# _" [80] 80) + "$# m == intrel `` {<natify(m), 0>}" + + intify :: "i=>i" --{*coercion from ANYTHING to int*} + "intify(m) == if m : int then m else $#0" + + raw_zminus :: "i=>i" + "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}" + + zminus :: "i=>i" ("$- _" [80] 80) + "$- z == raw_zminus (intify(z))" + + znegative :: "i=>o" + "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z" + + iszero :: "i=>o" + "iszero(z) == z = $# 0" + + raw_nat_of :: "i=>i" + "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)" + + nat_of :: "i=>i" + "nat_of(z) == raw_nat_of (intify(z))" + + zmagnitude :: "i=>i" + --{*could be replaced by an absolute value function from int to int?*} + "zmagnitude(z) == + THE m. m\<in>nat & ((~ znegative(z) & z = $# m) | + (znegative(z) & $- z = $# m))" + + raw_zmult :: "[i,i]=>i" + (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2. + Perhaps a "curried" or even polymorphic congruent predicate would be + better.*) + "raw_zmult(z1,z2) == + \<Union>p1\<in>z1. \<Union>p2\<in>z2. split(%x1 y1. split(%x2 y2. + intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)" + + zmult :: "[i,i]=>i" (infixl "$*" 70) + "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" + + raw_zadd :: "[i,i]=>i" + "raw_zadd (z1, z2) == + \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2 + in intrel``{<x1#+x2, y1#+y2>}" + + zadd :: "[i,i]=>i" (infixl "$+" 65) + "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" + + zdiff :: "[i,i]=>i" (infixl "$-" 65) + "z1 $- z2 == z1 $+ zminus(z2)" + + zless :: "[i,i]=>o" (infixl "$<" 50) + "z1 $< z2 == znegative(z1 $- z2)" + + zle :: "[i,i]=>o" (infixl "$<=" 50) + "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)" + + +syntax (xsymbols) + zmult :: "[i,i]=>i" (infixl "$\<times>" 70) + zle :: "[i,i]=>o" (infixl "$\<le>" 50) --{*less than or equals*} + +syntax (HTML output) + zmult :: "[i,i]=>i" (infixl "$\<times>" 70) + zle :: "[i,i]=>o" (infixl "$\<le>" 50) + + +declare quotientE [elim!] + +subsection{*Proving that @{term intrel} is an equivalence relation*} + +(** Natural deduction for intrel **) + +lemma intrel_iff [simp]: + "<<x1,y1>,<x2,y2>>: intrel <-> + x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1" +by (simp add: intrel_def) + +lemma intrelI [intro!]: + "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> <<x1,y1>,<x2,y2>>: intrel" +by (simp add: intrel_def) + +lemma intrelE [elim!]: + "[| p: intrel; + !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1; + x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] ==> Q |] + ==> Q" +by (simp add: intrel_def, blast) + +lemma int_trans_lemma: + "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" +apply (rule sym) +apply (erule add_left_cancel)+ +apply (simp_all (no_asm_simp)) +done + +lemma equiv_intrel: "equiv(nat*nat, intrel)" +apply (simp add: equiv_def refl_def sym_def trans_def) +apply (fast elim!: sym int_trans_lemma) +done + +lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : int" +by (simp add: int_def) + +declare equiv_intrel [THEN eq_equiv_class_iff, simp] +declare conj_cong [cong] + +lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel] + +(** int_of: the injection from nat to int **) + +lemma int_of_type [simp,TC]: "$#m : int" +by (simp add: int_def quotient_def int_of_def, auto) + +lemma int_of_eq [iff]: "($# m = $# n) <-> natify(m)=natify(n)" +by (simp add: int_of_def) + +lemma int_of_inject: "[| $#m = $#n; m\<in>nat; n\<in>nat |] ==> m=n" +by (drule int_of_eq [THEN iffD1], auto) + + +(** intify: coercion from anything to int **) + +lemma intify_in_int [iff,TC]: "intify(x) : int" +by (simp add: intify_def) + +lemma intify_ident [simp]: "n : int ==> intify(n) = n" +by (simp add: intify_def) + + +subsection{*Collapsing rules: to remove @{term intify} + from arithmetic expressions*} + +lemma intify_idem [simp]: "intify(intify(x)) = intify(x)" +by simp + +lemma int_of_natify [simp]: "$# (natify(m)) = $# m" +by (simp add: int_of_def) + +lemma zminus_intify [simp]: "$- (intify(m)) = $- m" +by (simp add: zminus_def) + +(** Addition **) + +lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y" +by (simp add: zadd_def) + +lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y" +by (simp add: zadd_def) + +(** Subtraction **) + +lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y" +by (simp add: zdiff_def) + +lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y" +by (simp add: zdiff_def) + +(** Multiplication **) + +lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y" +by (simp add: zmult_def) + +lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y" +by (simp add: zmult_def) + +(** Orderings **) + +lemma zless_intify1 [simp]:"intify(x) $< y <-> x $< y" +by (simp add: zless_def) + +lemma zless_intify2 [simp]:"x $< intify(y) <-> x $< y" +by (simp add: zless_def) + +lemma zle_intify1 [simp]:"intify(x) $<= y <-> x $<= y" +by (simp add: zle_def) + +lemma zle_intify2 [simp]:"x $<= intify(y) <-> x $<= y" +by (simp add: zle_def) + + +subsection{*@{term zminus}: unary negation on @{term int}*} + +lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel" +by (auto simp add: congruent_def add_ac) + +lemma raw_zminus_type: "z : int ==> raw_zminus(z) : int" +apply (simp add: int_def raw_zminus_def) +apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent]) +done + +lemma zminus_type [TC,iff]: "$-z : int" +by (simp add: zminus_def raw_zminus_type) + +lemma raw_zminus_inject: + "[| raw_zminus(z) = raw_zminus(w); z: int; w: int |] ==> z=w" +apply (simp add: int_def raw_zminus_def) +apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) +apply (auto dest: eq_intrelD simp add: add_ac) +done + +lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)" +apply (simp add: zminus_def) +apply (blast dest!: raw_zminus_inject) +done + +lemma zminus_inject: "[| $-z = $-w; z: int; w: int |] ==> z=w" +by auto + +lemma raw_zminus: + "[| x\<in>nat; y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}" +apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) +done + +lemma zminus: + "[| x\<in>nat; y\<in>nat |] + ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}" +by (simp add: zminus_def raw_zminus image_intrel_int) + +lemma raw_zminus_zminus: "z : int ==> raw_zminus (raw_zminus(z)) = z" +by (auto simp add: int_def raw_zminus) + +lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)" +by (simp add: zminus_def raw_zminus_type raw_zminus_zminus) + +lemma zminus_int0 [simp]: "$- ($#0) = $#0" +by (simp add: int_of_def zminus) + +lemma zminus_zminus: "z : int ==> $- ($- z) = z" +by simp + + +subsection{*@{term znegative}: the test for negative integers*} + +lemma znegative: "[| x\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y" +apply (cases "x<y") +apply (auto simp add: znegative_def not_lt_iff_le) +apply (subgoal_tac "y #+ x2 < x #+ y2", force) +apply (rule add_le_lt_mono, auto) +done + +(*No natural number is negative!*) +lemma not_znegative_int_of [iff]: "~ znegative($# n)" +by (simp add: znegative int_of_def) + +lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))" +by (simp add: znegative int_of_def zminus natify_succ) + +lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0" +by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym]) + + +subsection{*@{term nat_of}: Coercion of an Integer to a Natural Number*} + +lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)" +by (simp add: nat_of_def) + +lemma nat_of_congruent: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel" +by (auto simp add: congruent_def split add: nat_diff_split) + +lemma raw_nat_of: + "[| x\<in>nat; y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = x#-y" +by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) + +lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" +by (simp add: int_of_def raw_nat_of) + +lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)" +by (simp add: raw_nat_of_int_of nat_of_def) + +lemma raw_nat_of_type: "raw_nat_of(z) \<in> nat" +by (simp add: raw_nat_of_def) + +lemma nat_of_type [iff,TC]: "nat_of(z) \<in> nat" +by (simp add: nat_of_def raw_nat_of_type) + +subsection{*zmagnitude: magnitide of an integer, as a natural number*} + +lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)" +by (auto simp add: zmagnitude_def int_of_eq) + +lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n" +apply (drule sym) +apply (simp (no_asm_simp) add: int_of_eq) +done + +lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)" +apply (simp add: zmagnitude_def) +apply (rule the_equality) +apply (auto dest!: not_znegative_imp_zero natify_int_of_eq + iff del: int_of_eq, auto) +done + +lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\<in>nat" +apply (simp add: zmagnitude_def) +apply (rule theI2, auto) +done + +lemma not_zneg_int_of: + "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>nat. z = $# n" +apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) +apply (rename_tac x y) +apply (rule_tac x="x#-y" in bexI) +apply (auto simp add: add_diff_inverse2) +done + +lemma not_zneg_mag [simp]: + "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" +by (drule not_zneg_int_of, auto) + +lemma zneg_int_of: + "[| znegative(z); z: int |] ==> \<exists>n\<in>nat. z = $- ($# succ(n))" +by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) + +lemma zneg_mag [simp]: + "[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z" +by (drule zneg_int_of, auto) + +lemma int_cases: "z : int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))" +apply (case_tac "znegative (z) ") +prefer 2 apply (blast dest: not_zneg_mag sym) +apply (blast dest: zneg_int_of) +done + +lemma not_zneg_raw_nat_of: + "[| ~ znegative(z); z: int |] ==> $# (raw_nat_of(z)) = z" +apply (drule not_zneg_int_of) +apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) +done + +lemma not_zneg_nat_of_intify: + "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" +by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) + +lemma not_zneg_nat_of: "[| ~ znegative(z); z: int |] ==> $# (nat_of(z)) = z" +apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) +done + +lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" +apply (subgoal_tac "intify(z) \<in> int") +apply (simp add: int_def) +apply (auto simp add: znegative nat_of_def raw_nat_of + split add: nat_diff_split) +done + + +subsection{*@{term zadd}: addition on int*} + +text{*Congruence Property for Addition*} +lemma zadd_congruent2: + "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2 + in intrel``{<x1#+x2, y1#+y2>}) + respects2 intrel" +apply (simp add: congruent2_def) +(*Proof via congruent2_commuteI seems longer*) +apply safe +apply (simp (no_asm_simp) add: add_assoc Let_def) +(*The rest should be trivial, but rearranging terms is hard + add_ac does not help rewriting with the assumptions.*) +apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst]) +apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst]) +apply (simp (no_asm_simp) add: add_assoc [symmetric]) +done + +lemma raw_zadd_type: "[| z: int; w: int |] ==> raw_zadd(z,w) : int" +apply (simp add: int_def raw_zadd_def) +apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) +apply (simp add: Let_def) +done + +lemma zadd_type [iff,TC]: "z $+ w : int" +by (simp add: zadd_def raw_zadd_type) + +lemma raw_zadd: + "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) = + intrel `` {<x1#+x2, y1#+y2>}" +apply (simp add: raw_zadd_def + UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) +apply (simp add: Let_def) +done + +lemma zadd: + "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = + intrel `` {<x1#+x2, y1#+y2>}" +by (simp add: zadd_def raw_zadd image_intrel_int) + +lemma raw_zadd_int0: "z : int ==> raw_zadd ($#0,z) = z" +by (auto simp add: int_def int_of_def raw_zadd) + +lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" +by (simp add: zadd_def raw_zadd_int0) + +lemma zadd_int0: "z: int ==> $#0 $+ z = z" +by simp + +lemma raw_zminus_zadd_distrib: + "[| z: int; w: int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" +by (auto simp add: zminus raw_zadd int_def) + +lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" +by (simp add: zadd_def raw_zminus_zadd_distrib) + +lemma raw_zadd_commute: + "[| z: int; w: int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" +by (auto simp add: raw_zadd add_ac int_def) + +lemma zadd_commute: "z $+ w = w $+ z" +by (simp add: zadd_def raw_zadd_commute) + +lemma raw_zadd_assoc: + "[| z1: int; z2: int; z3: int |] + ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" +by (auto simp add: int_def raw_zadd add_assoc) + +lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)" +by (simp add: zadd_def raw_zadd_type raw_zadd_assoc) + +(*For AC rewriting*) +lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)" +apply (simp add: zadd_assoc [symmetric]) +apply (simp add: zadd_commute) +done + +(*Integer addition is an AC operator*) +lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute + +lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)" +by (simp add: int_of_def zadd) + +lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)" +by (simp add: int_of_add [symmetric] natify_succ) + +lemma int_of_diff: + "[| m\<in>nat; n le m |] ==> $# (m #- n) = ($#m) $- ($#n)" +apply (simp add: int_of_def zdiff_def) +apply (frule lt_nat_in_nat) +apply (simp_all add: zadd zminus add_diff_inverse2) +done + +lemma raw_zadd_zminus_inverse: "z : int ==> raw_zadd (z, $- z) = $#0" +by (auto simp add: int_def int_of_def zminus raw_zadd add_commute) + +lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0" +apply (simp add: zadd_def) +apply (subst zminus_intify [symmetric]) +apply (rule intify_in_int [THEN raw_zadd_zminus_inverse]) +done + +lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0" +by (simp add: zadd_commute zadd_zminus_inverse) + +lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" +by (rule trans [OF zadd_commute zadd_int0_intify]) + +lemma zadd_int0_right: "z:int ==> z $+ $#0 = z" +by simp + + +subsection{*@{term zmult}: Integer Multiplication*} + +text{*Congruence property for multiplication*} +lemma zmult_congruent2: + "(%p1 p2. split(%x1 y1. split(%x2 y2. + intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)) + respects2 intrel" +apply (rule equiv_intrel [THEN congruent2_commuteI], auto) +(*Proof that zmult is congruent in one argument*) +apply (rename_tac x y) +apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context]) +apply (drule_tac t = "%u. y#*u" in subst_context) +apply (erule add_left_cancel)+ +apply (simp_all add: add_mult_distrib_left) +done + + +lemma raw_zmult_type: "[| z: int; w: int |] ==> raw_zmult(z,w) : int" +apply (simp add: int_def raw_zmult_def) +apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) +apply (simp add: Let_def) +done + +lemma zmult_type [iff,TC]: "z $* w : int" +by (simp add: zmult_def raw_zmult_type) + +lemma raw_zmult: + "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) = + intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" +by (simp add: raw_zmult_def + UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) + +lemma zmult: + "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] + ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = + intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" +by (simp add: zmult_def raw_zmult image_intrel_int) + +lemma raw_zmult_int0: "z : int ==> raw_zmult ($#0,z) = $#0" +by (auto simp add: int_def int_of_def raw_zmult) + +lemma zmult_int0 [simp]: "$#0 $* z = $#0" +by (simp add: zmult_def raw_zmult_int0) + +lemma raw_zmult_int1: "z : int ==> raw_zmult ($#1,z) = z" +by (auto simp add: int_def int_of_def raw_zmult) + +lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)" +by (simp add: zmult_def raw_zmult_int1) + +lemma zmult_int1: "z : int ==> $#1 $* z = z" +by simp + +lemma raw_zmult_commute: + "[| z: int; w: int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" +by (auto simp add: int_def raw_zmult add_ac mult_ac) + +lemma zmult_commute: "z $* w = w $* z" +by (simp add: zmult_def raw_zmult_commute) + +lemma raw_zmult_zminus: + "[| z: int; w: int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" +by (auto simp add: int_def zminus raw_zmult add_ac) + +lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" +apply (simp add: zmult_def raw_zmult_zminus) +apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto) +done + +lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)" +by (simp add: zmult_commute [of w]) + +lemma raw_zmult_assoc: + "[| z1: int; z2: int; z3: int |] + ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" +by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) + +lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)" +by (simp add: zmult_def raw_zmult_type raw_zmult_assoc) + +(*For AC rewriting*) +lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)" +apply (simp add: zmult_assoc [symmetric]) +apply (simp add: zmult_commute) +done + +(*Integer multiplication is an AC operator*) +lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute + +lemma raw_zadd_zmult_distrib: + "[| z1: int; z2: int; w: int |] + ==> raw_zmult(raw_zadd(z1,z2), w) = + raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" +by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) + +lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)" +by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type + raw_zadd_zmult_distrib) + +lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)" +by (simp add: zmult_commute [of w] zadd_zmult_distrib) + +lemmas int_typechecks = + int_of_type zminus_type zmagnitude_type zadd_type zmult_type + + +(*** Subtraction laws ***) + +lemma zdiff_type [iff,TC]: "z $- w : int" +by (simp add: zdiff_def) + +lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z" +by (simp add: zdiff_def zadd_commute) + +lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)" +apply (simp add: zdiff_def) +apply (subst zadd_zmult_distrib) +apply (simp add: zmult_zminus) +done + +lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)" +by (simp add: zmult_commute [of w] zdiff_zmult_distrib) + +lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y" +by (simp add: zdiff_def zadd_ac) + + +subsection{*The "Less Than" Relation*} + +(*"Less than" is a linear ordering*) +lemma zless_linear_lemma: + "[| z: int; w: int |] ==> z$<w | z=w | w$<z" +apply (simp add: int_def zless_def znegative_def zdiff_def, auto) +apply (simp add: zadd zminus image_iff Bex_def) +apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt) +apply (force dest!: spec simp add: add_ac)+ +done + +lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z" +apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma) +apply auto +done + +lemma zless_not_refl [iff]: "~ (z$<z)" +by (auto simp add: zless_def znegative_def int_of_def zdiff_def) + +lemma neq_iff_zless: "[| x: int; y: int |] ==> (x ~= y) <-> (x $< y | y $< x)" +by (cut_tac z = x and w = y in zless_linear, auto) + +lemma zless_imp_intify_neq: "w $< z ==> intify(w) ~= intify(z)" +apply auto +apply (subgoal_tac "~ (intify (w) $< intify (z))") +apply (erule_tac [2] ssubst) +apply (simp (no_asm_use)) +apply auto +done + +(*This lemma allows direct proofs of other <-properties*) +lemma zless_imp_succ_zadd_lemma: + "[| w $< z; w: int; z: int |] ==> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) +apply (rule_tac x = k in bexI) +apply (erule add_left_cancel, auto) +done + +lemma zless_imp_succ_zadd: + "w $< z ==> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))" +apply (subgoal_tac "intify (w) $< intify (z) ") +apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma) +apply auto +done + +lemma zless_succ_zadd_lemma: + "w : int ==> w $< w $+ $# succ(n)" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto simp add: zadd zminus int_of_def image_iff) +apply (rule_tac x = 0 in exI, auto) +done + +lemma zless_succ_zadd: "w $< w $+ $# succ(n)" +by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto) + +lemma zless_iff_succ_zadd: + "w $< z <-> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))" +apply (rule iffI) +apply (erule zless_imp_succ_zadd, auto) +apply (rename_tac "n") +apply (cut_tac w = w and n = n in zless_succ_zadd, auto) +done + +lemma zless_int_of [simp]: "[| m\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)" +apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric]) +apply (blast intro: sym) +done + +lemma zless_trans_lemma: + "[| x $< y; y $< z; x: int; y : int; z: int |] ==> x $< z" +apply (simp add: zless_def znegative_def zdiff_def int_def) +apply (auto simp add: zadd zminus image_iff) +apply (rename_tac x1 x2 y1 y2) +apply (rule_tac x = "x1#+x2" in exI) +apply (rule_tac x = "y1#+y2" in exI) +apply (auto simp add: add_lt_mono) +apply (rule sym) +apply (erule add_left_cancel)+ +apply auto +done + +lemma zless_trans: "[| x $< y; y $< z |] ==> x $< z" +apply (subgoal_tac "intify (x) $< intify (z) ") +apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma) +apply auto +done + +lemma zless_not_sym: "z $< w ==> ~ (w $< z)" +by (blast dest: zless_trans) + +(* [| z $< w; ~ P ==> w $< z |] ==> P *) +lemmas zless_asym = zless_not_sym [THEN swap, standard] + +lemma zless_imp_zle: "z $< w ==> z $<= w" +by (simp add: zle_def) + +lemma zle_linear: "z $<= w | w $<= z" +apply (simp add: zle_def) +apply (cut_tac zless_linear, blast) +done + + +subsection{*Less Than or Equals*} + +lemma zle_refl: "z $<= z" +by (simp add: zle_def) + +lemma zle_eq_refl: "x=y ==> x $<= y" +by (simp add: zle_refl) + +lemma zle_anti_sym_intify: "[| x $<= y; y $<= x |] ==> intify(x) = intify(y)" +apply (simp add: zle_def, auto) +apply (blast dest: zless_trans) +done + +lemma zle_anti_sym: "[| x $<= y; y $<= x; x: int; y: int |] ==> x=y" +by (drule zle_anti_sym_intify, auto) + +lemma zle_trans_lemma: + "[| x: int; y: int; z: int; x $<= y; y $<= z |] ==> x $<= z" +apply (simp add: zle_def, auto) +apply (blast intro: zless_trans) +done + +lemma zle_trans: "[| x $<= y; y $<= z |] ==> x $<= z" +apply (subgoal_tac "intify (x) $<= intify (z) ") +apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma) +apply auto +done + +lemma zle_zless_trans: "[| i $<= j; j $< k |] ==> i $< k" +apply (auto simp add: zle_def) +apply (blast intro: zless_trans) +apply (simp add: zless_def zdiff_def zadd_def) +done + +lemma zless_zle_trans: "[| i $< j; j $<= k |] ==> i $< k" +apply (auto simp add: zle_def) +apply (blast intro: zless_trans) +apply (simp add: zless_def zdiff_def zminus_def) +done + +lemma not_zless_iff_zle: "~ (z $< w) <-> (w $<= z)" +apply (cut_tac z = z and w = w in zless_linear) +apply (auto dest: zless_trans simp add: zle_def) +apply (auto dest!: zless_imp_intify_neq) +done + +lemma not_zle_iff_zless: "~ (z $<= w) <-> (w $< z)" +by (simp add: not_zless_iff_zle [THEN iff_sym]) + + +subsection{*More subtraction laws (for @{text zcompare_rls})*} + +lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y" +by (simp add: zdiff_def zadd_ac) + +lemma zdiff_zless_iff: "(x$-y $< z) <-> (x $< z $+ y)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zless_zdiff_iff: "(x $< z$-y) <-> (x $+ y $< z)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zdiff_eq_iff: "[| x: int; z: int |] ==> (x$-y = z) <-> (x = z $+ y)" +by (auto simp add: zdiff_def zadd_assoc) + +lemma eq_zdiff_iff: "[| x: int; z: int |] ==> (x = z$-y) <-> (x $+ y = z)" +by (auto simp add: zdiff_def zadd_assoc) + +lemma zdiff_zle_iff_lemma: + "[| x: int; z: int |] ==> (x$-y $<= z) <-> (x $<= z $+ y)" +by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) + +lemma zdiff_zle_iff: "(x$-y $<= z) <-> (x $<= z $+ y)" +by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) + +lemma zle_zdiff_iff_lemma: + "[| x: int; z: int |] ==>(x $<= z$-y) <-> (x $+ y $<= z)" +apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) +apply (auto simp add: zdiff_def zadd_assoc) +done + +lemma zle_zdiff_iff: "(x $<= z$-y) <-> (x $+ y $<= z)" +by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp) + +text{*This list of rewrites simplifies (in)equalities by bringing subtractions + to the top and then moving negative terms to the other side. + Use with @{text zadd_ac}*} +lemmas zcompare_rls = + zdiff_def [symmetric] + zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 + zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff + zdiff_eq_iff eq_zdiff_iff + + +subsection{*Monotonicity and Cancellation Results for Instantiation + of the CancelNumerals Simprocs*} + +lemma zadd_left_cancel: + "[| w: int; w': int |] ==> (z $+ w' = z $+ w) <-> (w' = w)" +apply safe +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (simp add: zadd_ac) +done + +lemma zadd_left_cancel_intify [simp]: + "(z $+ w' = z $+ w) <-> intify(w') = intify(w)" +apply (rule iff_trans) +apply (rule_tac [2] zadd_left_cancel, auto) +done + +lemma zadd_right_cancel: + "[| w: int; w': int |] ==> (w' $+ z = w $+ z) <-> (w' = w)" +apply safe +apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) +apply (simp add: zadd_ac) +done + +lemma zadd_right_cancel_intify [simp]: + "(w' $+ z = w $+ z) <-> intify(w') = intify(w)" +apply (rule iff_trans) +apply (rule_tac [2] zadd_right_cancel, auto) +done + +lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) <-> (w' $< w)" +by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc) + +lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) <-> (w' $< w)" +by (simp add: zadd_commute [of z] zadd_right_cancel_zless) + +lemma zadd_right_cancel_zle [simp]: "(w' $+ z $<= w $+ z) <-> w' $<= w" +by (simp add: zle_def) + +lemma zadd_left_cancel_zle [simp]: "(z $+ w' $<= z $+ w) <-> w' $<= w" +by (simp add: zadd_commute [of z] zadd_right_cancel_zle) + + +(*"v $<= w ==> v$+z $<= w$+z"*) +lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2, standard] + +(*"v $<= w ==> z$+v $<= z$+w"*) +lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard] + +(*"v $<= w ==> v$+z $<= w$+z"*) +lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard] + +(*"v $<= w ==> z$+v $<= z$+w"*) +lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard] + +lemma zadd_zle_mono: "[| w' $<= w; z' $<= z |] ==> w' $+ z' $<= w $+ z" +by (erule zadd_zle_mono1 [THEN zle_trans], simp) + +lemma zadd_zless_mono: "[| w' $< w; z' $<= z |] ==> w' $+ z' $< w $+ z" +by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp) + + +subsection{*Comparison laws*} + +lemma zminus_zless_zminus [simp]: "($- x $< $- y) <-> (y $< x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zminus_zle_zminus [simp]: "($- x $<= $- y) <-> (y $<= x)" +by (simp add: not_zless_iff_zle [THEN iff_sym]) + +subsubsection{*More inequality lemmas*} + +lemma equation_zminus: "[| x: int; y: int |] ==> (x = $- y) <-> (y = $- x)" +by auto + +lemma zminus_equation: "[| x: int; y: int |] ==> ($- x = y) <-> ($- y = x)" +by auto + +lemma equation_zminus_intify: "(intify(x) = $- y) <-> (intify(y) = $- x)" +apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus) +apply auto +done + +lemma zminus_equation_intify: "($- x = intify(y)) <-> ($- y = intify(x))" +apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation) +apply auto +done + + +subsubsection{*The next several equations are permutative: watch out!*} + +lemma zless_zminus: "(x $< $- y) <-> (y $< $- x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zminus_zless: "($- x $< y) <-> ($- y $< x)" +by (simp add: zless_def zdiff_def zadd_ac) + +lemma zle_zminus: "(x $<= $- y) <-> (y $<= $- x)" +by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless) + +lemma zminus_zle: "($- x $<= y) <-> ($- y $<= x)" +by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus) + +ML +{* +val zdiff_def = thm "zdiff_def"; +val int_of_type = thm "int_of_type"; +val int_of_eq = thm "int_of_eq"; +val int_of_inject = thm "int_of_inject"; +val intify_in_int = thm "intify_in_int"; +val intify_ident = thm "intify_ident"; +val intify_idem = thm "intify_idem"; +val int_of_natify = thm "int_of_natify"; +val zminus_intify = thm "zminus_intify"; +val zadd_intify1 = thm "zadd_intify1"; +val zadd_intify2 = thm "zadd_intify2"; +val zdiff_intify1 = thm "zdiff_intify1"; +val zdiff_intify2 = thm "zdiff_intify2"; +val zmult_intify1 = thm "zmult_intify1"; +val zmult_intify2 = thm "zmult_intify2"; +val zless_intify1 = thm "zless_intify1"; +val zless_intify2 = thm "zless_intify2"; +val zle_intify1 = thm "zle_intify1"; +val zle_intify2 = thm "zle_intify2"; +val zminus_congruent = thm "zminus_congruent"; +val zminus_type = thm "zminus_type"; +val zminus_inject_intify = thm "zminus_inject_intify"; +val zminus_inject = thm "zminus_inject"; +val zminus = thm "zminus"; +val zminus_zminus_intify = thm "zminus_zminus_intify"; +val zminus_int0 = thm "zminus_int0"; +val zminus_zminus = thm "zminus_zminus"; +val not_znegative_int_of = thm "not_znegative_int_of"; +val znegative_zminus_int_of = thm "znegative_zminus_int_of"; +val not_znegative_imp_zero = thm "not_znegative_imp_zero"; +val nat_of_intify = thm "nat_of_intify"; +val nat_of_int_of = thm "nat_of_int_of"; +val nat_of_type = thm "nat_of_type"; +val zmagnitude_int_of = thm "zmagnitude_int_of"; +val natify_int_of_eq = thm "natify_int_of_eq"; +val zmagnitude_zminus_int_of = thm "zmagnitude_zminus_int_of"; +val zmagnitude_type = thm "zmagnitude_type"; +val not_zneg_int_of = thm "not_zneg_int_of"; +val not_zneg_mag = thm "not_zneg_mag"; +val zneg_int_of = thm "zneg_int_of"; +val zneg_mag = thm "zneg_mag"; +val int_cases = thm "int_cases"; +val not_zneg_nat_of_intify = thm "not_zneg_nat_of_intify"; +val not_zneg_nat_of = thm "not_zneg_nat_of"; +val zneg_nat_of = thm "zneg_nat_of"; +val zadd_congruent2 = thm "zadd_congruent2"; +val zadd_type = thm "zadd_type"; +val zadd = thm "zadd"; +val zadd_int0_intify = thm "zadd_int0_intify"; +val zadd_int0 = thm "zadd_int0"; +val zminus_zadd_distrib = thm "zminus_zadd_distrib"; +val zadd_commute = thm "zadd_commute"; +val zadd_assoc = thm "zadd_assoc"; +val zadd_left_commute = thm "zadd_left_commute"; +val zadd_ac = thms "zadd_ac"; +val int_of_add = thm "int_of_add"; +val int_succ_int_1 = thm "int_succ_int_1"; +val int_of_diff = thm "int_of_diff"; +val zadd_zminus_inverse = thm "zadd_zminus_inverse"; +val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2"; +val zadd_int0_right_intify = thm "zadd_int0_right_intify"; +val zadd_int0_right = thm "zadd_int0_right"; +val zmult_congruent2 = thm "zmult_congruent2"; +val zmult_type = thm "zmult_type"; +val zmult = thm "zmult"; +val zmult_int0 = thm "zmult_int0"; +val zmult_int1_intify = thm "zmult_int1_intify"; +val zmult_int1 = thm "zmult_int1"; +val zmult_commute = thm "zmult_commute"; +val zmult_zminus = thm "zmult_zminus"; +val zmult_zminus_right = thm "zmult_zminus_right"; +val zmult_assoc = thm "zmult_assoc"; +val zmult_left_commute = thm "zmult_left_commute"; +val zmult_ac = thms "zmult_ac"; +val zadd_zmult_distrib = thm "zadd_zmult_distrib"; +val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; +val int_typechecks = thms "int_typechecks"; +val zdiff_type = thm "zdiff_type"; +val zminus_zdiff_eq = thm "zminus_zdiff_eq"; +val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; +val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; +val zadd_zdiff_eq = thm "zadd_zdiff_eq"; +val zdiff_zadd_eq = thm "zdiff_zadd_eq"; +val zless_linear = thm "zless_linear"; +val zless_not_refl = thm "zless_not_refl"; +val neq_iff_zless = thm "neq_iff_zless"; +val zless_imp_intify_neq = thm "zless_imp_intify_neq"; +val zless_imp_succ_zadd = thm "zless_imp_succ_zadd"; +val zless_succ_zadd = thm "zless_succ_zadd"; +val zless_iff_succ_zadd = thm "zless_iff_succ_zadd"; +val zless_int_of = thm "zless_int_of"; +val zless_trans = thm "zless_trans"; +val zless_not_sym = thm "zless_not_sym"; +val zless_asym = thm "zless_asym"; +val zless_imp_zle = thm "zless_imp_zle"; +val zle_linear = thm "zle_linear"; +val zle_refl = thm "zle_refl"; +val zle_eq_refl = thm "zle_eq_refl"; +val zle_anti_sym_intify = thm "zle_anti_sym_intify"; +val zle_anti_sym = thm "zle_anti_sym"; +val zle_trans = thm "zle_trans"; +val zle_zless_trans = thm "zle_zless_trans"; +val zless_zle_trans = thm "zless_zle_trans"; +val not_zless_iff_zle = thm "not_zless_iff_zle"; +val not_zle_iff_zless = thm "not_zle_iff_zless"; +val zdiff_zdiff_eq = thm "zdiff_zdiff_eq"; +val zdiff_zdiff_eq2 = thm "zdiff_zdiff_eq2"; +val zdiff_zless_iff = thm "zdiff_zless_iff"; +val zless_zdiff_iff = thm "zless_zdiff_iff"; +val zdiff_eq_iff = thm "zdiff_eq_iff"; +val eq_zdiff_iff = thm "eq_zdiff_iff"; +val zdiff_zle_iff = thm "zdiff_zle_iff"; +val zle_zdiff_iff = thm "zle_zdiff_iff"; +val zcompare_rls = thms "zcompare_rls"; +val zadd_left_cancel = thm "zadd_left_cancel"; +val zadd_left_cancel_intify = thm "zadd_left_cancel_intify"; +val zadd_right_cancel = thm "zadd_right_cancel"; +val zadd_right_cancel_intify = thm "zadd_right_cancel_intify"; +val zadd_right_cancel_zless = thm "zadd_right_cancel_zless"; +val zadd_left_cancel_zless = thm "zadd_left_cancel_zless"; +val zadd_right_cancel_zle = thm "zadd_right_cancel_zle"; +val zadd_left_cancel_zle = thm "zadd_left_cancel_zle"; +val zadd_zless_mono1 = thm "zadd_zless_mono1"; +val zadd_zless_mono2 = thm "zadd_zless_mono2"; +val zadd_zle_mono1 = thm "zadd_zle_mono1"; +val zadd_zle_mono2 = thm "zadd_zle_mono2"; +val zadd_zle_mono = thm "zadd_zle_mono"; +val zadd_zless_mono = thm "zadd_zless_mono"; +val zminus_zless_zminus = thm "zminus_zless_zminus"; +val zminus_zle_zminus = thm "zminus_zle_zminus"; +val equation_zminus = thm "equation_zminus"; +val zminus_equation = thm "zminus_equation"; +val equation_zminus_intify = thm "equation_zminus_intify"; +val zminus_equation_intify = thm "zminus_equation_intify"; +val zless_zminus = thm "zless_zminus"; +val zminus_zless = thm "zminus_zless"; +val zle_zminus = thm "zle_zminus"; +val zminus_zle = thm "zminus_zle"; +*} + + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/IntArith.thy Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,5 @@ + +theory IntArith imports Bin +uses "int_arith.ML" begin + +end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/IntDiv.thy Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,1925 @@ +(* Title: ZF/IntDiv.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1999 University of Cambridge + +Here is the division algorithm in ML: + + fun posDivAlg (a,b) = + if a<b then (0,a) + else let val (q,r) = posDivAlg(a, 2*b) + in if 0<=r-b then (2*q+1, r-b) else (2*q, r) + end + + fun negDivAlg (a,b) = + if 0<=a+b then (~1,a+b) + else let val (q,r) = negDivAlg(a, 2*b) + in if 0<=r-b then (2*q+1, r-b) else (2*q, r) + end; + + fun negateSnd (q,r:int) = (q,~r); + + fun divAlg (a,b) = if 0<=a then + if b>0 then posDivAlg (a,b) + else if a=0 then (0,0) + else negateSnd (negDivAlg (~a,~b)) + else + if 0<b then negDivAlg (a,b) + else negateSnd (posDivAlg (~a,~b)); + +*) + +header{*The Division Operators Div and Mod*} + +theory IntDiv imports IntArith OrderArith begin + +constdefs + quorem :: "[i,i] => o" + "quorem == %<a,b> <q,r>. + a = b$*q $+ r & + (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" + + adjust :: "[i,i] => i" + "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b> + else <#2$*q,r>" + + +(** the division algorithm **) + +constdefs posDivAlg :: "i => i" +(*for the case a>=0, b>0*) +(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) + "posDivAlg(ab) == + wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)), + ab, + %<a,b> f. if (a$<b | b$<=#0) then <#0,a> + else adjust(b, f ` <a,#2$*b>))" + + +(*for the case a<0, b>0*) +constdefs negDivAlg :: "i => i" +(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) + "negDivAlg(ab) == + wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)), + ab, + %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> + else adjust(b, f ` <a,#2$*b>))" + +(*for the general case b\<noteq>0*) + +constdefs + negateSnd :: "i => i" + "negateSnd == %<q,r>. <q, $-r>" + + (*The full division algorithm considers all possible signs for a, b + including the special case a=0, b<0, because negDivAlg requires a<0*) + divAlg :: "i => i" + "divAlg == + %<a,b>. if #0 $<= a then + if #0 $<= b then posDivAlg (<a,b>) + else if a=#0 then <#0,#0> + else negateSnd (negDivAlg (<$-a,$-b>)) + else + if #0$<b then negDivAlg (<a,b>) + else negateSnd (posDivAlg (<$-a,$-b>))" + + zdiv :: "[i,i]=>i" (infixl "zdiv" 70) + "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" + + zmod :: "[i,i]=>i" (infixl "zmod" 70) + "a zmod b == snd (divAlg (<intify(a), intify(b)>))" + + +(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) + +lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" +apply (rule_tac y = "y" in zless_trans) +apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) +apply auto +done + +lemma zpos_add_zpos_imp_zpos: "[| #0 $<= x; #0 $<= y |] ==> #0 $<= x $+ y" +apply (rule_tac y = "y" in zle_trans) +apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) +apply auto +done + +lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" +apply (rule_tac y = "y" in zless_trans) +apply (rule zless_zdiff_iff [THEN iffD1]) +apply auto +done + +(* this theorem is used below *) +lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: + "[| x $<= #0; y $<= #0 |] ==> x $+ y $<= #0" +apply (rule_tac y = "y" in zle_trans) +apply (rule zle_zdiff_iff [THEN iffD1]) +apply auto +done + +lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)" +apply (drule zero_zless_imp_znegative_zminus) +apply (drule_tac [2] zneg_int_of) +apply (auto simp add: zminus_equation [of k]) +apply (subgoal_tac "0 < zmagnitude ($# succ (n))") + apply simp +apply (simp only: zmagnitude_int_of) +apply simp +done + + +(*** Inequality lemmas involving $#succ(m) ***) + +lemma zless_add_succ_iff: + "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)" +apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric]) +apply (rule_tac [3] x = "0" in bexI) +apply (cut_tac m = "m" in int_succ_int_1) +apply (cut_tac m = "n" in int_succ_int_1) +apply simp +apply (erule natE) +apply auto +apply (rule_tac x = "succ (n) " in bexI) +apply auto +done + +lemma zadd_succ_lemma: + "z \<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" +apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) +apply (auto intro: zle_anti_sym elim: zless_asym + simp add: zless_imp_zle not_zless_iff_zle) +done + +lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" +apply (cut_tac z = "intify (z)" in zadd_succ_lemma) +apply auto +done + +(** Inequality reasoning **) + +lemma zless_add1_iff_zle: "(w $< z $+ #1) <-> (w$<=z)" +apply (subgoal_tac "#1 = $# 1") +apply (simp only: zless_add_succ_iff zle_def) +apply auto +done + +lemma add1_zle_iff: "(w $+ #1 $<= z) <-> (w $< z)" +apply (subgoal_tac "#1 = $# 1") +apply (simp only: zadd_succ_zle_iff) +apply auto +done + +lemma add1_left_zle_iff: "(#1 $+ w $<= z) <-> (w $< z)" +apply (subst zadd_commute) +apply (rule add1_zle_iff) +done + + +(*** Monotonicity of Multiplication ***) + +lemma zmult_mono_lemma: "k \<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k" +apply (induct_tac "k") + prefer 2 apply (subst int_succ_int_1) +apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) +done + +lemma zmult_zle_mono1: "[| i $<= j; #0 $<= k |] ==> i$*k $<= j$*k" +apply (subgoal_tac "i $* intify (k) $<= j $* intify (k) ") +apply (simp (no_asm_use)) +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) +apply (rule_tac [3] zmult_mono_lemma) +apply auto +apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) +done + +lemma zmult_zle_mono1_neg: "[| i $<= j; k $<= #0 |] ==> j$*k $<= i$*k" +apply (rule zminus_zle_zminus [THEN iffD1]) +apply (simp del: zmult_zminus_right + add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) +done + +lemma zmult_zle_mono2: "[| i $<= j; #0 $<= k |] ==> k$*i $<= k$*j" +apply (drule zmult_zle_mono1) +apply (simp_all add: zmult_commute) +done + +lemma zmult_zle_mono2_neg: "[| i $<= j; k $<= #0 |] ==> k$*j $<= k$*i" +apply (drule zmult_zle_mono1_neg) +apply (simp_all add: zmult_commute) +done + +(* $<= monotonicity, BOTH arguments*) +lemma zmult_zle_mono: + "[| i $<= j; k $<= l; #0 $<= j; #0 $<= k |] ==> i$*k $<= j$*l" +apply (erule zmult_zle_mono1 [THEN zle_trans]) +apply assumption +apply (erule zmult_zle_mono2) +apply assumption +done + + +(** strict, in 1st argument; proof is by induction on k>0 **) + +lemma zmult_zless_mono2_lemma [rule_format]: + "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j" +apply (induct_tac "k") + prefer 2 + apply (subst int_succ_int_1) + apply (erule natE) +apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def) +apply (frule nat_0_le) +apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ") +apply (simp (no_asm_use)) +apply (rule zadd_zless_mono) +apply (simp_all (no_asm_simp) add: zle_def) +done + +lemma zmult_zless_mono2: "[| i$<j; #0 $< k |] ==> k$*i $< k$*j" +apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") +apply (simp (no_asm_use)) +apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) +apply (rule_tac [3] zmult_zless_mono2_lemma) +apply auto +apply (simp add: znegative_iff_zless_0) +apply (drule zless_trans, assumption) +apply (auto simp add: zero_lt_zmagnitude) +done + +lemma zmult_zless_mono1: "[| i$<j; #0 $< k |] ==> i$*k $< j$*k" +apply (drule zmult_zless_mono2) +apply (simp_all add: zmult_commute) +done + +(* < monotonicity, BOTH arguments*) +lemma zmult_zless_mono: + "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" +apply (erule zmult_zless_mono1 [THEN zless_trans]) +apply assumption +apply (erule zmult_zless_mono2) +apply assumption +done + +lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" +apply (rule zminus_zless_zminus [THEN iffD1]) +apply (simp del: zmult_zminus_right + add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) +done + +lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" +apply (rule zminus_zless_zminus [THEN iffD1]) +apply (simp del: zmult_zminus + add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) +done + + +(** Products of zeroes **) + +lemma zmult_eq_lemma: + "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) <-> (m$*n = #0)" +apply (case_tac "m $< #0") +apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) +apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ +done + +lemma zmult_eq_0_iff [iff]: "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)" +apply (simp add: zmult_eq_lemma) +done + + +(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, + but not (yet?) for k*m < n*k. **) + +lemma zmult_zless_lemma: + "[| k \<in> int; m \<in> int; n \<in> int |] + ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +apply (case_tac "k = #0") +apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg) +apply (auto simp add: not_zless_iff_zle + not_zle_iff_zless [THEN iff_sym, of "m$*k"] + not_zle_iff_zless [THEN iff_sym, of m]) +apply (auto elim: notE + simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg) +done + +lemma zmult_zless_cancel2: + "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" + in zmult_zless_lemma) +apply auto +done + +lemma zmult_zless_cancel1: + "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" +by (simp add: zmult_commute [of k] zmult_zless_cancel2) + +lemma zmult_zle_cancel2: + "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) + +lemma zmult_zle_cancel1: + "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" +by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) + +lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)" +apply (blast intro: zle_refl zle_anti_sym) +done + +lemma zmult_cancel2_lemma: + "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)" +apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) +apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) +done + +lemma zmult_cancel2 [simp]: + "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))" +apply (rule iff_trans) +apply (rule_tac [2] zmult_cancel2_lemma) +apply auto +done + +lemma zmult_cancel1 [simp]: + "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))" +by (simp add: zmult_commute [of k] zmult_cancel2) + + +subsection{* Uniqueness and monotonicity of quotients and remainders *} + +lemma unique_quotient_lemma: + "[| b$*q' $+ r' $<= b$*q $+ r; #0 $<= r'; #0 $< b; r $< b |] + ==> q' $<= q" +apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r") + prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) +apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") + prefer 2 + apply (erule zle_zless_trans) + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) + apply (erule zle_zless_trans) + apply (simp add: ); +apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)") + prefer 2 + apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) +apply (auto elim: zless_asym + simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls) +done + +lemma unique_quotient_lemma_neg: + "[| b$*q' $+ r' $<= b$*q $+ r; r $<= #0; b $< #0; b $< r' |] + ==> q $<= q'" +apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" + in unique_quotient_lemma) +apply (auto simp del: zminus_zadd_distrib + simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus) +done + + +lemma unique_quotient: + "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; + q \<in> int; q' \<in> int |] ==> q = q'" +apply (simp add: split_ifs quorem_def neq_iff_zless) +apply safe +apply simp_all +apply (blast intro: zle_anti_sym + dest: zle_eq_refl [THEN unique_quotient_lemma] + zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+ +done + +lemma unique_remainder: + "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; + q \<in> int; q' \<in> int; + r \<in> int; r' \<in> int |] ==> r = r'" +apply (subgoal_tac "q = q'") + prefer 2 apply (blast intro: unique_quotient) +apply (simp add: quorem_def) +done + + +subsection{*Correctness of posDivAlg, + the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *} + +lemma adjust_eq [simp]: + "adjust(b, <q,r>) = (let diff = r$-b in + if #0 $<= diff then <#2$*q $+ #1,diff> + else <#2$*q,r>)" +by (simp add: Let_def adjust_def) + + +lemma posDivAlg_termination: + "[| #0 $< b; ~ a $< b |] + ==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)" +apply (simp (no_asm) add: zless_nat_conj) +apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) +done + +lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] + +lemma posDivAlg_eqn: + "[| #0 $< b; a \<in> int; b \<in> int |] ==> + posDivAlg(<a,b>) = + (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))" +apply (rule posDivAlg_unfold [THEN trans]) +apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: posDivAlg_termination) +done + +lemma posDivAlg_induct_lemma [rule_format]: + assumes prem: + "!!a b. [| a \<in> int; b \<in> int; + ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)" + shows "<u,v> \<in> int*int --> P(<u,v>)" +apply (rule_tac a = "<u,v>" in wf_induct) +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" + in wf_measure) +apply clarify +apply (rule prem) +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) +apply auto +apply (simp add: not_zle_iff_zless posDivAlg_termination) +done + + +lemma posDivAlg_induct [consumes 2]: + assumes u_int: "u \<in> int" + and v_int: "v \<in> int" + and ih: "!!a b. [| a \<in> int; b \<in> int; + ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)") +apply simp +apply (rule posDivAlg_induct_lemma) +apply (simp (no_asm_use)) +apply (rule ih) +apply (auto simp add: u_int v_int) +done + +(*FIXME: use intify in integ_of so that we always have integ_of w \<in> int. + then this rewrite can work for ALL constants!!*) +lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)" +apply (simp (no_asm) add: int_eq_iff_zle) +done + + +subsection{* Some convenient biconditionals for products of signs *} + +lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" +apply (drule zmult_zless_mono1) +apply auto +done + +lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" +apply (drule zmult_zless_mono1_neg) +apply auto +done + +lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" +apply (drule zmult_zless_mono1_neg) +apply auto +done + +(** Inequality reasoning **) + +lemma int_0_less_lemma: + "[| x \<in> int; y \<in> int |] + ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" +apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) +apply (rule ccontr) +apply (rule_tac [2] ccontr) +apply (auto simp add: zle_def not_zless_iff_zle) +apply (erule_tac P = "#0$< x$* y" in rev_mp) +apply (erule_tac [2] P = "#0$< x$* y" in rev_mp) +apply (drule zmult_pos_neg, assumption) + prefer 2 + apply (drule zmult_pos_neg, assumption) +apply (auto dest: zless_not_sym simp add: zmult_commute) +done + +lemma int_0_less_mult_iff: + "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) +apply auto +done + +lemma int_0_le_lemma: + "[| x \<in> int; y \<in> int |] + ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)" +by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) + +lemma int_0_le_mult_iff: + "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))" +apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) +apply auto +done + +lemma zmult_less_0_iff: + "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" +apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) +apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) +done + +lemma zmult_le_0_iff: + "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)" +by (auto dest: zless_not_sym + simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) + + +(*Typechecking for posDivAlg*) +lemma posDivAlg_type [rule_format]: + "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int" +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) +apply assumption+ +apply (case_tac "#0 $< ba") + apply (simp add: posDivAlg_eqn adjust_def integ_of_type + split add: split_if_asm) + apply clarify + apply (simp add: int_0_less_mult_iff not_zle_iff_zless) +apply (simp add: not_zless_iff_zle) +apply (subst posDivAlg_unfold) +apply simp +done + +(*Correctness of posDivAlg: it computes quotients correctly*) +lemma posDivAlg_correct [rule_format]: + "[| a \<in> int; b \<in> int |] + ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))" +apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) +apply auto + apply (simp_all add: quorem_def) + txt{*base case: a<b*} + apply (simp add: posDivAlg_eqn) + apply (simp add: not_zless_iff_zle [THEN iff_sym]) + apply (simp add: int_0_less_mult_iff) +txt{*main argument*} +apply (subst posDivAlg_eqn) +apply (simp_all (no_asm_simp)) +apply (erule splitE) +apply (rule posDivAlg_type) +apply (simp_all add: int_0_less_mult_iff) +apply (auto simp add: zadd_zmult_distrib2 Let_def) +txt{*now just linear arithmetic*} +apply (simp add: not_zle_iff_zless zdiff_zless_iff) +done + + +subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*} + +lemma negDivAlg_termination: + "[| #0 $< b; a $+ b $< #0 |] + ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" +apply (simp (no_asm) add: zless_nat_conj) +apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] + zless_zminus) +done + +lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] + +lemma negDivAlg_eqn: + "[| #0 $< b; a : int; b : int |] ==> + negDivAlg(<a,b>) = + (if #0 $<= a$+b then <#-1,a$+b> + else adjust(b, negDivAlg (<a, #2$*b>)))" +apply (rule negDivAlg_unfold [THEN trans]) +apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: negDivAlg_termination) +done + +lemma negDivAlg_induct_lemma [rule_format]: + assumes prem: + "!!a b. [| a \<in> int; b \<in> int; + ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] + ==> P(<a,b>)" + shows "<u,v> \<in> int*int --> P(<u,v>)" +apply (rule_tac a = "<u,v>" in wf_induct) +apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" + in wf_measure) +apply clarify +apply (rule prem) +apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) +apply auto +apply (simp add: not_zle_iff_zless negDivAlg_termination) +done + +lemma negDivAlg_induct [consumes 2]: + assumes u_int: "u \<in> int" + and v_int: "v \<in> int" + and ih: "!!a b. [| a \<in> int; b \<in> int; + ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] + ==> P(a,b)" + shows "P(u,v)" +apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)") +apply simp +apply (rule negDivAlg_induct_lemma) +apply (simp (no_asm_use)) +apply (rule ih) +apply (auto simp add: u_int v_int) +done + + +(*Typechecking for negDivAlg*) +lemma negDivAlg_type: + "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int" +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) +apply assumption+ +apply (case_tac "#0 $< ba") + apply (simp add: negDivAlg_eqn adjust_def integ_of_type + split add: split_if_asm) + apply clarify + apply (simp add: int_0_less_mult_iff not_zle_iff_zless) +apply (simp add: not_zless_iff_zle) +apply (subst negDivAlg_unfold) +apply simp +done + + +(*Correctness of negDivAlg: it computes quotients correctly + It doesn't work if a=0 because the 0/b=0 rather than -1*) +lemma negDivAlg_correct [rule_format]: + "[| a \<in> int; b \<in> int |] + ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))" +apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) + apply auto + apply (simp_all add: quorem_def) + txt{*base case: @{term "0$<=a$+b"}*} + apply (simp add: negDivAlg_eqn) + apply (simp add: not_zless_iff_zle [THEN iff_sym]) + apply (simp add: int_0_less_mult_iff) +txt{*main argument*} +apply (subst negDivAlg_eqn) +apply (simp_all (no_asm_simp)) +apply (erule splitE) +apply (rule negDivAlg_type) +apply (simp_all add: int_0_less_mult_iff) +apply (auto simp add: zadd_zmult_distrib2 Let_def) +txt{*now just linear arithmetic*} +apply (simp add: not_zle_iff_zless zdiff_zless_iff) +done + + +subsection{* Existence shown by proving the division algorithm to be correct *} + +(*the case a=0*) +lemma quorem_0: "[|b \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)" +by (force simp add: quorem_def neq_iff_zless) + +lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>" +apply (subst posDivAlg_unfold) +apply simp +done + +lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>" +apply (subst posDivAlg_unfold) +apply (simp add: not_zle_iff_zless) +done + + +(*Needed below. Actually it's an equivalence.*) +lemma linear_arith_lemma: "~ (#0 $<= #-1 $+ b) ==> (b $<= #0)" +apply (simp add: not_zle_iff_zless) +apply (drule zminus_zless_zminus [THEN iffD2]) +apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) +done + +lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>" +apply (subst negDivAlg_unfold) +apply (simp add: linear_arith_lemma integ_of_type vimage_iff) +done + +lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>" +apply (unfold negateSnd_def) +apply auto +done + +lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int" +apply (unfold negateSnd_def) +apply auto +done + +lemma quorem_neg: + "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|] + ==> quorem (<a,b>, negateSnd(qr))" +apply clarify +apply (auto elim: zless_asym simp add: quorem_def zless_zminus) +txt{*linear arithmetic from here on*} +apply (simp_all add: zminus_equation [of a] zminus_zless) +apply (cut_tac [2] z = "b" and w = "#0" in zless_linear) +apply (cut_tac [1] z = "b" and w = "#0" in zless_linear) +apply auto +apply (blast dest: zle_zless_trans)+ +done + +lemma divAlg_correct: + "[|b \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))" +apply (auto simp add: quorem_0 divAlg_def) +apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct + posDivAlg_type negDivAlg_type) +apply (auto simp add: quorem_def neq_iff_zless) +txt{*linear arithmetic from here on*} +apply (auto simp add: zle_def) +done + +lemma divAlg_type: "[|a \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> int * int" +apply (auto simp add: divAlg_def) +apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) +done + + +(** intify cancellation **) + +lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y" +apply (simp (no_asm) add: zdiv_def) +done + +lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" +apply (simp (no_asm) add: zdiv_def) +done + +lemma zdiv_type [iff,TC]: "z zdiv w \<in> int" +apply (unfold zdiv_def) +apply (blast intro: fst_type divAlg_type) +done + +lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" +apply (simp (no_asm) add: zmod_def) +done + +lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" +apply (simp (no_asm) add: zmod_def) +done + +lemma zmod_type [iff,TC]: "z zmod w \<in> int" +apply (unfold zmod_def) +apply (rule snd_type) +apply (blast intro: divAlg_type) +done + + +(** Arbitrary definitions for division by zero. Useful to simplify + certain equations **) + +lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0" +apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor) +done (*NOT for adding to default simpset*) + +lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" +apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor) +done (*NOT for adding to default simpset*) + + + +(** Basic laws about division and remainder **) + +lemma raw_zmod_zdiv_equality: + "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (cut_tac a = "a" and b = "b" in divAlg_correct) +apply (auto simp add: quorem_def zdiv_def zmod_def split_def) +done + +lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)" +apply (rule trans) +apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality) +apply auto +done + +lemma pos_mod: "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) +apply (blast dest: zle_zless_trans)+ +done + +lemmas pos_mod_sign = pos_mod [THEN conjunct1, standard] +and pos_mod_bound = pos_mod [THEN conjunct2, standard] + +lemma neg_mod: "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) +apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) +apply (blast dest: zle_zless_trans) +apply (blast dest: zless_trans)+ +done + +lemmas neg_mod_sign = neg_mod [THEN conjunct1, standard] +and neg_mod_bound = neg_mod [THEN conjunct2, standard] + + +(** proving general properties of zdiv and zmod **) + +lemma quorem_div_mod: + "[|b \<noteq> #0; a \<in> int; b \<in> int |] + ==> quorem (<a,b>, <a zdiv b, a zmod b>)" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound + neg_mod_sign neg_mod_bound) +done + +(*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*) +lemma quorem_div: + "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |] + ==> a zdiv b = q" +by (blast intro: quorem_div_mod [THEN unique_quotient]) + +lemma quorem_mod: + "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] + ==> a zmod b = r" +by (blast intro: quorem_div_mod [THEN unique_remainder]) + +lemma zdiv_pos_pos_trivial_raw: + "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zdiv b = #0" +apply (rule quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans)+ +done + +lemma zdiv_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_pos_pos_trivial_raw) +apply auto +done + +lemma zdiv_neg_neg_trivial_raw: + "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zdiv b = #0" +apply (rule_tac r = "a" in quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans zless_trans)+ +done + +lemma zdiv_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_neg_neg_trivial_raw) +apply auto +done + +lemma zadd_le_0_lemma: "[| a$+b $<= #0; #0 $< a; #0 $< b |] ==> False" +apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) +apply (auto simp add: zle_def) +apply (blast dest: zless_trans) +done + +lemma zdiv_pos_neg_trivial_raw: + "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" +apply (rule_tac r = "a $+ b" in quorem_div) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ +done + +lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zdiv_pos_neg_trivial_raw) +apply auto +done + +(*There is no zdiv_neg_pos_trivial because #0 zdiv b = #0 would supersede it*) + + +lemma zmod_pos_pos_trivial_raw: + "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zmod b = a" +apply (rule_tac q = "#0" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans)+ +done + +lemma zmod_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zmod b = intify(a)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_pos_pos_trivial_raw) +apply auto +done + +lemma zmod_neg_neg_trivial_raw: + "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zmod b = a" +apply (rule_tac q = "#0" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zle_zless_trans zless_trans)+ +done + +lemma zmod_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zmod b = intify(a)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_neg_neg_trivial_raw) +apply auto +done + +lemma zmod_pos_neg_trivial_raw: + "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" +apply (rule_tac q = "#-1" in quorem_mod) +apply (auto simp add: quorem_def) +(*linear arithmetic*) +apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ +done + +lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" + in zmod_pos_neg_trivial_raw) +apply auto +done + +(*There is no zmod_neg_pos_trivial...*) + + +(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) + +lemma zdiv_zminus_zminus_raw: + "[|a \<in> int; b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) +apply auto +done + +lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw) +apply auto +done + +(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) +lemma zmod_zminus_zminus_raw: + "[|a \<in> int; b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) +apply auto +done + +lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw) +apply auto +done + + +subsection{* division of a number by itself *} + +lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q" +apply (subgoal_tac "#0 $< a$*q") +apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) +apply (simp add: int_0_less_mult_iff) +apply (blast dest: zless_trans) +(*linear arithmetic...*) +apply (drule_tac t = "%x. x $- r" in subst_context) +apply (drule sym) +apply (simp add: zcompare_rls) +done + +lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1" +apply (subgoal_tac "#0 $<= a$* (#1$-q)") + apply (simp add: int_0_le_mult_iff zcompare_rls) + apply (blast dest: zle_zless_trans) +apply (simp add: zdiff_zmult_distrib2) +apply (drule_tac t = "%x. x $- a $* q" in subst_context) +apply (simp add: zcompare_rls) +done + +lemma self_quotient: + "[| quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #0|] ==> q = #1" +apply (simp add: split_ifs quorem_def neq_iff_zless) +apply (rule zle_anti_sym) +apply safe +apply auto +prefer 4 apply (blast dest: zless_trans) +apply (blast dest: zless_trans) +apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1) +apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2) +apply (rule_tac [6] zminus_equation [THEN iffD1]) +apply (rule_tac [2] zminus_equation [THEN iffD1]) +apply (force intro: self_quotient_aux1 self_quotient_aux2 + simp add: zadd_commute zmult_zminus)+ +done + +lemma self_remainder: + "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0" +apply (frule self_quotient) +apply (auto simp add: quorem_def) +done + +lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1" +apply (blast intro: quorem_div_mod [THEN self_quotient]) +done + +lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1" +apply (drule zdiv_self_raw) +apply auto +done + +(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) +lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: quorem_div_mod [THEN self_remainder]) +done + +lemma zmod_self [simp]: "a zmod a = #0" +apply (cut_tac a = "intify (a)" in zmod_self_raw) +apply auto +done + + +subsection{* Computation of division and remainder *} + +lemma zdiv_zero [simp]: "#0 zdiv b = #0" +apply (simp (no_asm) add: zdiv_def divAlg_def) +done + +lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +done + +lemma zmod_zero [simp]: "#0 zmod b = #0" +apply (simp (no_asm) add: zmod_def divAlg_def) +done + +lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +done + +lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +done + +(** a positive, b positive **) + +lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |] + ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply (auto simp add: zle_def) +done + +lemma zmod_pos_pos: + "[| #0 $< a; #0 $<= b |] + ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply (auto simp add: zle_def) +done + +(** a negative, b positive **) + +lemma zdiv_neg_pos: + "[| a $< #0; #0 $< b |] + ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply (blast dest: zle_zless_trans) +done + +lemma zmod_neg_pos: + "[| a $< #0; #0 $< b |] + ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply (blast dest: zle_zless_trans) +done + +(** a positive, b negative **) + +lemma zdiv_pos_neg: + "[| #0 $< a; b $< #0 |] + ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) +apply auto +apply (blast dest: zle_zless_trans)+ +apply (blast dest: zless_trans) +apply (blast intro: zless_imp_zle) +done + +lemma zmod_pos_neg: + "[| #0 $< a; b $< #0 |] + ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) +apply auto +apply (blast dest: zle_zless_trans)+ +apply (blast dest: zless_trans) +apply (blast intro: zless_imp_zle) +done + +(** a negative, b negative **) + +lemma zdiv_neg_neg: + "[| a $< #0; b $<= #0 |] + ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zdiv_def divAlg_def) +apply auto +apply (blast dest!: zle_zless_trans)+ +done + +lemma zmod_neg_neg: + "[| a $< #0; b $<= #0 |] + ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" +apply (simp (no_asm_simp) add: zmod_def divAlg_def) +apply auto +apply (blast dest!: zle_zless_trans)+ +done + +declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] +declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] +declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] + + +(** Special-case simplification **) + +lemma zmod_1 [simp]: "a zmod #1 = #0" +apply (cut_tac a = "a" and b = "#1" in pos_mod_sign) +apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound) +apply auto +(*arithmetic*) +apply (drule add1_zle_iff [THEN iffD2]) +apply (rule zle_anti_sym) +apply auto +done + +lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)" +apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality) +apply auto +done + +lemma zmod_minus1_right [simp]: "a zmod #-1 = #0" +apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign) +apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound) +apply auto +(*arithmetic*) +apply (drule add1_zle_iff [THEN iffD2]) +apply (rule zle_anti_sym) +apply auto +done + +lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a" +apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) +apply auto +apply (rule equation_zminus [THEN iffD2]) +apply auto +done + +lemma zdiv_minus1_right: "a zdiv #-1 = $-a" +apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw) +apply auto +done +declare zdiv_minus1_right [simp] + + +subsection{* Monotonicity in the first argument (divisor) *} + +lemma zdiv_mono1: "[| a $<= a'; #0 $< b |] ==> a zdiv b $<= a' zdiv b" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) +apply (rule unique_quotient_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono1_neg: "[| a $<= a'; b $< #0 |] ==> a' zdiv b $<= a zdiv b" +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) +apply (rule unique_quotient_lemma_neg) +apply (erule subst) +apply (erule subst) +apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound) +done + + +subsection{* Monotonicity in the second argument (dividend) *} + +lemma q_pos_lemma: + "[| #0 $<= b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $<= q'" +apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") + apply (simp add: int_0_less_mult_iff) + apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) +apply (simp add: zadd_zmult_distrib2) +apply (erule zle_zless_trans) +apply (erule zadd_zless_mono2) +done + +lemma zdiv_mono2_lemma: + "[| b$*q $+ r = b'$*q' $+ r'; #0 $<= b'$*q' $+ r'; + r' $< b'; #0 $<= r; #0 $< b'; b' $<= b |] + ==> q $<= q'" +apply (frule q_pos_lemma, assumption+) +apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") + apply (simp add: zmult_zless_cancel1) + apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans) +apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") + prefer 2 apply (simp add: zcompare_rls) +apply (simp (no_asm_simp) add: zadd_zmult_distrib2) +apply (subst zadd_commute [of "b $\<times> q'"], rule zadd_zless_mono) + prefer 2 apply (blast intro: zmult_zle_mono1) +apply (subgoal_tac "r' $+ #0 $< b $+ r") + apply (simp add: zcompare_rls) +apply (rule zadd_zless_mono) + apply auto +apply (blast dest: zless_zle_trans) +done + + +lemma zdiv_mono2_raw: + "[| #0 $<= a; #0 $< b'; b' $<= b; a \<in> int |] + ==> a zdiv b $<= a zdiv b'" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) +apply (rule zdiv_mono2_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono2: + "[| #0 $<= a; #0 $< b'; b' $<= b |] + ==> a zdiv b $<= a zdiv b'" +apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) +apply auto +done + +lemma q_neg_lemma: + "[| b'$*q' $+ r' $< #0; #0 $<= r'; #0 $< b' |] ==> q' $< #0" +apply (subgoal_tac "b'$*q' $< #0") + prefer 2 apply (force intro: zle_zless_trans) +apply (simp add: zmult_less_0_iff) +apply (blast dest: zless_trans) +done + + + +lemma zdiv_mono2_neg_lemma: + "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; + r $< b; #0 $<= r'; #0 $< b'; b' $<= b |] + ==> q' $<= q" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (frule q_neg_lemma, assumption+) +apply (subgoal_tac "b$*q' $< b$* (q $+ #1)") + apply (simp add: zmult_zless_cancel1) + apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1]) +apply (simp (no_asm_simp) add: zadd_zmult_distrib2) +apply (subgoal_tac "b$*q' $<= b'$*q'") + prefer 2 + apply (simp add: zmult_zle_cancel2) + apply (blast dest: zless_trans) +apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)") + prefer 2 + apply (erule ssubst) + apply simp + apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono) + apply (assumption) + apply simp +apply (simp (no_asm_use) add: zadd_commute) +apply (rule zle_zless_trans) + prefer 2 apply (assumption) +apply (simp (no_asm_simp) add: zmult_zle_cancel2) +apply (blast dest: zless_trans) +done + +lemma zdiv_mono2_neg_raw: + "[| a $< #0; #0 $< b'; b' $<= b; a \<in> int |] + ==> a zdiv b' $<= a zdiv b" +apply (subgoal_tac "#0 $< b") + prefer 2 apply (blast dest: zless_zle_trans) +apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) +apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) +apply (rule zdiv_mono2_neg_lemma) +apply (erule subst) +apply (erule subst) +apply (simp_all add: pos_mod_sign pos_mod_bound) +done + +lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $<= b |] + ==> a zdiv b' $<= a zdiv b" +apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) +apply auto +done + + + +subsection{* More algebraic laws for zdiv and zmod *} + +(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) + +lemma zmult1_lemma: + "[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |] + ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)" +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 + pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) +apply (auto intro: raw_zmod_zdiv_equality) +done + +lemma zdiv_zmult1_eq_raw: + "[|b \<in> int; c \<in> int|] + ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) +apply auto +done + +lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw) +apply auto +done + +lemma zmod_zmult1_eq_raw: + "[|b \<in> int; c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) +apply auto +done + +lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw) +apply auto +done + +lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c" +apply (rule trans) +apply (rule_tac b = " (b $* a) zmod c" in trans) +apply (rule_tac [2] zmod_zmult1_eq) +apply (simp_all (no_asm) add: zmult_commute) +done + +lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c" +apply (rule zmod_zmult1_eq' [THEN trans]) +apply (rule zmod_zmult1_eq) +done + +lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)" +apply (simp (no_asm_simp) add: zdiv_zmult1_eq) +done + +lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)" +apply (subst zmult_commute , erule zdiv_zmult_self1) +done + +lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" +apply (simp (no_asm) add: zmod_zmult1_eq) +done + +lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" +apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq) +done + + +(** proving (a$+b) zdiv c = + a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) + +lemma zadd1_lemma: + "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>); + c \<in> int; c \<noteq> #0 |] + ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)" +apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 + pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) +apply (auto intro: raw_zmod_zdiv_equality) +done + +(*NOT suitable for rewriting: the RHS has an instance of the LHS*) +lemma zdiv_zadd1_eq_raw: + "[|a \<in> int; b \<in> int; c \<in> int|] ==> + (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, + THEN quorem_div]) +done + +lemma zdiv_zadd1_eq: + "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" + in zdiv_zadd1_eq_raw) +apply auto +done + +lemma zmod_zadd1_eq_raw: + "[|a \<in> int; b \<in> int; c \<in> int|] + ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, + THEN quorem_mod]) +done + +lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" +apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" + in zmod_zadd1_eq_raw) +apply auto +done + +lemma zmod_div_trivial_raw: + "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound + zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial) +done + +lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw) +apply auto +done + +lemma zmod_mod_trivial_raw: + "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound + zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial) +done + +lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw) +apply auto +done + +lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c" +apply (rule trans [symmetric]) +apply (rule zmod_zadd1_eq) +apply (simp (no_asm)) +apply (rule zmod_zadd1_eq [symmetric]) +done + +lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c" +apply (rule trans [symmetric]) +apply (rule zmod_zadd1_eq) +apply (simp (no_asm)) +apply (rule zmod_zadd1_eq [symmetric]) +done + + +lemma zdiv_zadd_self1 [simp]: + "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" +by (simp (no_asm_simp) add: zdiv_zadd1_eq) + +lemma zdiv_zadd_self2 [simp]: + "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" +by (simp (no_asm_simp) add: zdiv_zadd1_eq) + +lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (simp (no_asm_simp) add: zmod_zadd1_eq) +done + +lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a" +apply (case_tac "a = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (simp (no_asm_simp) add: zmod_zadd1_eq) +done + + +subsection{* proving a zdiv (b*c) = (a zdiv b) zdiv c *} + +(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but + 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems + to cause particular problems.*) + +(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) + +lemma zdiv_zmult2_aux1: + "[| #0 $< c; b $< r; r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r" +apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) +apply (rule zle_zless_trans) +apply (erule_tac [2] zmult_zless_mono1) +apply (rule zmult_zle_mono2_neg) +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) +apply (blast intro: zless_imp_zle dest: zless_zle_trans) +done + +lemma zdiv_zmult2_aux2: + "[| #0 $< c; b $< r; r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0" +apply (subgoal_tac "b $* (q zmod c) $<= #0") + prefer 2 + apply (simp add: zmult_le_0_iff pos_mod_sign) + apply (blast intro: zless_imp_zle dest: zless_zle_trans) +(*arithmetic*) +apply (drule zadd_zle_mono) +apply assumption +apply (simp add: zadd_commute) +done + +lemma zdiv_zmult2_aux3: + "[| #0 $< c; #0 $<= r; r $< b |] ==> #0 $<= b $* (q zmod c) $+ r" +apply (subgoal_tac "#0 $<= b $* (q zmod c)") + prefer 2 + apply (simp add: int_0_le_mult_iff pos_mod_sign) + apply (blast intro: zless_imp_zle dest: zle_zless_trans) +(*arithmetic*) +apply (drule zadd_zle_mono) +apply assumption +apply (simp add: zadd_commute) +done + +lemma zdiv_zmult2_aux4: + "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" +apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") +apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) +apply (rule zless_zle_trans) +apply (erule zmult_zless_mono1) +apply (rule_tac [2] zmult_zle_mono2) +apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) +apply (blast intro: zless_imp_zle dest: zle_zless_trans) +done + +lemma zdiv_zmult2_lemma: + "[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |] + ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)" +apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def + neq_iff_zless int_0_less_mult_iff + zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 + zdiv_zmult2_aux3 zdiv_zmult2_aux4) +apply (blast dest: zless_trans)+ +done + +lemma zdiv_zmult2_eq_raw: + "[|#0 $< c; a \<in> int; b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) +apply (auto simp add: intify_eq_0_iff_zle) +apply (blast dest: zle_zless_trans) +done + +lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) +apply auto +done + +lemma zmod_zmult2_eq_raw: + "[|#0 $< c; a \<in> int; b \<in> int|] + ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) +apply (auto simp add: intify_eq_0_iff_zle) +apply (blast dest: zle_zless_trans) +done + +lemma zmod_zmult2_eq: + "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" +apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) +apply auto +done + +subsection{* Cancellation of common factors in "zdiv" *} + +lemma zdiv_zmult_zmult1_aux1: + "[| #0 $< b; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (subst zdiv_zmult2_eq) +apply auto +done + +lemma zdiv_zmult_zmult1_aux2: + "[| b $< #0; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") +apply (rule_tac [2] zdiv_zmult_zmult1_aux1) +apply auto +done + +lemma zdiv_zmult_zmult1_raw: + "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless [of b] + zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) +done + +lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" +apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) +apply auto +done + +lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" +apply (drule zdiv_zmult_zmult1) +apply (auto simp add: zmult_commute) +done + + +subsection{* Distribution of factors over "zmod" *} + +lemma zmod_zmult_zmult1_aux1: + "[| #0 $< b; intify(c) \<noteq> #0 |] + ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (subst zmod_zmult2_eq) +apply auto +done + +lemma zmod_zmult_zmult1_aux2: + "[| b $< #0; intify(c) \<noteq> #0 |] + ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") +apply (rule_tac [2] zmod_zmult_zmult1_aux1) +apply auto +done + +lemma zmod_zmult_zmult1_raw: + "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (case_tac "b = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (case_tac "c = #0") + apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) +apply (auto simp add: neq_iff_zless [of b] + zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) +done + +lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)" +apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw) +apply auto +done + +lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c" +apply (cut_tac c = "c" in zmod_zmult_zmult1) +apply (auto simp add: zmult_commute) +done + + +(** Quotients of signs **) + +lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" +apply (subgoal_tac "a zdiv b $<= #-1") +apply (erule zle_zless_trans) +apply (simp (no_asm)) +apply (rule zle_trans) +apply (rule_tac a' = "#-1" in zdiv_mono1) +apply (rule zless_add1_iff_zle [THEN iffD1]) +apply (simp (no_asm)) +apply (auto simp add: zdiv_minus1) +done + +lemma zdiv_nonneg_neg_le0: "[| #0 $<= a; b $< #0 |] ==> a zdiv b $<= #0" +apply (drule zdiv_mono1_neg) +apply auto +done + +lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)" +apply auto +apply (drule_tac [2] zdiv_mono1) +apply (auto simp add: neq_iff_zless) +apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym]) +apply (blast intro: zdiv_neg_pos_less0) +done + +lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)" +apply (subst zdiv_zminus_zminus [symmetric]) +apply (rule iff_trans) +apply (rule pos_imp_zdiv_nonneg_iff) +apply auto +done + +(*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*) +lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)" +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) +apply (erule pos_imp_zdiv_nonneg_iff) +done + +(*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*) +lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)" +apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) +apply (erule neg_imp_zdiv_nonneg_iff) +done + +(* + THESE REMAIN TO BE CONVERTED -- but aren't that useful! + + subsection{* Speeding up the division algorithm with shifting *} + + (** computing "zdiv" by shifting **) + + lemma pos_zdiv_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" + apply (case_tac "a = #0") + apply (subgoal_tac "#1 $<= a") + apply (arith_tac 2) + apply (subgoal_tac "#1 $< a $* #2") + apply (arith_tac 2) + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") + apply (rule_tac [2] zmult_zle_mono2) + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) + apply (subst zdiv_zadd1_eq) + apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial) + apply (subst zdiv_pos_pos_trivial) + apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) + apply (auto simp add: zmod_pos_pos_trivial) + apply (subgoal_tac "#0 $<= b zmod a") + apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) + apply arith + done + + + lemma neg_zdiv_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a" + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)") + apply (rule_tac [2] pos_zdiv_mult_2) + apply (auto simp add: zmult_zminus_right) + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") + apply (Simp_tac 2) + apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) + done + + + (*Not clear why this must be proved separately; probably integ_of causes + simplification problems*) + lemma lemma: "~ #0 $<= x ==> x $<= #0" + apply auto + done + + lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = + (if ~b | #0 $<= integ_of w + then integ_of v zdiv (integ_of w) + else (integ_of v $+ #1) zdiv (integ_of w))" + apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) + apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) + done + + declare zdiv_integ_of_BIT [simp] + + + (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) + + lemma pos_zmod_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" + apply (case_tac "a = #0") + apply (subgoal_tac "#1 $<= a") + apply (arith_tac 2) + apply (subgoal_tac "#1 $< a $* #2") + apply (arith_tac 2) + apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") + apply (rule_tac [2] zmult_zle_mono2) + apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) + apply (subst zmod_zadd1_eq) + apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial) + apply (rule zmod_pos_pos_trivial) + apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) + apply (auto simp add: zmod_pos_pos_trivial) + apply (subgoal_tac "#0 $<= b zmod a") + apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) + apply arith + done + + + lemma neg_zmod_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" + apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") + apply (rule_tac [2] pos_zmod_mult_2) + apply (auto simp add: zmult_zminus_right) + apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") + apply (Simp_tac 2) + apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) + apply (dtac (zminus_equation [THEN iffD1, symmetric]) + apply auto + done + + lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) = + (if b then + if #0 $<= integ_of w + then #2 $* (integ_of v zmod integ_of w) $+ #1 + else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 + else #2 $* (integ_of v zmod integ_of w))" + apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) + apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) + done + + declare zmod_integ_of_BIT [simp] +*) + +ML{* +val zspos_add_zspos_imp_zspos = thm "zspos_add_zspos_imp_zspos"; +val zpos_add_zpos_imp_zpos = thm "zpos_add_zpos_imp_zpos"; +val zneg_add_zneg_imp_zneg = thm "zneg_add_zneg_imp_zneg"; +val zneg_or_0_add_zneg_or_0_imp_zneg_or_0 = thm "zneg_or_0_add_zneg_or_0_imp_zneg_or_0"; +val zero_lt_zmagnitude = thm "zero_lt_zmagnitude"; +val zless_add_succ_iff = thm "zless_add_succ_iff"; +val zadd_succ_zle_iff = thm "zadd_succ_zle_iff"; +val zless_add1_iff_zle = thm "zless_add1_iff_zle"; +val add1_zle_iff = thm "add1_zle_iff"; +val add1_left_zle_iff = thm "add1_left_zle_iff"; +val zmult_zle_mono1 = thm "zmult_zle_mono1"; +val zmult_zle_mono1_neg = thm "zmult_zle_mono1_neg"; +val zmult_zle_mono2 = thm "zmult_zle_mono2"; +val zmult_zle_mono2_neg = thm "zmult_zle_mono2_neg"; +val zmult_zle_mono = thm "zmult_zle_mono"; +val zmult_zless_mono2 = thm "zmult_zless_mono2"; +val zmult_zless_mono1 = thm "zmult_zless_mono1"; +val zmult_zless_mono = thm "zmult_zless_mono"; +val zmult_zless_mono1_neg = thm "zmult_zless_mono1_neg"; +val zmult_zless_mono2_neg = thm "zmult_zless_mono2_neg"; +val zmult_eq_0_iff = thm "zmult_eq_0_iff"; +val zmult_zless_cancel2 = thm "zmult_zless_cancel2"; +val zmult_zless_cancel1 = thm "zmult_zless_cancel1"; +val zmult_zle_cancel2 = thm "zmult_zle_cancel2"; +val zmult_zle_cancel1 = thm "zmult_zle_cancel1"; +val int_eq_iff_zle = thm "int_eq_iff_zle"; +val zmult_cancel2 = thm "zmult_cancel2"; +val zmult_cancel1 = thm "zmult_cancel1"; +val unique_quotient = thm "unique_quotient"; +val unique_remainder = thm "unique_remainder"; +val adjust_eq = thm "adjust_eq"; +val posDivAlg_termination = thm "posDivAlg_termination"; +val posDivAlg_unfold = thm "posDivAlg_unfold"; +val posDivAlg_eqn = thm "posDivAlg_eqn"; +val posDivAlg_induct = thm "posDivAlg_induct"; +val intify_eq_0_iff_zle = thm "intify_eq_0_iff_zle"; +val zmult_pos = thm "zmult_pos"; +val zmult_neg = thm "zmult_neg"; +val zmult_pos_neg = thm "zmult_pos_neg"; +val int_0_less_mult_iff = thm "int_0_less_mult_iff"; +val int_0_le_mult_iff = thm "int_0_le_mult_iff"; +val zmult_less_0_iff = thm "zmult_less_0_iff"; +val zmult_le_0_iff = thm "zmult_le_0_iff"; +val posDivAlg_type = thm "posDivAlg_type"; +val posDivAlg_correct = thm "posDivAlg_correct"; +val negDivAlg_termination = thm "negDivAlg_termination"; +val negDivAlg_unfold = thm "negDivAlg_unfold"; +val negDivAlg_eqn = thm "negDivAlg_eqn"; +val negDivAlg_induct = thm "negDivAlg_induct"; +val negDivAlg_type = thm "negDivAlg_type"; +val negDivAlg_correct = thm "negDivAlg_correct"; +val quorem_0 = thm "quorem_0"; +val posDivAlg_zero_divisor = thm "posDivAlg_zero_divisor"; +val posDivAlg_0 = thm "posDivAlg_0"; +val negDivAlg_minus1 = thm "negDivAlg_minus1"; +val negateSnd_eq = thm "negateSnd_eq"; +val negateSnd_type = thm "negateSnd_type"; +val quorem_neg = thm "quorem_neg"; +val divAlg_correct = thm "divAlg_correct"; +val divAlg_type = thm "divAlg_type"; +val zdiv_intify1 = thm "zdiv_intify1"; +val zdiv_intify2 = thm "zdiv_intify2"; +val zdiv_type = thm "zdiv_type"; +val zmod_intify1 = thm "zmod_intify1"; +val zmod_intify2 = thm "zmod_intify2"; +val zmod_type = thm "zmod_type"; +val DIVISION_BY_ZERO_ZDIV = thm "DIVISION_BY_ZERO_ZDIV"; +val DIVISION_BY_ZERO_ZMOD = thm "DIVISION_BY_ZERO_ZMOD"; +val zmod_zdiv_equality = thm "zmod_zdiv_equality"; +val pos_mod = thm "pos_mod"; +val pos_mod_sign = thm "pos_mod_sign"; +val neg_mod = thm "neg_mod"; +val neg_mod_sign = thm "neg_mod_sign"; +val quorem_div_mod = thm "quorem_div_mod"; +val quorem_div = thm "quorem_div"; +val quorem_mod = thm "quorem_mod"; +val zdiv_pos_pos_trivial = thm "zdiv_pos_pos_trivial"; +val zdiv_neg_neg_trivial = thm "zdiv_neg_neg_trivial"; +val zdiv_pos_neg_trivial = thm "zdiv_pos_neg_trivial"; +val zmod_pos_pos_trivial = thm "zmod_pos_pos_trivial"; +val zmod_neg_neg_trivial = thm "zmod_neg_neg_trivial"; +val zmod_pos_neg_trivial = thm "zmod_pos_neg_trivial"; +val zdiv_zminus_zminus = thm "zdiv_zminus_zminus"; +val zmod_zminus_zminus = thm "zmod_zminus_zminus"; +val self_quotient = thm "self_quotient"; +val self_remainder = thm "self_remainder"; +val zdiv_self = thm "zdiv_self"; +val zmod_self = thm "zmod_self"; +val zdiv_zero = thm "zdiv_zero"; +val zdiv_eq_minus1 = thm "zdiv_eq_minus1"; +val zmod_zero = thm "zmod_zero"; +val zdiv_minus1 = thm "zdiv_minus1"; +val zmod_minus1 = thm "zmod_minus1"; +val zdiv_pos_pos = thm "zdiv_pos_pos"; +val zmod_pos_pos = thm "zmod_pos_pos"; +val zdiv_neg_pos = thm "zdiv_neg_pos"; +val zmod_neg_pos = thm "zmod_neg_pos"; +val zdiv_pos_neg = thm "zdiv_pos_neg"; +val zmod_pos_neg = thm "zmod_pos_neg"; +val zdiv_neg_neg = thm "zdiv_neg_neg"; +val zmod_neg_neg = thm "zmod_neg_neg"; +val zmod_1 = thm "zmod_1"; +val zdiv_1 = thm "zdiv_1"; +val zmod_minus1_right = thm "zmod_minus1_right"; +val zdiv_minus1_right = thm "zdiv_minus1_right"; +val zdiv_mono1 = thm "zdiv_mono1"; +val zdiv_mono1_neg = thm "zdiv_mono1_neg"; +val zdiv_mono2 = thm "zdiv_mono2"; +val zdiv_mono2_neg = thm "zdiv_mono2_neg"; +val zdiv_zmult1_eq = thm "zdiv_zmult1_eq"; +val zmod_zmult1_eq = thm "zmod_zmult1_eq"; +val zmod_zmult1_eq' = thm "zmod_zmult1_eq'"; +val zmod_zmult_distrib = thm "zmod_zmult_distrib"; +val zdiv_zmult_self1 = thm "zdiv_zmult_self1"; +val zdiv_zmult_self2 = thm "zdiv_zmult_self2"; +val zmod_zmult_self1 = thm "zmod_zmult_self1"; +val zmod_zmult_self2 = thm "zmod_zmult_self2"; +val zdiv_zadd1_eq = thm "zdiv_zadd1_eq"; +val zmod_zadd1_eq = thm "zmod_zadd1_eq"; +val zmod_div_trivial = thm "zmod_div_trivial"; +val zmod_mod_trivial = thm "zmod_mod_trivial"; +val zmod_zadd_left_eq = thm "zmod_zadd_left_eq"; +val zmod_zadd_right_eq = thm "zmod_zadd_right_eq"; +val zdiv_zadd_self1 = thm "zdiv_zadd_self1"; +val zdiv_zadd_self2 = thm "zdiv_zadd_self2"; +val zmod_zadd_self1 = thm "zmod_zadd_self1"; +val zmod_zadd_self2 = thm "zmod_zadd_self2"; +val zdiv_zmult2_eq = thm "zdiv_zmult2_eq"; +val zmod_zmult2_eq = thm "zmod_zmult2_eq"; +val zdiv_zmult_zmult1 = thm "zdiv_zmult_zmult1"; +val zdiv_zmult_zmult2 = thm "zdiv_zmult_zmult2"; +val zmod_zmult_zmult1 = thm "zmod_zmult_zmult1"; +val zmod_zmult_zmult2 = thm "zmod_zmult_zmult2"; +val zdiv_neg_pos_less0 = thm "zdiv_neg_pos_less0"; +val zdiv_nonneg_neg_le0 = thm "zdiv_nonneg_neg_le0"; +val pos_imp_zdiv_nonneg_iff = thm "pos_imp_zdiv_nonneg_iff"; +val neg_imp_zdiv_nonneg_iff = thm "neg_imp_zdiv_nonneg_iff"; +val pos_imp_zdiv_neg_iff = thm "pos_imp_zdiv_neg_iff"; +val neg_imp_zdiv_neg_iff = thm "neg_imp_zdiv_neg_iff"; +*} + +end +
--- a/src/ZF/Integ/Bin.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,692 +0,0 @@ -(* Title: ZF/ex/Bin.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge - - The sign Pls stands for an infinite string of leading 0's. - The sign Min stands for an infinite string of leading 1's. - -A number can have multiple representations, namely leading 0's with sign -Pls and leading 1's with sign Min. See twos-compl.ML/int_of_binary for -the numerical interpretation. - -The representation expects that (m mod 2) is 0 or 1, even if m is negative; -For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 -*) - -header{*Arithmetic on Binary Integers*} - -theory Bin -imports Int Datatype -uses "Tools/numeral_syntax.ML" -begin - -consts bin :: i -datatype - "bin" = Pls - | Min - | Bit ("w: bin", "b: bool") (infixl "BIT" 90) - -syntax - "_Int" :: "xnum => i" ("_") - -consts - integ_of :: "i=>i" - NCons :: "[i,i]=>i" - bin_succ :: "i=>i" - bin_pred :: "i=>i" - bin_minus :: "i=>i" - bin_adder :: "i=>i" - bin_mult :: "[i,i]=>i" - -primrec - integ_of_Pls: "integ_of (Pls) = $# 0" - integ_of_Min: "integ_of (Min) = $-($#1)" - integ_of_BIT: "integ_of (w BIT b) = $#b $+ integ_of(w) $+ integ_of(w)" - - (** recall that cond(1,b,c)=b and cond(0,b,c)=0 **) - -primrec (*NCons adds a bit, suppressing leading 0s and 1s*) - NCons_Pls: "NCons (Pls,b) = cond(b,Pls BIT b,Pls)" - NCons_Min: "NCons (Min,b) = cond(b,Min,Min BIT b)" - NCons_BIT: "NCons (w BIT c,b) = w BIT c BIT b" - -primrec (*successor. If a BIT, can change a 0 to a 1 without recursion.*) - bin_succ_Pls: "bin_succ (Pls) = Pls BIT 1" - bin_succ_Min: "bin_succ (Min) = Pls" - bin_succ_BIT: "bin_succ (w BIT b) = cond(b, bin_succ(w) BIT 0, NCons(w,1))" - -primrec (*predecessor*) - bin_pred_Pls: "bin_pred (Pls) = Min" - bin_pred_Min: "bin_pred (Min) = Min BIT 0" - bin_pred_BIT: "bin_pred (w BIT b) = cond(b, NCons(w,0), bin_pred(w) BIT 1)" - -primrec (*unary negation*) - bin_minus_Pls: - "bin_minus (Pls) = Pls" - bin_minus_Min: - "bin_minus (Min) = Pls BIT 1" - bin_minus_BIT: - "bin_minus (w BIT b) = cond(b, bin_pred(NCons(bin_minus(w),0)), - bin_minus(w) BIT 0)" - -primrec (*sum*) - bin_adder_Pls: - "bin_adder (Pls) = (lam w:bin. w)" - bin_adder_Min: - "bin_adder (Min) = (lam w:bin. bin_pred(w))" - bin_adder_BIT: - "bin_adder (v BIT x) = - (lam w:bin. - bin_case (v BIT x, bin_pred(v BIT x), - %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w), - x xor y), - w))" - -(*The bin_case above replaces the following mutually recursive function: -primrec - "adding (v,x,Pls) = v BIT x" - "adding (v,x,Min) = bin_pred(v BIT x)" - "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)), - x xor y)" -*) - -constdefs - bin_add :: "[i,i]=>i" - "bin_add(v,w) == bin_adder(v)`w" - - -primrec - bin_mult_Pls: - "bin_mult (Pls,w) = Pls" - bin_mult_Min: - "bin_mult (Min,w) = bin_minus(w)" - bin_mult_BIT: - "bin_mult (v BIT b,w) = cond(b, bin_add(NCons(bin_mult(v,w),0),w), - NCons(bin_mult(v,w),0))" - -setup NumeralSyntax.setup - - -declare bin.intros [simp,TC] - -lemma NCons_Pls_0: "NCons(Pls,0) = Pls" -by simp - -lemma NCons_Pls_1: "NCons(Pls,1) = Pls BIT 1" -by simp - -lemma NCons_Min_0: "NCons(Min,0) = Min BIT 0" -by simp - -lemma NCons_Min_1: "NCons(Min,1) = Min" -by simp - -lemma NCons_BIT: "NCons(w BIT x,b) = w BIT x BIT b" -by (simp add: bin.case_eqns) - -lemmas NCons_simps [simp] = - NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT - - - -(** Type checking **) - -lemma integ_of_type [TC]: "w: bin ==> integ_of(w) : int" -apply (induct_tac "w") -apply (simp_all add: bool_into_nat) -done - -lemma NCons_type [TC]: "[| w: bin; b: bool |] ==> NCons(w,b) : bin" -by (induct_tac "w", auto) - -lemma bin_succ_type [TC]: "w: bin ==> bin_succ(w) : bin" -by (induct_tac "w", auto) - -lemma bin_pred_type [TC]: "w: bin ==> bin_pred(w) : bin" -by (induct_tac "w", auto) - -lemma bin_minus_type [TC]: "w: bin ==> bin_minus(w) : bin" -by (induct_tac "w", auto) - -(*This proof is complicated by the mutual recursion*) -lemma bin_add_type [rule_format,TC]: - "v: bin ==> ALL w: bin. bin_add(v,w) : bin" -apply (unfold bin_add_def) -apply (induct_tac "v") -apply (rule_tac [3] ballI) -apply (rename_tac [3] "w'") -apply (induct_tac [3] "w'") -apply (simp_all add: NCons_type) -done - -lemma bin_mult_type [TC]: "[| v: bin; w: bin |] ==> bin_mult(v,w) : bin" -by (induct_tac "v", auto) - - -subsubsection{*The Carry and Borrow Functions, - @{term bin_succ} and @{term bin_pred}*} - -(*NCons preserves the integer value of its argument*) -lemma integ_of_NCons [simp]: - "[| w: bin; b: bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)" -apply (erule bin.cases) -apply (auto elim!: boolE) -done - -lemma integ_of_succ [simp]: - "w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)" -apply (erule bin.induct) -apply (auto simp add: zadd_ac elim!: boolE) -done - -lemma integ_of_pred [simp]: - "w: bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)" -apply (erule bin.induct) -apply (auto simp add: zadd_ac elim!: boolE) -done - - -subsubsection{*@{term bin_minus}: Unary Negation of Binary Integers*} - -lemma integ_of_minus: "w: bin ==> integ_of(bin_minus(w)) = $- integ_of(w)" -apply (erule bin.induct) -apply (auto simp add: zadd_ac zminus_zadd_distrib elim!: boolE) -done - - -subsubsection{*@{term bin_add}: Binary Addition*} - -lemma bin_add_Pls [simp]: "w: bin ==> bin_add(Pls,w) = w" -by (unfold bin_add_def, simp) - -lemma bin_add_Pls_right: "w: bin ==> bin_add(w,Pls) = w" -apply (unfold bin_add_def) -apply (erule bin.induct, auto) -done - -lemma bin_add_Min [simp]: "w: bin ==> bin_add(Min,w) = bin_pred(w)" -by (unfold bin_add_def, simp) - -lemma bin_add_Min_right: "w: bin ==> bin_add(w,Min) = bin_pred(w)" -apply (unfold bin_add_def) -apply (erule bin.induct, auto) -done - -lemma bin_add_BIT_Pls [simp]: "bin_add(v BIT x,Pls) = v BIT x" -by (unfold bin_add_def, simp) - -lemma bin_add_BIT_Min [simp]: "bin_add(v BIT x,Min) = bin_pred(v BIT x)" -by (unfold bin_add_def, simp) - -lemma bin_add_BIT_BIT [simp]: - "[| w: bin; y: bool |] - ==> bin_add(v BIT x, w BIT y) = - NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)" -by (unfold bin_add_def, simp) - -lemma integ_of_add [rule_format]: - "v: bin ==> - ALL w: bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)" -apply (erule bin.induct, simp, simp) -apply (rule ballI) -apply (induct_tac "wa") -apply (auto simp add: zadd_ac elim!: boolE) -done - -(*Subtraction*) -lemma diff_integ_of_eq: - "[| v: bin; w: bin |] - ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))" -apply (unfold zdiff_def) -apply (simp add: integ_of_add integ_of_minus) -done - - -subsubsection{*@{term bin_mult}: Binary Multiplication*} - -lemma integ_of_mult: - "[| v: bin; w: bin |] - ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)" -apply (induct_tac "v", simp) -apply (simp add: integ_of_minus) -apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib elim!: boolE) -done - - -subsection{*Computations*} - -(** extra rules for bin_succ, bin_pred **) - -lemma bin_succ_1: "bin_succ(w BIT 1) = bin_succ(w) BIT 0" -by simp - -lemma bin_succ_0: "bin_succ(w BIT 0) = NCons(w,1)" -by simp - -lemma bin_pred_1: "bin_pred(w BIT 1) = NCons(w,0)" -by simp - -lemma bin_pred_0: "bin_pred(w BIT 0) = bin_pred(w) BIT 1" -by simp - -(** extra rules for bin_minus **) - -lemma bin_minus_1: "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))" -by simp - -lemma bin_minus_0: "bin_minus(w BIT 0) = bin_minus(w) BIT 0" -by simp - -(** extra rules for bin_add **) - -lemma bin_add_BIT_11: "w: bin ==> bin_add(v BIT 1, w BIT 1) = - NCons(bin_add(v, bin_succ(w)), 0)" -by simp - -lemma bin_add_BIT_10: "w: bin ==> bin_add(v BIT 1, w BIT 0) = - NCons(bin_add(v,w), 1)" -by simp - -lemma bin_add_BIT_0: "[| w: bin; y: bool |] - ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)" -by simp - -(** extra rules for bin_mult **) - -lemma bin_mult_1: "bin_mult(v BIT 1, w) = bin_add(NCons(bin_mult(v,w),0), w)" -by simp - -lemma bin_mult_0: "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)" -by simp - - -(** Simplification rules with integer constants **) - -lemma int_of_0: "$#0 = #0" -by simp - -lemma int_of_succ: "$# succ(n) = #1 $+ $#n" -by (simp add: int_of_add [symmetric] natify_succ) - -lemma zminus_0 [simp]: "$- #0 = #0" -by simp - -lemma zadd_0_intify [simp]: "#0 $+ z = intify(z)" -by simp - -lemma zadd_0_right_intify [simp]: "z $+ #0 = intify(z)" -by simp - -lemma zmult_1_intify [simp]: "#1 $* z = intify(z)" -by simp - -lemma zmult_1_right_intify [simp]: "z $* #1 = intify(z)" -by (subst zmult_commute, simp) - -lemma zmult_0 [simp]: "#0 $* z = #0" -by simp - -lemma zmult_0_right [simp]: "z $* #0 = #0" -by (subst zmult_commute, simp) - -lemma zmult_minus1 [simp]: "#-1 $* z = $-z" -by (simp add: zcompare_rls) - -lemma zmult_minus1_right [simp]: "z $* #-1 = $-z" -apply (subst zmult_commute) -apply (rule zmult_minus1) -done - - -subsection{*Simplification Rules for Comparison of Binary Numbers*} -text{*Thanks to Norbert Voelker*} - -(** Equals (=) **) - -lemma eq_integ_of_eq: - "[| v: bin; w: bin |] - ==> ((integ_of(v)) = integ_of(w)) <-> - iszero (integ_of (bin_add (v, bin_minus(w))))" -apply (unfold iszero_def) -apply (simp add: zcompare_rls integ_of_add integ_of_minus) -done - -lemma iszero_integ_of_Pls: "iszero (integ_of(Pls))" -by (unfold iszero_def, simp) - - -lemma nonzero_integ_of_Min: "~ iszero (integ_of(Min))" -apply (unfold iszero_def) -apply (simp add: zminus_equation) -done - -lemma iszero_integ_of_BIT: - "[| w: bin; x: bool |] - ==> iszero (integ_of (w BIT x)) <-> (x=0 & iszero (integ_of(w)))" -apply (unfold iszero_def, simp) -apply (subgoal_tac "integ_of (w) : int") -apply typecheck -apply (drule int_cases) -apply (safe elim!: boolE) -apply (simp_all (asm_lr) add: zcompare_rls zminus_zadd_distrib [symmetric] - int_of_add [symmetric]) -done - -lemma iszero_integ_of_0: - "w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))" -by (simp only: iszero_integ_of_BIT, blast) - -lemma iszero_integ_of_1: "w: bin ==> ~ iszero (integ_of (w BIT 1))" -by (simp only: iszero_integ_of_BIT, blast) - - - -(** Less-than (<) **) - -lemma less_integ_of_eq_neg: - "[| v: bin; w: bin |] - ==> integ_of(v) $< integ_of(w) - <-> znegative (integ_of (bin_add (v, bin_minus(w))))" -apply (unfold zless_def zdiff_def) -apply (simp add: integ_of_minus integ_of_add) -done - -lemma not_neg_integ_of_Pls: "~ znegative (integ_of(Pls))" -by simp - -lemma neg_integ_of_Min: "znegative (integ_of(Min))" -by simp - -lemma neg_integ_of_BIT: - "[| w: bin; x: bool |] - ==> znegative (integ_of (w BIT x)) <-> znegative (integ_of(w))" -apply simp -apply (subgoal_tac "integ_of (w) : int") -apply typecheck -apply (drule int_cases) -apply (auto elim!: boolE simp add: int_of_add [symmetric] zcompare_rls) -apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def - int_of_add [symmetric]) -apply (subgoal_tac "$#1 $- $# succ (succ (n #+ n)) = $- $# succ (n #+ n) ") - apply (simp add: zdiff_def) -apply (simp add: equation_zminus int_of_diff [symmetric]) -done - -(** Less-than-or-equals (<=) **) - -lemma le_integ_of_eq_not_less: - "(integ_of(x) $<= (integ_of(w))) <-> ~ (integ_of(w) $< (integ_of(x)))" -by (simp add: not_zless_iff_zle [THEN iff_sym]) - - -(*Delete the original rewrites, with their clumsy conditional expressions*) -declare bin_succ_BIT [simp del] - bin_pred_BIT [simp del] - bin_minus_BIT [simp del] - NCons_Pls [simp del] - NCons_Min [simp del] - bin_adder_BIT [simp del] - bin_mult_BIT [simp del] - -(*Hide the binary representation of integer constants*) -declare integ_of_Pls [simp del] integ_of_Min [simp del] integ_of_BIT [simp del] - - -lemmas bin_arith_extra_simps = - integ_of_add [symmetric] - integ_of_minus [symmetric] - integ_of_mult [symmetric] - bin_succ_1 bin_succ_0 - bin_pred_1 bin_pred_0 - bin_minus_1 bin_minus_0 - bin_add_Pls_right bin_add_Min_right - bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11 - diff_integ_of_eq - bin_mult_1 bin_mult_0 NCons_simps - - -(*For making a minimal simpset, one must include these default simprules - of thy. Also include simp_thms, or at least (~False)=True*) -lemmas bin_arith_simps = - bin_pred_Pls bin_pred_Min - bin_succ_Pls bin_succ_Min - bin_add_Pls bin_add_Min - bin_minus_Pls bin_minus_Min - bin_mult_Pls bin_mult_Min - bin_arith_extra_simps - -(*Simplification of relational operations*) -lemmas bin_rel_simps = - eq_integ_of_eq iszero_integ_of_Pls nonzero_integ_of_Min - iszero_integ_of_0 iszero_integ_of_1 - less_integ_of_eq_neg - not_neg_integ_of_Pls neg_integ_of_Min neg_integ_of_BIT - le_integ_of_eq_not_less - -declare bin_arith_simps [simp] -declare bin_rel_simps [simp] - - -(** Simplification of arithmetic when nested to the right **) - -lemma add_integ_of_left [simp]: - "[| v: bin; w: bin |] - ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)" -by (simp add: zadd_assoc [symmetric]) - -lemma mult_integ_of_left [simp]: - "[| v: bin; w: bin |] - ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)" -by (simp add: zmult_assoc [symmetric]) - -lemma add_integ_of_diff1 [simp]: - "[| v: bin; w: bin |] - ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)" -apply (unfold zdiff_def) -apply (rule add_integ_of_left, auto) -done - -lemma add_integ_of_diff2 [simp]: - "[| v: bin; w: bin |] - ==> integ_of(v) $+ (c $- integ_of(w)) = - integ_of (bin_add (v, bin_minus(w))) $+ (c)" -apply (subst diff_integ_of_eq [symmetric]) -apply (simp_all add: zdiff_def zadd_ac) -done - - -(** More for integer constants **) - -declare int_of_0 [simp] int_of_succ [simp] - -lemma zdiff0 [simp]: "#0 $- x = $-x" -by (simp add: zdiff_def) - -lemma zdiff0_right [simp]: "x $- #0 = intify(x)" -by (simp add: zdiff_def) - -lemma zdiff_self [simp]: "x $- x = #0" -by (simp add: zdiff_def) - -lemma znegative_iff_zless_0: "k: int ==> znegative(k) <-> k $< #0" -by (simp add: zless_def) - -lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k: int|] ==> znegative($-k)" -by (simp add: zless_def) - -lemma zero_zle_int_of [simp]: "#0 $<= $# n" -by (simp add: not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym]) - -lemma nat_of_0 [simp]: "nat_of(#0) = 0" -by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of) - -lemma nat_le_int0_lemma: "[| z $<= $#0; z: int |] ==> nat_of(z) = 0" -by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of) - -lemma nat_le_int0: "z $<= $#0 ==> nat_of(z) = 0" -apply (subgoal_tac "nat_of (intify (z)) = 0") -apply (rule_tac [2] nat_le_int0_lemma, auto) -done - -lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 ==> natify(n) = 0" -by (rule not_znegative_imp_zero, auto) - -lemma nat_of_zminus_int_of: "nat_of($- $# n) = 0" -by (simp add: nat_of_def int_of_def raw_nat_of zminus image_intrel_int) - -lemma int_of_nat_of: "#0 $<= z ==> $# nat_of(z) = intify(z)" -apply (rule not_zneg_nat_of_intify) -apply (simp add: znegative_iff_zless_0 not_zless_iff_zle) -done - -declare int_of_nat_of [simp] nat_of_zminus_int_of [simp] - -lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $<= z then intify(z) else #0)" -by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless) - -lemma zless_nat_iff_int_zless: "[| m: nat; z: int |] ==> (m < nat_of(z)) <-> ($#m $< z)" -apply (case_tac "znegative (z) ") -apply (erule_tac [2] not_zneg_nat_of [THEN subst]) -apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans] - simp add: znegative_iff_zless_0) -done - - -(** nat_of and zless **) - -(*An alternative condition is $#0 <= w *) -lemma zless_nat_conj_lemma: "$#0 $< z ==> (nat_of(w) < nat_of(z)) <-> (w $< z)" -apply (rule iff_trans) -apply (rule zless_int_of [THEN iff_sym]) -apply (auto simp add: int_of_nat_of_if simp del: zless_int_of) -apply (auto elim: zless_asym simp add: not_zle_iff_zless) -apply (blast intro: zless_zle_trans) -done - -lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) <-> ($#0 $< z & w $< z)" -apply (case_tac "$#0 $< z") -apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle) -done - -(*This simprule cannot be added unless we can find a way to make eq_integ_of_eq - unconditional! - [The condition "True" is a hack to prevent looping. - Conditional rewrite rules are tried after unconditional ones, so a rule - like eq_nat_number_of will be tried first to eliminate #mm=#nn.] - lemma integ_of_reorient [simp]: - "True ==> (integ_of(w) = x) <-> (x = integ_of(w))" - by auto -*) - -lemma integ_of_minus_reorient [simp]: - "(integ_of(w) = $- x) <-> ($- x = integ_of(w))" -by auto - -lemma integ_of_add_reorient [simp]: - "(integ_of(w) = x $+ y) <-> (x $+ y = integ_of(w))" -by auto - -lemma integ_of_diff_reorient [simp]: - "(integ_of(w) = x $- y) <-> (x $- y = integ_of(w))" -by auto - -lemma integ_of_mult_reorient [simp]: - "(integ_of(w) = x $* y) <-> (x $* y = integ_of(w))" -by auto - -ML -{* -val bin_pred_Pls = thm "bin_pred_Pls"; -val bin_pred_Min = thm "bin_pred_Min"; -val bin_minus_Pls = thm "bin_minus_Pls"; -val bin_minus_Min = thm "bin_minus_Min"; - -val NCons_Pls_0 = thm "NCons_Pls_0"; -val NCons_Pls_1 = thm "NCons_Pls_1"; -val NCons_Min_0 = thm "NCons_Min_0"; -val NCons_Min_1 = thm "NCons_Min_1"; -val NCons_BIT = thm "NCons_BIT"; -val NCons_simps = thms "NCons_simps"; -val integ_of_type = thm "integ_of_type"; -val NCons_type = thm "NCons_type"; -val bin_succ_type = thm "bin_succ_type"; -val bin_pred_type = thm "bin_pred_type"; -val bin_minus_type = thm "bin_minus_type"; -val bin_add_type = thm "bin_add_type"; -val bin_mult_type = thm "bin_mult_type"; -val integ_of_NCons = thm "integ_of_NCons"; -val integ_of_succ = thm "integ_of_succ"; -val integ_of_pred = thm "integ_of_pred"; -val integ_of_minus = thm "integ_of_minus"; -val bin_add_Pls = thm "bin_add_Pls"; -val bin_add_Pls_right = thm "bin_add_Pls_right"; -val bin_add_Min = thm "bin_add_Min"; -val bin_add_Min_right = thm "bin_add_Min_right"; -val bin_add_BIT_Pls = thm "bin_add_BIT_Pls"; -val bin_add_BIT_Min = thm "bin_add_BIT_Min"; -val bin_add_BIT_BIT = thm "bin_add_BIT_BIT"; -val integ_of_add = thm "integ_of_add"; -val diff_integ_of_eq = thm "diff_integ_of_eq"; -val integ_of_mult = thm "integ_of_mult"; -val bin_succ_1 = thm "bin_succ_1"; -val bin_succ_0 = thm "bin_succ_0"; -val bin_pred_1 = thm "bin_pred_1"; -val bin_pred_0 = thm "bin_pred_0"; -val bin_minus_1 = thm "bin_minus_1"; -val bin_minus_0 = thm "bin_minus_0"; -val bin_add_BIT_11 = thm "bin_add_BIT_11"; -val bin_add_BIT_10 = thm "bin_add_BIT_10"; -val bin_add_BIT_0 = thm "bin_add_BIT_0"; -val bin_mult_1 = thm "bin_mult_1"; -val bin_mult_0 = thm "bin_mult_0"; -val int_of_0 = thm "int_of_0"; -val int_of_succ = thm "int_of_succ"; -val zminus_0 = thm "zminus_0"; -val zadd_0_intify = thm "zadd_0_intify"; -val zadd_0_right_intify = thm "zadd_0_right_intify"; -val zmult_1_intify = thm "zmult_1_intify"; -val zmult_1_right_intify = thm "zmult_1_right_intify"; -val zmult_0 = thm "zmult_0"; -val zmult_0_right = thm "zmult_0_right"; -val zmult_minus1 = thm "zmult_minus1"; -val zmult_minus1_right = thm "zmult_minus1_right"; -val eq_integ_of_eq = thm "eq_integ_of_eq"; -val iszero_integ_of_Pls = thm "iszero_integ_of_Pls"; -val nonzero_integ_of_Min = thm "nonzero_integ_of_Min"; -val iszero_integ_of_BIT = thm "iszero_integ_of_BIT"; -val iszero_integ_of_0 = thm "iszero_integ_of_0"; -val iszero_integ_of_1 = thm "iszero_integ_of_1"; -val less_integ_of_eq_neg = thm "less_integ_of_eq_neg"; -val not_neg_integ_of_Pls = thm "not_neg_integ_of_Pls"; -val neg_integ_of_Min = thm "neg_integ_of_Min"; -val neg_integ_of_BIT = thm "neg_integ_of_BIT"; -val le_integ_of_eq_not_less = thm "le_integ_of_eq_not_less"; -val bin_arith_extra_simps = thms "bin_arith_extra_simps"; -val bin_arith_simps = thms "bin_arith_simps"; -val bin_rel_simps = thms "bin_rel_simps"; -val add_integ_of_left = thm "add_integ_of_left"; -val mult_integ_of_left = thm "mult_integ_of_left"; -val add_integ_of_diff1 = thm "add_integ_of_diff1"; -val add_integ_of_diff2 = thm "add_integ_of_diff2"; -val zdiff0 = thm "zdiff0"; -val zdiff0_right = thm "zdiff0_right"; -val zdiff_self = thm "zdiff_self"; -val znegative_iff_zless_0 = thm "znegative_iff_zless_0"; -val zero_zless_imp_znegative_zminus = thm "zero_zless_imp_znegative_zminus"; -val zero_zle_int_of = thm "zero_zle_int_of"; -val nat_of_0 = thm "nat_of_0"; -val nat_le_int0 = thm "nat_le_int0"; -val int_of_eq_0_imp_natify_eq_0 = thm "int_of_eq_0_imp_natify_eq_0"; -val nat_of_zminus_int_of = thm "nat_of_zminus_int_of"; -val int_of_nat_of = thm "int_of_nat_of"; -val int_of_nat_of_if = thm "int_of_nat_of_if"; -val zless_nat_iff_int_zless = thm "zless_nat_iff_int_zless"; -val zless_nat_conj = thm "zless_nat_conj"; -val integ_of_minus_reorient = thm "integ_of_minus_reorient"; -val integ_of_add_reorient = thm "integ_of_add_reorient"; -val integ_of_diff_reorient = thm "integ_of_diff_reorient"; -val integ_of_mult_reorient = thm "integ_of_mult_reorient"; -*} - -end
--- a/src/ZF/Integ/EquivClass.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,265 +0,0 @@ -(* Title: ZF/EquivClass.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1994 University of Cambridge - -*) - -header{*Equivalence Relations*} - -theory EquivClass imports Trancl Perm begin - -constdefs - - quotient :: "[i,i]=>i" (infixl "'/'/" 90) (*set of equiv classes*) - "A//r == {r``{x} . x:A}" - - congruent :: "[i,i=>i]=>o" - "congruent(r,b) == ALL y z. <y,z>:r --> b(y)=b(z)" - - congruent2 :: "[i,i,[i,i]=>i]=>o" - "congruent2(r1,r2,b) == ALL y1 z1 y2 z2. - <y1,z1>:r1 --> <y2,z2>:r2 --> b(y1,y2) = b(z1,z2)" - -syntax - RESPECTS ::"[i=>i, i] => o" (infixr "respects" 80) - RESPECTS2 ::"[i=>i, i] => o" (infixr "respects2 " 80) - --{*Abbreviation for the common case where the relations are identical*} - -translations - "f respects r" == "congruent(r,f)" - "f respects2 r" => "congruent2(r,r,f)" - -subsection{*Suppes, Theorem 70: - @{term r} is an equiv relation iff @{term "converse(r) O r = r"}*} - -(** first half: equiv(A,r) ==> converse(r) O r = r **) - -lemma sym_trans_comp_subset: - "[| sym(r); trans(r) |] ==> converse(r) O r <= r" -by (unfold trans_def sym_def, blast) - -lemma refl_comp_subset: - "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r" -by (unfold refl_def, blast) - -lemma equiv_comp_eq: - "equiv(A,r) ==> converse(r) O r = r" -apply (unfold equiv_def) -apply (blast del: subsetI intro!: sym_trans_comp_subset refl_comp_subset) -done - -(*second half*) -lemma comp_equivI: - "[| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)" -apply (unfold equiv_def refl_def sym_def trans_def) -apply (erule equalityE) -apply (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r", blast+) -done - -(** Equivalence classes **) - -(*Lemma for the next result*) -lemma equiv_class_subset: - "[| sym(r); trans(r); <a,b>: r |] ==> r``{a} <= r``{b}" -by (unfold trans_def sym_def, blast) - -lemma equiv_class_eq: - "[| equiv(A,r); <a,b>: r |] ==> r``{a} = r``{b}" -apply (unfold equiv_def) -apply (safe del: subsetI intro!: equalityI equiv_class_subset) -apply (unfold sym_def, blast) -done - -lemma equiv_class_self: - "[| equiv(A,r); a: A |] ==> a: r``{a}" -by (unfold equiv_def refl_def, blast) - -(*Lemma for the next result*) -lemma subset_equiv_class: - "[| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> <a,b>: r" -by (unfold equiv_def refl_def, blast) - -lemma eq_equiv_class: "[| r``{a} = r``{b}; equiv(A,r); b: A |] ==> <a,b>: r" -by (assumption | rule equalityD2 subset_equiv_class)+ - -(*thus r``{a} = r``{b} as well*) -lemma equiv_class_nondisjoint: - "[| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> <a,b>: r" -by (unfold equiv_def trans_def sym_def, blast) - -lemma equiv_type: "equiv(A,r) ==> r <= A*A" -by (unfold equiv_def, blast) - -lemma equiv_class_eq_iff: - "equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A" -by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) - -lemma eq_equiv_class_iff: - "[| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> <x,y>: r" -by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type) - -(*** Quotients ***) - -(** Introduction/elimination rules -- needed? **) - -lemma quotientI [TC]: "x:A ==> r``{x}: A//r" -apply (unfold quotient_def) -apply (erule RepFunI) -done - -lemma quotientE: - "[| X: A//r; !!x. [| X = r``{x}; x:A |] ==> P |] ==> P" -by (unfold quotient_def, blast) - -lemma Union_quotient: - "equiv(A,r) ==> Union(A//r) = A" -by (unfold equiv_def refl_def quotient_def, blast) - -lemma quotient_disj: - "[| equiv(A,r); X: A//r; Y: A//r |] ==> X=Y | (X Int Y <= 0)" -apply (unfold quotient_def) -apply (safe intro!: equiv_class_eq, assumption) -apply (unfold equiv_def trans_def sym_def, blast) -done - -subsection{*Defining Unary Operations upon Equivalence Classes*} - -(** Could have a locale with the premises equiv(A,r) and congruent(r,b) -**) - -(*Conversion rule*) -lemma UN_equiv_class: - "[| equiv(A,r); b respects r; a: A |] ==> (UN x:r``{a}. b(x)) = b(a)" -apply (subgoal_tac "\<forall>x \<in> r``{a}. b(x) = b(a)") - apply simp - apply (blast intro: equiv_class_self) -apply (unfold equiv_def sym_def congruent_def, blast) -done - -(*type checking of UN x:r``{a}. b(x) *) -lemma UN_equiv_class_type: - "[| equiv(A,r); b respects r; X: A//r; !!x. x : A ==> b(x) : B |] - ==> (UN x:X. b(x)) : B" -apply (unfold quotient_def, safe) -apply (simp (no_asm_simp) add: UN_equiv_class) -done - -(*Sufficient conditions for injectiveness. Could weaken premises! - major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B -*) -lemma UN_equiv_class_inject: - "[| equiv(A,r); b respects r; - (UN x:X. b(x))=(UN y:Y. b(y)); X: A//r; Y: A//r; - !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |] - ==> X=Y" -apply (unfold quotient_def, safe) -apply (rule equiv_class_eq, assumption) -apply (simp add: UN_equiv_class [of A r b]) -done - - -subsection{*Defining Binary Operations upon Equivalence Classes*} - -lemma congruent2_implies_congruent: - "[| equiv(A,r1); congruent2(r1,r2,b); a: A |] ==> congruent(r2,b(a))" -by (unfold congruent_def congruent2_def equiv_def refl_def, blast) - -lemma congruent2_implies_congruent_UN: - "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a: A2 |] ==> - congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))" -apply (unfold congruent_def, safe) -apply (frule equiv_type [THEN subsetD], assumption) -apply clarify -apply (simp add: UN_equiv_class congruent2_implies_congruent) -apply (unfold congruent2_def equiv_def refl_def, blast) -done - -lemma UN_equiv_class2: - "[| equiv(A1,r1); equiv(A2,r2); congruent2(r1,r2,b); a1: A1; a2: A2 |] - ==> (\<Union>x1 \<in> r1``{a1}. \<Union>x2 \<in> r2``{a2}. b(x1,x2)) = b(a1,a2)" -by (simp add: UN_equiv_class congruent2_implies_congruent - congruent2_implies_congruent_UN) - -(*type checking*) -lemma UN_equiv_class_type2: - "[| equiv(A,r); b respects2 r; - X1: A//r; X2: A//r; - !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B - |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B" -apply (unfold quotient_def, safe) -apply (blast intro: UN_equiv_class_type congruent2_implies_congruent_UN - congruent2_implies_congruent quotientI) -done - - -(*Suggested by John Harrison -- the two subproofs may be MUCH simpler - than the direct proof*) -lemma congruent2I: - "[| equiv(A1,r1); equiv(A2,r2); - !! y z w. [| w \<in> A2; <y,z> \<in> r1 |] ==> b(y,w) = b(z,w); - !! y z w. [| w \<in> A1; <y,z> \<in> r2 |] ==> b(w,y) = b(w,z) - |] ==> congruent2(r1,r2,b)" -apply (unfold congruent2_def equiv_def refl_def, safe) -apply (blast intro: trans) -done - -lemma congruent2_commuteI: - assumes equivA: "equiv(A,r)" - and commute: "!! y z. [| y: A; z: A |] ==> b(y,z) = b(z,y)" - and congt: "!! y z w. [| w: A; <y,z>: r |] ==> b(w,y) = b(w,z)" - shows "b respects2 r" -apply (insert equivA [THEN equiv_type, THEN subsetD]) -apply (rule congruent2I [OF equivA equivA]) -apply (rule commute [THEN trans]) -apply (rule_tac [3] commute [THEN trans, symmetric]) -apply (rule_tac [5] sym) -apply (blast intro: congt)+ -done - -(*Obsolete?*) -lemma congruent_commuteI: - "[| equiv(A,r); Z: A//r; - !!w. [| w: A |] ==> congruent(r, %z. b(w,z)); - !!x y. [| x: A; y: A |] ==> b(y,x) = b(x,y) - |] ==> congruent(r, %w. UN z: Z. b(w,z))" -apply (simp (no_asm) add: congruent_def) -apply (safe elim!: quotientE) -apply (frule equiv_type [THEN subsetD], assumption) -apply (simp add: UN_equiv_class [of A r]) -apply (simp add: congruent_def) -done - -ML -{* -val sym_trans_comp_subset = thm "sym_trans_comp_subset"; -val refl_comp_subset = thm "refl_comp_subset"; -val equiv_comp_eq = thm "equiv_comp_eq"; -val comp_equivI = thm "comp_equivI"; -val equiv_class_subset = thm "equiv_class_subset"; -val equiv_class_eq = thm "equiv_class_eq"; -val equiv_class_self = thm "equiv_class_self"; -val subset_equiv_class = thm "subset_equiv_class"; -val eq_equiv_class = thm "eq_equiv_class"; -val equiv_class_nondisjoint = thm "equiv_class_nondisjoint"; -val equiv_type = thm "equiv_type"; -val equiv_class_eq_iff = thm "equiv_class_eq_iff"; -val eq_equiv_class_iff = thm "eq_equiv_class_iff"; -val quotientI = thm "quotientI"; -val quotientE = thm "quotientE"; -val Union_quotient = thm "Union_quotient"; -val quotient_disj = thm "quotient_disj"; -val UN_equiv_class = thm "UN_equiv_class"; -val UN_equiv_class_type = thm "UN_equiv_class_type"; -val UN_equiv_class_inject = thm "UN_equiv_class_inject"; -val congruent2_implies_congruent = thm "congruent2_implies_congruent"; -val congruent2_implies_congruent_UN = thm "congruent2_implies_congruent_UN"; -val congruent_commuteI = thm "congruent_commuteI"; -val UN_equiv_class2 = thm "UN_equiv_class2"; -val UN_equiv_class_type2 = thm "UN_equiv_class_type2"; -val congruent2I = thm "congruent2I"; -val congruent2_commuteI = thm "congruent2_commuteI"; -val congruent_commuteI = thm "congruent_commuteI"; -*} - -end
--- a/src/ZF/Integ/Int.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1057 +0,0 @@ -(* Title: ZF/Integ/Int.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge - -*) - -header{*The Integers as Equivalence Classes Over Pairs of Natural Numbers*} - -theory Int imports EquivClass ArithSimp begin - -constdefs - intrel :: i - "intrel == {p : (nat*nat)*(nat*nat). - \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1}" - - int :: i - "int == (nat*nat)//intrel" - - int_of :: "i=>i" --{*coercion from nat to int*} ("$# _" [80] 80) - "$# m == intrel `` {<natify(m), 0>}" - - intify :: "i=>i" --{*coercion from ANYTHING to int*} - "intify(m) == if m : int then m else $#0" - - raw_zminus :: "i=>i" - "raw_zminus(z) == \<Union><x,y>\<in>z. intrel``{<y,x>}" - - zminus :: "i=>i" ("$- _" [80] 80) - "$- z == raw_zminus (intify(z))" - - znegative :: "i=>o" - "znegative(z) == \<exists>x y. x<y & y\<in>nat & <x,y>\<in>z" - - iszero :: "i=>o" - "iszero(z) == z = $# 0" - - raw_nat_of :: "i=>i" - "raw_nat_of(z) == natify (\<Union><x,y>\<in>z. x#-y)" - - nat_of :: "i=>i" - "nat_of(z) == raw_nat_of (intify(z))" - - zmagnitude :: "i=>i" - --{*could be replaced by an absolute value function from int to int?*} - "zmagnitude(z) == - THE m. m\<in>nat & ((~ znegative(z) & z = $# m) | - (znegative(z) & $- z = $# m))" - - raw_zmult :: "[i,i]=>i" - (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2. - Perhaps a "curried" or even polymorphic congruent predicate would be - better.*) - "raw_zmult(z1,z2) == - \<Union>p1\<in>z1. \<Union>p2\<in>z2. split(%x1 y1. split(%x2 y2. - intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)" - - zmult :: "[i,i]=>i" (infixl "$*" 70) - "z1 $* z2 == raw_zmult (intify(z1),intify(z2))" - - raw_zadd :: "[i,i]=>i" - "raw_zadd (z1, z2) == - \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2 - in intrel``{<x1#+x2, y1#+y2>}" - - zadd :: "[i,i]=>i" (infixl "$+" 65) - "z1 $+ z2 == raw_zadd (intify(z1),intify(z2))" - - zdiff :: "[i,i]=>i" (infixl "$-" 65) - "z1 $- z2 == z1 $+ zminus(z2)" - - zless :: "[i,i]=>o" (infixl "$<" 50) - "z1 $< z2 == znegative(z1 $- z2)" - - zle :: "[i,i]=>o" (infixl "$<=" 50) - "z1 $<= z2 == z1 $< z2 | intify(z1)=intify(z2)" - - -syntax (xsymbols) - zmult :: "[i,i]=>i" (infixl "$\<times>" 70) - zle :: "[i,i]=>o" (infixl "$\<le>" 50) --{*less than or equals*} - -syntax (HTML output) - zmult :: "[i,i]=>i" (infixl "$\<times>" 70) - zle :: "[i,i]=>o" (infixl "$\<le>" 50) - - -declare quotientE [elim!] - -subsection{*Proving that @{term intrel} is an equivalence relation*} - -(** Natural deduction for intrel **) - -lemma intrel_iff [simp]: - "<<x1,y1>,<x2,y2>>: intrel <-> - x1\<in>nat & y1\<in>nat & x2\<in>nat & y2\<in>nat & x1#+y2 = x2#+y1" -by (simp add: intrel_def) - -lemma intrelI [intro!]: - "[| x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> <<x1,y1>,<x2,y2>>: intrel" -by (simp add: intrel_def) - -lemma intrelE [elim!]: - "[| p: intrel; - !!x1 y1 x2 y2. [| p = <<x1,y1>,<x2,y2>>; x1#+y2 = x2#+y1; - x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] ==> Q |] - ==> Q" -by (simp add: intrel_def, blast) - -lemma int_trans_lemma: - "[| x1 #+ y2 = x2 #+ y1; x2 #+ y3 = x3 #+ y2 |] ==> x1 #+ y3 = x3 #+ y1" -apply (rule sym) -apply (erule add_left_cancel)+ -apply (simp_all (no_asm_simp)) -done - -lemma equiv_intrel: "equiv(nat*nat, intrel)" -apply (simp add: equiv_def refl_def sym_def trans_def) -apply (fast elim!: sym int_trans_lemma) -done - -lemma image_intrel_int: "[| m\<in>nat; n\<in>nat |] ==> intrel `` {<m,n>} : int" -by (simp add: int_def) - -declare equiv_intrel [THEN eq_equiv_class_iff, simp] -declare conj_cong [cong] - -lemmas eq_intrelD = eq_equiv_class [OF _ equiv_intrel] - -(** int_of: the injection from nat to int **) - -lemma int_of_type [simp,TC]: "$#m : int" -by (simp add: int_def quotient_def int_of_def, auto) - -lemma int_of_eq [iff]: "($# m = $# n) <-> natify(m)=natify(n)" -by (simp add: int_of_def) - -lemma int_of_inject: "[| $#m = $#n; m\<in>nat; n\<in>nat |] ==> m=n" -by (drule int_of_eq [THEN iffD1], auto) - - -(** intify: coercion from anything to int **) - -lemma intify_in_int [iff,TC]: "intify(x) : int" -by (simp add: intify_def) - -lemma intify_ident [simp]: "n : int ==> intify(n) = n" -by (simp add: intify_def) - - -subsection{*Collapsing rules: to remove @{term intify} - from arithmetic expressions*} - -lemma intify_idem [simp]: "intify(intify(x)) = intify(x)" -by simp - -lemma int_of_natify [simp]: "$# (natify(m)) = $# m" -by (simp add: int_of_def) - -lemma zminus_intify [simp]: "$- (intify(m)) = $- m" -by (simp add: zminus_def) - -(** Addition **) - -lemma zadd_intify1 [simp]: "intify(x) $+ y = x $+ y" -by (simp add: zadd_def) - -lemma zadd_intify2 [simp]: "x $+ intify(y) = x $+ y" -by (simp add: zadd_def) - -(** Subtraction **) - -lemma zdiff_intify1 [simp]:"intify(x) $- y = x $- y" -by (simp add: zdiff_def) - -lemma zdiff_intify2 [simp]:"x $- intify(y) = x $- y" -by (simp add: zdiff_def) - -(** Multiplication **) - -lemma zmult_intify1 [simp]:"intify(x) $* y = x $* y" -by (simp add: zmult_def) - -lemma zmult_intify2 [simp]:"x $* intify(y) = x $* y" -by (simp add: zmult_def) - -(** Orderings **) - -lemma zless_intify1 [simp]:"intify(x) $< y <-> x $< y" -by (simp add: zless_def) - -lemma zless_intify2 [simp]:"x $< intify(y) <-> x $< y" -by (simp add: zless_def) - -lemma zle_intify1 [simp]:"intify(x) $<= y <-> x $<= y" -by (simp add: zle_def) - -lemma zle_intify2 [simp]:"x $<= intify(y) <-> x $<= y" -by (simp add: zle_def) - - -subsection{*@{term zminus}: unary negation on @{term int}*} - -lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel" -by (auto simp add: congruent_def add_ac) - -lemma raw_zminus_type: "z : int ==> raw_zminus(z) : int" -apply (simp add: int_def raw_zminus_def) -apply (typecheck add: UN_equiv_class_type [OF equiv_intrel zminus_congruent]) -done - -lemma zminus_type [TC,iff]: "$-z : int" -by (simp add: zminus_def raw_zminus_type) - -lemma raw_zminus_inject: - "[| raw_zminus(z) = raw_zminus(w); z: int; w: int |] ==> z=w" -apply (simp add: int_def raw_zminus_def) -apply (erule UN_equiv_class_inject [OF equiv_intrel zminus_congruent], safe) -apply (auto dest: eq_intrelD simp add: add_ac) -done - -lemma zminus_inject_intify [dest!]: "$-z = $-w ==> intify(z) = intify(w)" -apply (simp add: zminus_def) -apply (blast dest!: raw_zminus_inject) -done - -lemma zminus_inject: "[| $-z = $-w; z: int; w: int |] ==> z=w" -by auto - -lemma raw_zminus: - "[| x\<in>nat; y\<in>nat |] ==> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}" -apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent]) -done - -lemma zminus: - "[| x\<in>nat; y\<in>nat |] - ==> $- (intrel``{<x,y>}) = intrel `` {<y,x>}" -by (simp add: zminus_def raw_zminus image_intrel_int) - -lemma raw_zminus_zminus: "z : int ==> raw_zminus (raw_zminus(z)) = z" -by (auto simp add: int_def raw_zminus) - -lemma zminus_zminus_intify [simp]: "$- ($- z) = intify(z)" -by (simp add: zminus_def raw_zminus_type raw_zminus_zminus) - -lemma zminus_int0 [simp]: "$- ($#0) = $#0" -by (simp add: int_of_def zminus) - -lemma zminus_zminus: "z : int ==> $- ($- z) = z" -by simp - - -subsection{*@{term znegative}: the test for negative integers*} - -lemma znegative: "[| x\<in>nat; y\<in>nat |] ==> znegative(intrel``{<x,y>}) <-> x<y" -apply (cases "x<y") -apply (auto simp add: znegative_def not_lt_iff_le) -apply (subgoal_tac "y #+ x2 < x #+ y2", force) -apply (rule add_le_lt_mono, auto) -done - -(*No natural number is negative!*) -lemma not_znegative_int_of [iff]: "~ znegative($# n)" -by (simp add: znegative int_of_def) - -lemma znegative_zminus_int_of [simp]: "znegative($- $# succ(n))" -by (simp add: znegative int_of_def zminus natify_succ) - -lemma not_znegative_imp_zero: "~ znegative($- $# n) ==> natify(n)=0" -by (simp add: znegative int_of_def zminus Ord_0_lt_iff [THEN iff_sym]) - - -subsection{*@{term nat_of}: Coercion of an Integer to a Natural Number*} - -lemma nat_of_intify [simp]: "nat_of(intify(z)) = nat_of(z)" -by (simp add: nat_of_def) - -lemma nat_of_congruent: "(\<lambda>x. (\<lambda>\<langle>x,y\<rangle>. x #- y)(x)) respects intrel" -by (auto simp add: congruent_def split add: nat_diff_split) - -lemma raw_nat_of: - "[| x\<in>nat; y\<in>nat |] ==> raw_nat_of(intrel``{<x,y>}) = x#-y" -by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent]) - -lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)" -by (simp add: int_of_def raw_nat_of) - -lemma nat_of_int_of [simp]: "nat_of($# n) = natify(n)" -by (simp add: raw_nat_of_int_of nat_of_def) - -lemma raw_nat_of_type: "raw_nat_of(z) \<in> nat" -by (simp add: raw_nat_of_def) - -lemma nat_of_type [iff,TC]: "nat_of(z) \<in> nat" -by (simp add: nat_of_def raw_nat_of_type) - -subsection{*zmagnitude: magnitide of an integer, as a natural number*} - -lemma zmagnitude_int_of [simp]: "zmagnitude($# n) = natify(n)" -by (auto simp add: zmagnitude_def int_of_eq) - -lemma natify_int_of_eq: "natify(x)=n ==> $#x = $# n" -apply (drule sym) -apply (simp (no_asm_simp) add: int_of_eq) -done - -lemma zmagnitude_zminus_int_of [simp]: "zmagnitude($- $# n) = natify(n)" -apply (simp add: zmagnitude_def) -apply (rule the_equality) -apply (auto dest!: not_znegative_imp_zero natify_int_of_eq - iff del: int_of_eq, auto) -done - -lemma zmagnitude_type [iff,TC]: "zmagnitude(z)\<in>nat" -apply (simp add: zmagnitude_def) -apply (rule theI2, auto) -done - -lemma not_zneg_int_of: - "[| z: int; ~ znegative(z) |] ==> \<exists>n\<in>nat. z = $# n" -apply (auto simp add: int_def znegative int_of_def not_lt_iff_le) -apply (rename_tac x y) -apply (rule_tac x="x#-y" in bexI) -apply (auto simp add: add_diff_inverse2) -done - -lemma not_zneg_mag [simp]: - "[| z: int; ~ znegative(z) |] ==> $# (zmagnitude(z)) = z" -by (drule not_zneg_int_of, auto) - -lemma zneg_int_of: - "[| znegative(z); z: int |] ==> \<exists>n\<in>nat. z = $- ($# succ(n))" -by (auto simp add: int_def znegative zminus int_of_def dest!: less_imp_succ_add) - -lemma zneg_mag [simp]: - "[| znegative(z); z: int |] ==> $# (zmagnitude(z)) = $- z" -by (drule zneg_int_of, auto) - -lemma int_cases: "z : int ==> \<exists>n\<in>nat. z = $# n | z = $- ($# succ(n))" -apply (case_tac "znegative (z) ") -prefer 2 apply (blast dest: not_zneg_mag sym) -apply (blast dest: zneg_int_of) -done - -lemma not_zneg_raw_nat_of: - "[| ~ znegative(z); z: int |] ==> $# (raw_nat_of(z)) = z" -apply (drule not_zneg_int_of) -apply (auto simp add: raw_nat_of_type raw_nat_of_int_of) -done - -lemma not_zneg_nat_of_intify: - "~ znegative(intify(z)) ==> $# (nat_of(z)) = intify(z)" -by (simp (no_asm_simp) add: nat_of_def not_zneg_raw_nat_of) - -lemma not_zneg_nat_of: "[| ~ znegative(z); z: int |] ==> $# (nat_of(z)) = z" -apply (simp (no_asm_simp) add: not_zneg_nat_of_intify) -done - -lemma zneg_nat_of [simp]: "znegative(intify(z)) ==> nat_of(z) = 0" -apply (subgoal_tac "intify(z) \<in> int") -apply (simp add: int_def) -apply (auto simp add: znegative nat_of_def raw_nat_of - split add: nat_diff_split) -done - - -subsection{*@{term zadd}: addition on int*} - -text{*Congruence Property for Addition*} -lemma zadd_congruent2: - "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2 - in intrel``{<x1#+x2, y1#+y2>}) - respects2 intrel" -apply (simp add: congruent2_def) -(*Proof via congruent2_commuteI seems longer*) -apply safe -apply (simp (no_asm_simp) add: add_assoc Let_def) -(*The rest should be trivial, but rearranging terms is hard - add_ac does not help rewriting with the assumptions.*) -apply (rule_tac m1 = x1a in add_left_commute [THEN ssubst]) -apply (rule_tac m1 = x2a in add_left_commute [THEN ssubst]) -apply (simp (no_asm_simp) add: add_assoc [symmetric]) -done - -lemma raw_zadd_type: "[| z: int; w: int |] ==> raw_zadd(z,w) : int" -apply (simp add: int_def raw_zadd_def) -apply (rule UN_equiv_class_type2 [OF equiv_intrel zadd_congruent2], assumption+) -apply (simp add: Let_def) -done - -lemma zadd_type [iff,TC]: "z $+ w : int" -by (simp add: zadd_def raw_zadd_type) - -lemma raw_zadd: - "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) = - intrel `` {<x1#+x2, y1#+y2>}" -apply (simp add: raw_zadd_def - UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2]) -apply (simp add: Let_def) -done - -lemma zadd: - "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = - intrel `` {<x1#+x2, y1#+y2>}" -by (simp add: zadd_def raw_zadd image_intrel_int) - -lemma raw_zadd_int0: "z : int ==> raw_zadd ($#0,z) = z" -by (auto simp add: int_def int_of_def raw_zadd) - -lemma zadd_int0_intify [simp]: "$#0 $+ z = intify(z)" -by (simp add: zadd_def raw_zadd_int0) - -lemma zadd_int0: "z: int ==> $#0 $+ z = z" -by simp - -lemma raw_zminus_zadd_distrib: - "[| z: int; w: int |] ==> $- raw_zadd(z,w) = raw_zadd($- z, $- w)" -by (auto simp add: zminus raw_zadd int_def) - -lemma zminus_zadd_distrib [simp]: "$- (z $+ w) = $- z $+ $- w" -by (simp add: zadd_def raw_zminus_zadd_distrib) - -lemma raw_zadd_commute: - "[| z: int; w: int |] ==> raw_zadd(z,w) = raw_zadd(w,z)" -by (auto simp add: raw_zadd add_ac int_def) - -lemma zadd_commute: "z $+ w = w $+ z" -by (simp add: zadd_def raw_zadd_commute) - -lemma raw_zadd_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zadd (raw_zadd(z1,z2),z3) = raw_zadd(z1,raw_zadd(z2,z3))" -by (auto simp add: int_def raw_zadd add_assoc) - -lemma zadd_assoc: "(z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)" -by (simp add: zadd_def raw_zadd_type raw_zadd_assoc) - -(*For AC rewriting*) -lemma zadd_left_commute: "z1$+(z2$+z3) = z2$+(z1$+z3)" -apply (simp add: zadd_assoc [symmetric]) -apply (simp add: zadd_commute) -done - -(*Integer addition is an AC operator*) -lemmas zadd_ac = zadd_assoc zadd_commute zadd_left_commute - -lemma int_of_add: "$# (m #+ n) = ($#m) $+ ($#n)" -by (simp add: int_of_def zadd) - -lemma int_succ_int_1: "$# succ(m) = $# 1 $+ ($# m)" -by (simp add: int_of_add [symmetric] natify_succ) - -lemma int_of_diff: - "[| m\<in>nat; n le m |] ==> $# (m #- n) = ($#m) $- ($#n)" -apply (simp add: int_of_def zdiff_def) -apply (frule lt_nat_in_nat) -apply (simp_all add: zadd zminus add_diff_inverse2) -done - -lemma raw_zadd_zminus_inverse: "z : int ==> raw_zadd (z, $- z) = $#0" -by (auto simp add: int_def int_of_def zminus raw_zadd add_commute) - -lemma zadd_zminus_inverse [simp]: "z $+ ($- z) = $#0" -apply (simp add: zadd_def) -apply (subst zminus_intify [symmetric]) -apply (rule intify_in_int [THEN raw_zadd_zminus_inverse]) -done - -lemma zadd_zminus_inverse2 [simp]: "($- z) $+ z = $#0" -by (simp add: zadd_commute zadd_zminus_inverse) - -lemma zadd_int0_right_intify [simp]: "z $+ $#0 = intify(z)" -by (rule trans [OF zadd_commute zadd_int0_intify]) - -lemma zadd_int0_right: "z:int ==> z $+ $#0 = z" -by simp - - -subsection{*@{term zmult}: Integer Multiplication*} - -text{*Congruence property for multiplication*} -lemma zmult_congruent2: - "(%p1 p2. split(%x1 y1. split(%x2 y2. - intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)) - respects2 intrel" -apply (rule equiv_intrel [THEN congruent2_commuteI], auto) -(*Proof that zmult is congruent in one argument*) -apply (rename_tac x y) -apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context]) -apply (drule_tac t = "%u. y#*u" in subst_context) -apply (erule add_left_cancel)+ -apply (simp_all add: add_mult_distrib_left) -done - - -lemma raw_zmult_type: "[| z: int; w: int |] ==> raw_zmult(z,w) : int" -apply (simp add: int_def raw_zmult_def) -apply (rule UN_equiv_class_type2 [OF equiv_intrel zmult_congruent2], assumption+) -apply (simp add: Let_def) -done - -lemma zmult_type [iff,TC]: "z $* w : int" -by (simp add: zmult_def raw_zmult_type) - -lemma raw_zmult: - "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) = - intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" -by (simp add: raw_zmult_def - UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2]) - -lemma zmult: - "[| x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat |] - ==> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = - intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}" -by (simp add: zmult_def raw_zmult image_intrel_int) - -lemma raw_zmult_int0: "z : int ==> raw_zmult ($#0,z) = $#0" -by (auto simp add: int_def int_of_def raw_zmult) - -lemma zmult_int0 [simp]: "$#0 $* z = $#0" -by (simp add: zmult_def raw_zmult_int0) - -lemma raw_zmult_int1: "z : int ==> raw_zmult ($#1,z) = z" -by (auto simp add: int_def int_of_def raw_zmult) - -lemma zmult_int1_intify [simp]: "$#1 $* z = intify(z)" -by (simp add: zmult_def raw_zmult_int1) - -lemma zmult_int1: "z : int ==> $#1 $* z = z" -by simp - -lemma raw_zmult_commute: - "[| z: int; w: int |] ==> raw_zmult(z,w) = raw_zmult(w,z)" -by (auto simp add: int_def raw_zmult add_ac mult_ac) - -lemma zmult_commute: "z $* w = w $* z" -by (simp add: zmult_def raw_zmult_commute) - -lemma raw_zmult_zminus: - "[| z: int; w: int |] ==> raw_zmult($- z, w) = $- raw_zmult(z, w)" -by (auto simp add: int_def zminus raw_zmult add_ac) - -lemma zmult_zminus [simp]: "($- z) $* w = $- (z $* w)" -apply (simp add: zmult_def raw_zmult_zminus) -apply (subst zminus_intify [symmetric], rule raw_zmult_zminus, auto) -done - -lemma zmult_zminus_right [simp]: "w $* ($- z) = $- (w $* z)" -by (simp add: zmult_commute [of w]) - -lemma raw_zmult_assoc: - "[| z1: int; z2: int; z3: int |] - ==> raw_zmult (raw_zmult(z1,z2),z3) = raw_zmult(z1,raw_zmult(z2,z3))" -by (auto simp add: int_def raw_zmult add_mult_distrib_left add_ac mult_ac) - -lemma zmult_assoc: "(z1 $* z2) $* z3 = z1 $* (z2 $* z3)" -by (simp add: zmult_def raw_zmult_type raw_zmult_assoc) - -(*For AC rewriting*) -lemma zmult_left_commute: "z1$*(z2$*z3) = z2$*(z1$*z3)" -apply (simp add: zmult_assoc [symmetric]) -apply (simp add: zmult_commute) -done - -(*Integer multiplication is an AC operator*) -lemmas zmult_ac = zmult_assoc zmult_commute zmult_left_commute - -lemma raw_zadd_zmult_distrib: - "[| z1: int; z2: int; w: int |] - ==> raw_zmult(raw_zadd(z1,z2), w) = - raw_zadd (raw_zmult(z1,w), raw_zmult(z2,w))" -by (auto simp add: int_def raw_zadd raw_zmult add_mult_distrib_left add_ac mult_ac) - -lemma zadd_zmult_distrib: "(z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)" -by (simp add: zmult_def zadd_def raw_zadd_type raw_zmult_type - raw_zadd_zmult_distrib) - -lemma zadd_zmult_distrib2: "w $* (z1 $+ z2) = (w $* z1) $+ (w $* z2)" -by (simp add: zmult_commute [of w] zadd_zmult_distrib) - -lemmas int_typechecks = - int_of_type zminus_type zmagnitude_type zadd_type zmult_type - - -(*** Subtraction laws ***) - -lemma zdiff_type [iff,TC]: "z $- w : int" -by (simp add: zdiff_def) - -lemma zminus_zdiff_eq [simp]: "$- (z $- y) = y $- z" -by (simp add: zdiff_def zadd_commute) - -lemma zdiff_zmult_distrib: "(z1 $- z2) $* w = (z1 $* w) $- (z2 $* w)" -apply (simp add: zdiff_def) -apply (subst zadd_zmult_distrib) -apply (simp add: zmult_zminus) -done - -lemma zdiff_zmult_distrib2: "w $* (z1 $- z2) = (w $* z1) $- (w $* z2)" -by (simp add: zmult_commute [of w] zdiff_zmult_distrib) - -lemma zadd_zdiff_eq: "x $+ (y $- z) = (x $+ y) $- z" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zadd_eq: "(x $- y) $+ z = (x $+ z) $- y" -by (simp add: zdiff_def zadd_ac) - - -subsection{*The "Less Than" Relation*} - -(*"Less than" is a linear ordering*) -lemma zless_linear_lemma: - "[| z: int; w: int |] ==> z$<w | z=w | w$<z" -apply (simp add: int_def zless_def znegative_def zdiff_def, auto) -apply (simp add: zadd zminus image_iff Bex_def) -apply (rule_tac i = "xb#+ya" and j = "xc #+ y" in Ord_linear_lt) -apply (force dest!: spec simp add: add_ac)+ -done - -lemma zless_linear: "z$<w | intify(z)=intify(w) | w$<z" -apply (cut_tac z = " intify (z) " and w = " intify (w) " in zless_linear_lemma) -apply auto -done - -lemma zless_not_refl [iff]: "~ (z$<z)" -by (auto simp add: zless_def znegative_def int_of_def zdiff_def) - -lemma neq_iff_zless: "[| x: int; y: int |] ==> (x ~= y) <-> (x $< y | y $< x)" -by (cut_tac z = x and w = y in zless_linear, auto) - -lemma zless_imp_intify_neq: "w $< z ==> intify(w) ~= intify(z)" -apply auto -apply (subgoal_tac "~ (intify (w) $< intify (z))") -apply (erule_tac [2] ssubst) -apply (simp (no_asm_use)) -apply auto -done - -(*This lemma allows direct proofs of other <-properties*) -lemma zless_imp_succ_zadd_lemma: - "[| w $< z; w: int; z: int |] ==> (\<exists>n\<in>nat. z = w $+ $#(succ(n)))" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto dest!: less_imp_succ_add simp add: zadd zminus int_of_def) -apply (rule_tac x = k in bexI) -apply (erule add_left_cancel, auto) -done - -lemma zless_imp_succ_zadd: - "w $< z ==> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))" -apply (subgoal_tac "intify (w) $< intify (z) ") -apply (drule_tac w = "intify (w) " in zless_imp_succ_zadd_lemma) -apply auto -done - -lemma zless_succ_zadd_lemma: - "w : int ==> w $< w $+ $# succ(n)" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto simp add: zadd zminus int_of_def image_iff) -apply (rule_tac x = 0 in exI, auto) -done - -lemma zless_succ_zadd: "w $< w $+ $# succ(n)" -by (cut_tac intify_in_int [THEN zless_succ_zadd_lemma], auto) - -lemma zless_iff_succ_zadd: - "w $< z <-> (\<exists>n\<in>nat. w $+ $#(succ(n)) = intify(z))" -apply (rule iffI) -apply (erule zless_imp_succ_zadd, auto) -apply (rename_tac "n") -apply (cut_tac w = w and n = n in zless_succ_zadd, auto) -done - -lemma zless_int_of [simp]: "[| m\<in>nat; n\<in>nat |] ==> ($#m $< $#n) <-> (m<n)" -apply (simp add: less_iff_succ_add zless_iff_succ_zadd int_of_add [symmetric]) -apply (blast intro: sym) -done - -lemma zless_trans_lemma: - "[| x $< y; y $< z; x: int; y : int; z: int |] ==> x $< z" -apply (simp add: zless_def znegative_def zdiff_def int_def) -apply (auto simp add: zadd zminus image_iff) -apply (rename_tac x1 x2 y1 y2) -apply (rule_tac x = "x1#+x2" in exI) -apply (rule_tac x = "y1#+y2" in exI) -apply (auto simp add: add_lt_mono) -apply (rule sym) -apply (erule add_left_cancel)+ -apply auto -done - -lemma zless_trans: "[| x $< y; y $< z |] ==> x $< z" -apply (subgoal_tac "intify (x) $< intify (z) ") -apply (rule_tac [2] y = "intify (y) " in zless_trans_lemma) -apply auto -done - -lemma zless_not_sym: "z $< w ==> ~ (w $< z)" -by (blast dest: zless_trans) - -(* [| z $< w; ~ P ==> w $< z |] ==> P *) -lemmas zless_asym = zless_not_sym [THEN swap, standard] - -lemma zless_imp_zle: "z $< w ==> z $<= w" -by (simp add: zle_def) - -lemma zle_linear: "z $<= w | w $<= z" -apply (simp add: zle_def) -apply (cut_tac zless_linear, blast) -done - - -subsection{*Less Than or Equals*} - -lemma zle_refl: "z $<= z" -by (simp add: zle_def) - -lemma zle_eq_refl: "x=y ==> x $<= y" -by (simp add: zle_refl) - -lemma zle_anti_sym_intify: "[| x $<= y; y $<= x |] ==> intify(x) = intify(y)" -apply (simp add: zle_def, auto) -apply (blast dest: zless_trans) -done - -lemma zle_anti_sym: "[| x $<= y; y $<= x; x: int; y: int |] ==> x=y" -by (drule zle_anti_sym_intify, auto) - -lemma zle_trans_lemma: - "[| x: int; y: int; z: int; x $<= y; y $<= z |] ==> x $<= z" -apply (simp add: zle_def, auto) -apply (blast intro: zless_trans) -done - -lemma zle_trans: "[| x $<= y; y $<= z |] ==> x $<= z" -apply (subgoal_tac "intify (x) $<= intify (z) ") -apply (rule_tac [2] y = "intify (y) " in zle_trans_lemma) -apply auto -done - -lemma zle_zless_trans: "[| i $<= j; j $< k |] ==> i $< k" -apply (auto simp add: zle_def) -apply (blast intro: zless_trans) -apply (simp add: zless_def zdiff_def zadd_def) -done - -lemma zless_zle_trans: "[| i $< j; j $<= k |] ==> i $< k" -apply (auto simp add: zle_def) -apply (blast intro: zless_trans) -apply (simp add: zless_def zdiff_def zminus_def) -done - -lemma not_zless_iff_zle: "~ (z $< w) <-> (w $<= z)" -apply (cut_tac z = z and w = w in zless_linear) -apply (auto dest: zless_trans simp add: zle_def) -apply (auto dest!: zless_imp_intify_neq) -done - -lemma not_zle_iff_zless: "~ (z $<= w) <-> (w $< z)" -by (simp add: not_zless_iff_zle [THEN iff_sym]) - - -subsection{*More subtraction laws (for @{text zcompare_rls})*} - -lemma zdiff_zdiff_eq: "(x $- y) $- z = x $- (y $+ z)" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zdiff_eq2: "x $- (y $- z) = (x $+ z) $- y" -by (simp add: zdiff_def zadd_ac) - -lemma zdiff_zless_iff: "(x$-y $< z) <-> (x $< z $+ y)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zless_zdiff_iff: "(x $< z$-y) <-> (x $+ y $< z)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zdiff_eq_iff: "[| x: int; z: int |] ==> (x$-y = z) <-> (x = z $+ y)" -by (auto simp add: zdiff_def zadd_assoc) - -lemma eq_zdiff_iff: "[| x: int; z: int |] ==> (x = z$-y) <-> (x $+ y = z)" -by (auto simp add: zdiff_def zadd_assoc) - -lemma zdiff_zle_iff_lemma: - "[| x: int; z: int |] ==> (x$-y $<= z) <-> (x $<= z $+ y)" -by (auto simp add: zle_def zdiff_eq_iff zdiff_zless_iff) - -lemma zdiff_zle_iff: "(x$-y $<= z) <-> (x $<= z $+ y)" -by (cut_tac zdiff_zle_iff_lemma [OF intify_in_int intify_in_int], simp) - -lemma zle_zdiff_iff_lemma: - "[| x: int; z: int |] ==>(x $<= z$-y) <-> (x $+ y $<= z)" -apply (auto simp add: zle_def zdiff_eq_iff zless_zdiff_iff) -apply (auto simp add: zdiff_def zadd_assoc) -done - -lemma zle_zdiff_iff: "(x $<= z$-y) <-> (x $+ y $<= z)" -by (cut_tac zle_zdiff_iff_lemma [ OF intify_in_int intify_in_int], simp) - -text{*This list of rewrites simplifies (in)equalities by bringing subtractions - to the top and then moving negative terms to the other side. - Use with @{text zadd_ac}*} -lemmas zcompare_rls = - zdiff_def [symmetric] - zadd_zdiff_eq zdiff_zadd_eq zdiff_zdiff_eq zdiff_zdiff_eq2 - zdiff_zless_iff zless_zdiff_iff zdiff_zle_iff zle_zdiff_iff - zdiff_eq_iff eq_zdiff_iff - - -subsection{*Monotonicity and Cancellation Results for Instantiation - of the CancelNumerals Simprocs*} - -lemma zadd_left_cancel: - "[| w: int; w': int |] ==> (z $+ w' = z $+ w) <-> (w' = w)" -apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) -apply (simp add: zadd_ac) -done - -lemma zadd_left_cancel_intify [simp]: - "(z $+ w' = z $+ w) <-> intify(w') = intify(w)" -apply (rule iff_trans) -apply (rule_tac [2] zadd_left_cancel, auto) -done - -lemma zadd_right_cancel: - "[| w: int; w': int |] ==> (w' $+ z = w $+ z) <-> (w' = w)" -apply safe -apply (drule_tac t = "%x. x $+ ($-z) " in subst_context) -apply (simp add: zadd_ac) -done - -lemma zadd_right_cancel_intify [simp]: - "(w' $+ z = w $+ z) <-> intify(w') = intify(w)" -apply (rule iff_trans) -apply (rule_tac [2] zadd_right_cancel, auto) -done - -lemma zadd_right_cancel_zless [simp]: "(w' $+ z $< w $+ z) <-> (w' $< w)" -by (simp add: zdiff_zless_iff [THEN iff_sym] zdiff_def zadd_assoc) - -lemma zadd_left_cancel_zless [simp]: "(z $+ w' $< z $+ w) <-> (w' $< w)" -by (simp add: zadd_commute [of z] zadd_right_cancel_zless) - -lemma zadd_right_cancel_zle [simp]: "(w' $+ z $<= w $+ z) <-> w' $<= w" -by (simp add: zle_def) - -lemma zadd_left_cancel_zle [simp]: "(z $+ w' $<= z $+ w) <-> w' $<= w" -by (simp add: zadd_commute [of z] zadd_right_cancel_zle) - - -(*"v $<= w ==> v$+z $<= w$+z"*) -lemmas zadd_zless_mono1 = zadd_right_cancel_zless [THEN iffD2, standard] - -(*"v $<= w ==> z$+v $<= z$+w"*) -lemmas zadd_zless_mono2 = zadd_left_cancel_zless [THEN iffD2, standard] - -(*"v $<= w ==> v$+z $<= w$+z"*) -lemmas zadd_zle_mono1 = zadd_right_cancel_zle [THEN iffD2, standard] - -(*"v $<= w ==> z$+v $<= z$+w"*) -lemmas zadd_zle_mono2 = zadd_left_cancel_zle [THEN iffD2, standard] - -lemma zadd_zle_mono: "[| w' $<= w; z' $<= z |] ==> w' $+ z' $<= w $+ z" -by (erule zadd_zle_mono1 [THEN zle_trans], simp) - -lemma zadd_zless_mono: "[| w' $< w; z' $<= z |] ==> w' $+ z' $< w $+ z" -by (erule zadd_zless_mono1 [THEN zless_zle_trans], simp) - - -subsection{*Comparison laws*} - -lemma zminus_zless_zminus [simp]: "($- x $< $- y) <-> (y $< x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zminus_zle_zminus [simp]: "($- x $<= $- y) <-> (y $<= x)" -by (simp add: not_zless_iff_zle [THEN iff_sym]) - -subsubsection{*More inequality lemmas*} - -lemma equation_zminus: "[| x: int; y: int |] ==> (x = $- y) <-> (y = $- x)" -by auto - -lemma zminus_equation: "[| x: int; y: int |] ==> ($- x = y) <-> ($- y = x)" -by auto - -lemma equation_zminus_intify: "(intify(x) = $- y) <-> (intify(y) = $- x)" -apply (cut_tac x = "intify (x) " and y = "intify (y) " in equation_zminus) -apply auto -done - -lemma zminus_equation_intify: "($- x = intify(y)) <-> ($- y = intify(x))" -apply (cut_tac x = "intify (x) " and y = "intify (y) " in zminus_equation) -apply auto -done - - -subsubsection{*The next several equations are permutative: watch out!*} - -lemma zless_zminus: "(x $< $- y) <-> (y $< $- x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zminus_zless: "($- x $< y) <-> ($- y $< x)" -by (simp add: zless_def zdiff_def zadd_ac) - -lemma zle_zminus: "(x $<= $- y) <-> (y $<= $- x)" -by (simp add: not_zless_iff_zle [THEN iff_sym] zminus_zless) - -lemma zminus_zle: "($- x $<= y) <-> ($- y $<= x)" -by (simp add: not_zless_iff_zle [THEN iff_sym] zless_zminus) - -ML -{* -val zdiff_def = thm "zdiff_def"; -val int_of_type = thm "int_of_type"; -val int_of_eq = thm "int_of_eq"; -val int_of_inject = thm "int_of_inject"; -val intify_in_int = thm "intify_in_int"; -val intify_ident = thm "intify_ident"; -val intify_idem = thm "intify_idem"; -val int_of_natify = thm "int_of_natify"; -val zminus_intify = thm "zminus_intify"; -val zadd_intify1 = thm "zadd_intify1"; -val zadd_intify2 = thm "zadd_intify2"; -val zdiff_intify1 = thm "zdiff_intify1"; -val zdiff_intify2 = thm "zdiff_intify2"; -val zmult_intify1 = thm "zmult_intify1"; -val zmult_intify2 = thm "zmult_intify2"; -val zless_intify1 = thm "zless_intify1"; -val zless_intify2 = thm "zless_intify2"; -val zle_intify1 = thm "zle_intify1"; -val zle_intify2 = thm "zle_intify2"; -val zminus_congruent = thm "zminus_congruent"; -val zminus_type = thm "zminus_type"; -val zminus_inject_intify = thm "zminus_inject_intify"; -val zminus_inject = thm "zminus_inject"; -val zminus = thm "zminus"; -val zminus_zminus_intify = thm "zminus_zminus_intify"; -val zminus_int0 = thm "zminus_int0"; -val zminus_zminus = thm "zminus_zminus"; -val not_znegative_int_of = thm "not_znegative_int_of"; -val znegative_zminus_int_of = thm "znegative_zminus_int_of"; -val not_znegative_imp_zero = thm "not_znegative_imp_zero"; -val nat_of_intify = thm "nat_of_intify"; -val nat_of_int_of = thm "nat_of_int_of"; -val nat_of_type = thm "nat_of_type"; -val zmagnitude_int_of = thm "zmagnitude_int_of"; -val natify_int_of_eq = thm "natify_int_of_eq"; -val zmagnitude_zminus_int_of = thm "zmagnitude_zminus_int_of"; -val zmagnitude_type = thm "zmagnitude_type"; -val not_zneg_int_of = thm "not_zneg_int_of"; -val not_zneg_mag = thm "not_zneg_mag"; -val zneg_int_of = thm "zneg_int_of"; -val zneg_mag = thm "zneg_mag"; -val int_cases = thm "int_cases"; -val not_zneg_nat_of_intify = thm "not_zneg_nat_of_intify"; -val not_zneg_nat_of = thm "not_zneg_nat_of"; -val zneg_nat_of = thm "zneg_nat_of"; -val zadd_congruent2 = thm "zadd_congruent2"; -val zadd_type = thm "zadd_type"; -val zadd = thm "zadd"; -val zadd_int0_intify = thm "zadd_int0_intify"; -val zadd_int0 = thm "zadd_int0"; -val zminus_zadd_distrib = thm "zminus_zadd_distrib"; -val zadd_commute = thm "zadd_commute"; -val zadd_assoc = thm "zadd_assoc"; -val zadd_left_commute = thm "zadd_left_commute"; -val zadd_ac = thms "zadd_ac"; -val int_of_add = thm "int_of_add"; -val int_succ_int_1 = thm "int_succ_int_1"; -val int_of_diff = thm "int_of_diff"; -val zadd_zminus_inverse = thm "zadd_zminus_inverse"; -val zadd_zminus_inverse2 = thm "zadd_zminus_inverse2"; -val zadd_int0_right_intify = thm "zadd_int0_right_intify"; -val zadd_int0_right = thm "zadd_int0_right"; -val zmult_congruent2 = thm "zmult_congruent2"; -val zmult_type = thm "zmult_type"; -val zmult = thm "zmult"; -val zmult_int0 = thm "zmult_int0"; -val zmult_int1_intify = thm "zmult_int1_intify"; -val zmult_int1 = thm "zmult_int1"; -val zmult_commute = thm "zmult_commute"; -val zmult_zminus = thm "zmult_zminus"; -val zmult_zminus_right = thm "zmult_zminus_right"; -val zmult_assoc = thm "zmult_assoc"; -val zmult_left_commute = thm "zmult_left_commute"; -val zmult_ac = thms "zmult_ac"; -val zadd_zmult_distrib = thm "zadd_zmult_distrib"; -val zadd_zmult_distrib2 = thm "zadd_zmult_distrib2"; -val int_typechecks = thms "int_typechecks"; -val zdiff_type = thm "zdiff_type"; -val zminus_zdiff_eq = thm "zminus_zdiff_eq"; -val zdiff_zmult_distrib = thm "zdiff_zmult_distrib"; -val zdiff_zmult_distrib2 = thm "zdiff_zmult_distrib2"; -val zadd_zdiff_eq = thm "zadd_zdiff_eq"; -val zdiff_zadd_eq = thm "zdiff_zadd_eq"; -val zless_linear = thm "zless_linear"; -val zless_not_refl = thm "zless_not_refl"; -val neq_iff_zless = thm "neq_iff_zless"; -val zless_imp_intify_neq = thm "zless_imp_intify_neq"; -val zless_imp_succ_zadd = thm "zless_imp_succ_zadd"; -val zless_succ_zadd = thm "zless_succ_zadd"; -val zless_iff_succ_zadd = thm "zless_iff_succ_zadd"; -val zless_int_of = thm "zless_int_of"; -val zless_trans = thm "zless_trans"; -val zless_not_sym = thm "zless_not_sym"; -val zless_asym = thm "zless_asym"; -val zless_imp_zle = thm "zless_imp_zle"; -val zle_linear = thm "zle_linear"; -val zle_refl = thm "zle_refl"; -val zle_eq_refl = thm "zle_eq_refl"; -val zle_anti_sym_intify = thm "zle_anti_sym_intify"; -val zle_anti_sym = thm "zle_anti_sym"; -val zle_trans = thm "zle_trans"; -val zle_zless_trans = thm "zle_zless_trans"; -val zless_zle_trans = thm "zless_zle_trans"; -val not_zless_iff_zle = thm "not_zless_iff_zle"; -val not_zle_iff_zless = thm "not_zle_iff_zless"; -val zdiff_zdiff_eq = thm "zdiff_zdiff_eq"; -val zdiff_zdiff_eq2 = thm "zdiff_zdiff_eq2"; -val zdiff_zless_iff = thm "zdiff_zless_iff"; -val zless_zdiff_iff = thm "zless_zdiff_iff"; -val zdiff_eq_iff = thm "zdiff_eq_iff"; -val eq_zdiff_iff = thm "eq_zdiff_iff"; -val zdiff_zle_iff = thm "zdiff_zle_iff"; -val zle_zdiff_iff = thm "zle_zdiff_iff"; -val zcompare_rls = thms "zcompare_rls"; -val zadd_left_cancel = thm "zadd_left_cancel"; -val zadd_left_cancel_intify = thm "zadd_left_cancel_intify"; -val zadd_right_cancel = thm "zadd_right_cancel"; -val zadd_right_cancel_intify = thm "zadd_right_cancel_intify"; -val zadd_right_cancel_zless = thm "zadd_right_cancel_zless"; -val zadd_left_cancel_zless = thm "zadd_left_cancel_zless"; -val zadd_right_cancel_zle = thm "zadd_right_cancel_zle"; -val zadd_left_cancel_zle = thm "zadd_left_cancel_zle"; -val zadd_zless_mono1 = thm "zadd_zless_mono1"; -val zadd_zless_mono2 = thm "zadd_zless_mono2"; -val zadd_zle_mono1 = thm "zadd_zle_mono1"; -val zadd_zle_mono2 = thm "zadd_zle_mono2"; -val zadd_zle_mono = thm "zadd_zle_mono"; -val zadd_zless_mono = thm "zadd_zless_mono"; -val zminus_zless_zminus = thm "zminus_zless_zminus"; -val zminus_zle_zminus = thm "zminus_zle_zminus"; -val equation_zminus = thm "equation_zminus"; -val zminus_equation = thm "zminus_equation"; -val equation_zminus_intify = thm "equation_zminus_intify"; -val zminus_equation_intify = thm "zminus_equation_intify"; -val zless_zminus = thm "zless_zminus"; -val zminus_zless = thm "zminus_zless"; -val zle_zminus = thm "zle_zminus"; -val zminus_zle = thm "zminus_zle"; -*} - - -end
--- a/src/ZF/Integ/IntArith.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,5 +0,0 @@ - -theory IntArith imports Bin -uses "int_arith.ML" begin - -end
--- a/src/ZF/Integ/IntDiv.thy Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1925 +0,0 @@ -(* Title: ZF/IntDiv.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge - -Here is the division algorithm in ML: - - fun posDivAlg (a,b) = - if a<b then (0,a) - else let val (q,r) = posDivAlg(a, 2*b) - in if 0<=r-b then (2*q+1, r-b) else (2*q, r) - end - - fun negDivAlg (a,b) = - if 0<=a+b then (~1,a+b) - else let val (q,r) = negDivAlg(a, 2*b) - in if 0<=r-b then (2*q+1, r-b) else (2*q, r) - end; - - fun negateSnd (q,r:int) = (q,~r); - - fun divAlg (a,b) = if 0<=a then - if b>0 then posDivAlg (a,b) - else if a=0 then (0,0) - else negateSnd (negDivAlg (~a,~b)) - else - if 0<b then negDivAlg (a,b) - else negateSnd (posDivAlg (~a,~b)); - -*) - -header{*The Division Operators Div and Mod*} - -theory IntDiv imports IntArith OrderArith begin - -constdefs - quorem :: "[i,i] => o" - "quorem == %<a,b> <q,r>. - a = b$*q $+ r & - (#0$<b & #0$<=r & r$<b | ~(#0$<b) & b$<r & r $<= #0)" - - adjust :: "[i,i] => i" - "adjust(b) == %<q,r>. if #0 $<= r$-b then <#2$*q $+ #1,r$-b> - else <#2$*q,r>" - - -(** the division algorithm **) - -constdefs posDivAlg :: "i => i" -(*for the case a>=0, b>0*) -(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*) - "posDivAlg(ab) == - wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)), - ab, - %<a,b> f. if (a$<b | b$<=#0) then <#0,a> - else adjust(b, f ` <a,#2$*b>))" - - -(*for the case a<0, b>0*) -constdefs negDivAlg :: "i => i" -(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*) - "negDivAlg(ab) == - wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)), - ab, - %<a,b> f. if (#0 $<= a$+b | b$<=#0) then <#-1,a$+b> - else adjust(b, f ` <a,#2$*b>))" - -(*for the general case b\<noteq>0*) - -constdefs - negateSnd :: "i => i" - "negateSnd == %<q,r>. <q, $-r>" - - (*The full division algorithm considers all possible signs for a, b - including the special case a=0, b<0, because negDivAlg requires a<0*) - divAlg :: "i => i" - "divAlg == - %<a,b>. if #0 $<= a then - if #0 $<= b then posDivAlg (<a,b>) - else if a=#0 then <#0,#0> - else negateSnd (negDivAlg (<$-a,$-b>)) - else - if #0$<b then negDivAlg (<a,b>) - else negateSnd (posDivAlg (<$-a,$-b>))" - - zdiv :: "[i,i]=>i" (infixl "zdiv" 70) - "a zdiv b == fst (divAlg (<intify(a), intify(b)>))" - - zmod :: "[i,i]=>i" (infixl "zmod" 70) - "a zmod b == snd (divAlg (<intify(a), intify(b)>))" - - -(** Some basic laws by Sidi Ehmety (need linear arithmetic!) **) - -lemma zspos_add_zspos_imp_zspos: "[| #0 $< x; #0 $< y |] ==> #0 $< x $+ y" -apply (rule_tac y = "y" in zless_trans) -apply (rule_tac [2] zdiff_zless_iff [THEN iffD1]) -apply auto -done - -lemma zpos_add_zpos_imp_zpos: "[| #0 $<= x; #0 $<= y |] ==> #0 $<= x $+ y" -apply (rule_tac y = "y" in zle_trans) -apply (rule_tac [2] zdiff_zle_iff [THEN iffD1]) -apply auto -done - -lemma zneg_add_zneg_imp_zneg: "[| x $< #0; y $< #0 |] ==> x $+ y $< #0" -apply (rule_tac y = "y" in zless_trans) -apply (rule zless_zdiff_iff [THEN iffD1]) -apply auto -done - -(* this theorem is used below *) -lemma zneg_or_0_add_zneg_or_0_imp_zneg_or_0: - "[| x $<= #0; y $<= #0 |] ==> x $+ y $<= #0" -apply (rule_tac y = "y" in zle_trans) -apply (rule zle_zdiff_iff [THEN iffD1]) -apply auto -done - -lemma zero_lt_zmagnitude: "[| #0 $< k; k \<in> int |] ==> 0 < zmagnitude(k)" -apply (drule zero_zless_imp_znegative_zminus) -apply (drule_tac [2] zneg_int_of) -apply (auto simp add: zminus_equation [of k]) -apply (subgoal_tac "0 < zmagnitude ($# succ (n))") - apply simp -apply (simp only: zmagnitude_int_of) -apply simp -done - - -(*** Inequality lemmas involving $#succ(m) ***) - -lemma zless_add_succ_iff: - "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)" -apply (auto simp add: zless_iff_succ_zadd zadd_assoc int_of_add [symmetric]) -apply (rule_tac [3] x = "0" in bexI) -apply (cut_tac m = "m" in int_succ_int_1) -apply (cut_tac m = "n" in int_succ_int_1) -apply simp -apply (erule natE) -apply auto -apply (rule_tac x = "succ (n) " in bexI) -apply auto -done - -lemma zadd_succ_lemma: - "z \<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" -apply (simp only: not_zless_iff_zle [THEN iff_sym] zless_add_succ_iff) -apply (auto intro: zle_anti_sym elim: zless_asym - simp add: zless_imp_zle not_zless_iff_zle) -done - -lemma zadd_succ_zle_iff: "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)" -apply (cut_tac z = "intify (z)" in zadd_succ_lemma) -apply auto -done - -(** Inequality reasoning **) - -lemma zless_add1_iff_zle: "(w $< z $+ #1) <-> (w$<=z)" -apply (subgoal_tac "#1 = $# 1") -apply (simp only: zless_add_succ_iff zle_def) -apply auto -done - -lemma add1_zle_iff: "(w $+ #1 $<= z) <-> (w $< z)" -apply (subgoal_tac "#1 = $# 1") -apply (simp only: zadd_succ_zle_iff) -apply auto -done - -lemma add1_left_zle_iff: "(#1 $+ w $<= z) <-> (w $< z)" -apply (subst zadd_commute) -apply (rule add1_zle_iff) -done - - -(*** Monotonicity of Multiplication ***) - -lemma zmult_mono_lemma: "k \<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k" -apply (induct_tac "k") - prefer 2 apply (subst int_succ_int_1) -apply (simp_all (no_asm_simp) add: zadd_zmult_distrib2 zadd_zle_mono) -done - -lemma zmult_zle_mono1: "[| i $<= j; #0 $<= k |] ==> i$*k $<= j$*k" -apply (subgoal_tac "i $* intify (k) $<= j $* intify (k) ") -apply (simp (no_asm_use)) -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) -apply (rule_tac [3] zmult_mono_lemma) -apply auto -apply (simp add: znegative_iff_zless_0 not_zless_iff_zle [THEN iff_sym]) -done - -lemma zmult_zle_mono1_neg: "[| i $<= j; k $<= #0 |] ==> j$*k $<= i$*k" -apply (rule zminus_zle_zminus [THEN iffD1]) -apply (simp del: zmult_zminus_right - add: zmult_zminus_right [symmetric] zmult_zle_mono1 zle_zminus) -done - -lemma zmult_zle_mono2: "[| i $<= j; #0 $<= k |] ==> k$*i $<= k$*j" -apply (drule zmult_zle_mono1) -apply (simp_all add: zmult_commute) -done - -lemma zmult_zle_mono2_neg: "[| i $<= j; k $<= #0 |] ==> k$*j $<= k$*i" -apply (drule zmult_zle_mono1_neg) -apply (simp_all add: zmult_commute) -done - -(* $<= monotonicity, BOTH arguments*) -lemma zmult_zle_mono: - "[| i $<= j; k $<= l; #0 $<= j; #0 $<= k |] ==> i$*k $<= j$*l" -apply (erule zmult_zle_mono1 [THEN zle_trans]) -apply assumption -apply (erule zmult_zle_mono2) -apply assumption -done - - -(** strict, in 1st argument; proof is by induction on k>0 **) - -lemma zmult_zless_mono2_lemma [rule_format]: - "[| i$<j; k \<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j" -apply (induct_tac "k") - prefer 2 - apply (subst int_succ_int_1) - apply (erule natE) -apply (simp_all add: zadd_zmult_distrib zadd_zless_mono zle_def) -apply (frule nat_0_le) -apply (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j) ") -apply (simp (no_asm_use)) -apply (rule zadd_zless_mono) -apply (simp_all (no_asm_simp) add: zle_def) -done - -lemma zmult_zless_mono2: "[| i$<j; #0 $< k |] ==> k$*i $< k$*j" -apply (subgoal_tac "intify (k) $* i $< intify (k) $* j") -apply (simp (no_asm_use)) -apply (rule_tac b = "intify (k)" in not_zneg_mag [THEN subst]) -apply (rule_tac [3] zmult_zless_mono2_lemma) -apply auto -apply (simp add: znegative_iff_zless_0) -apply (drule zless_trans, assumption) -apply (auto simp add: zero_lt_zmagnitude) -done - -lemma zmult_zless_mono1: "[| i$<j; #0 $< k |] ==> i$*k $< j$*k" -apply (drule zmult_zless_mono2) -apply (simp_all add: zmult_commute) -done - -(* < monotonicity, BOTH arguments*) -lemma zmult_zless_mono: - "[| i $< j; k $< l; #0 $< j; #0 $< k |] ==> i$*k $< j$*l" -apply (erule zmult_zless_mono1 [THEN zless_trans]) -apply assumption -apply (erule zmult_zless_mono2) -apply assumption -done - -lemma zmult_zless_mono1_neg: "[| i $< j; k $< #0 |] ==> j$*k $< i$*k" -apply (rule zminus_zless_zminus [THEN iffD1]) -apply (simp del: zmult_zminus_right - add: zmult_zminus_right [symmetric] zmult_zless_mono1 zless_zminus) -done - -lemma zmult_zless_mono2_neg: "[| i $< j; k $< #0 |] ==> k$*j $< k$*i" -apply (rule zminus_zless_zminus [THEN iffD1]) -apply (simp del: zmult_zminus - add: zmult_zminus [symmetric] zmult_zless_mono2 zless_zminus) -done - - -(** Products of zeroes **) - -lemma zmult_eq_lemma: - "[| m \<in> int; n \<in> int |] ==> (m = #0 | n = #0) <-> (m$*n = #0)" -apply (case_tac "m $< #0") -apply (auto simp add: not_zless_iff_zle zle_def neq_iff_zless) -apply (force dest: zmult_zless_mono1_neg zmult_zless_mono1)+ -done - -lemma zmult_eq_0_iff [iff]: "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)" -apply (simp add: zmult_eq_lemma) -done - - -(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, - but not (yet?) for k*m < n*k. **) - -lemma zmult_zless_lemma: - "[| k \<in> int; m \<in> int; n \<in> int |] - ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -apply (case_tac "k = #0") -apply (auto simp add: neq_iff_zless zmult_zless_mono1 zmult_zless_mono1_neg) -apply (auto simp add: not_zless_iff_zle - not_zle_iff_zless [THEN iff_sym, of "m$*k"] - not_zle_iff_zless [THEN iff_sym, of m]) -apply (auto elim: notE - simp add: zless_imp_zle zmult_zle_mono1 zmult_zle_mono1_neg) -done - -lemma zmult_zless_cancel2: - "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -apply (cut_tac k = "intify (k)" and m = "intify (m)" and n = "intify (n)" - in zmult_zless_lemma) -apply auto -done - -lemma zmult_zless_cancel1: - "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))" -by (simp add: zmult_commute [of k] zmult_zless_cancel2) - -lemma zmult_zle_cancel2: - "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel2) - -lemma zmult_zle_cancel1: - "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))" -by (auto simp add: not_zless_iff_zle [THEN iff_sym] zmult_zless_cancel1) - -lemma int_eq_iff_zle: "[| m \<in> int; n \<in> int |] ==> m=n <-> (m $<= n & n $<= m)" -apply (blast intro: zle_refl zle_anti_sym) -done - -lemma zmult_cancel2_lemma: - "[| k \<in> int; m \<in> int; n \<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)" -apply (simp add: int_eq_iff_zle [of "m$*k"] int_eq_iff_zle [of m]) -apply (auto simp add: zmult_zle_cancel2 neq_iff_zless) -done - -lemma zmult_cancel2 [simp]: - "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))" -apply (rule iff_trans) -apply (rule_tac [2] zmult_cancel2_lemma) -apply auto -done - -lemma zmult_cancel1 [simp]: - "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))" -by (simp add: zmult_commute [of k] zmult_cancel2) - - -subsection{* Uniqueness and monotonicity of quotients and remainders *} - -lemma unique_quotient_lemma: - "[| b$*q' $+ r' $<= b$*q $+ r; #0 $<= r'; #0 $< b; r $< b |] - ==> q' $<= q" -apply (subgoal_tac "r' $+ b $* (q'$-q) $<= r") - prefer 2 apply (simp add: zdiff_zmult_distrib2 zadd_ac zcompare_rls) -apply (subgoal_tac "#0 $< b $* (#1 $+ q $- q') ") - prefer 2 - apply (erule zle_zless_trans) - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) - apply (erule zle_zless_trans) - apply (simp add: ); -apply (subgoal_tac "b $* q' $< b $* (#1 $+ q)") - prefer 2 - apply (simp add: zdiff_zmult_distrib2 zadd_zmult_distrib2 zadd_ac zcompare_rls) -apply (auto elim: zless_asym - simp add: zmult_zless_cancel1 zless_add1_iff_zle zadd_ac zcompare_rls) -done - -lemma unique_quotient_lemma_neg: - "[| b$*q' $+ r' $<= b$*q $+ r; r $<= #0; b $< #0; b $< r' |] - ==> q $<= q'" -apply (rule_tac b = "$-b" and r = "$-r'" and r' = "$-r" - in unique_quotient_lemma) -apply (auto simp del: zminus_zadd_distrib - simp add: zminus_zadd_distrib [symmetric] zle_zminus zless_zminus) -done - - -lemma unique_quotient: - "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; - q \<in> int; q' \<in> int |] ==> q = q'" -apply (simp add: split_ifs quorem_def neq_iff_zless) -apply safe -apply simp_all -apply (blast intro: zle_anti_sym - dest: zle_eq_refl [THEN unique_quotient_lemma] - zle_eq_refl [THEN unique_quotient_lemma_neg] sym)+ -done - -lemma unique_remainder: - "[| quorem (<a,b>, <q,r>); quorem (<a,b>, <q',r'>); b \<in> int; b ~= #0; - q \<in> int; q' \<in> int; - r \<in> int; r' \<in> int |] ==> r = r'" -apply (subgoal_tac "q = q'") - prefer 2 apply (blast intro: unique_quotient) -apply (simp add: quorem_def) -done - - -subsection{*Correctness of posDivAlg, - the Division Algorithm for @{text "a\<ge>0"} and @{text "b>0"} *} - -lemma adjust_eq [simp]: - "adjust(b, <q,r>) = (let diff = r$-b in - if #0 $<= diff then <#2$*q $+ #1,diff> - else <#2$*q,r>)" -by (simp add: Let_def adjust_def) - - -lemma posDivAlg_termination: - "[| #0 $< b; ~ a $< b |] - ==> nat_of(a $- #2 $\<times> b $+ #1) < nat_of(a $- b $+ #1)" -apply (simp (no_asm) add: zless_nat_conj) -apply (simp add: not_zless_iff_zle zless_add1_iff_zle zcompare_rls) -done - -lemmas posDivAlg_unfold = def_wfrec [OF posDivAlg_def wf_measure] - -lemma posDivAlg_eqn: - "[| #0 $< b; a \<in> int; b \<in> int |] ==> - posDivAlg(<a,b>) = - (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))" -apply (rule posDivAlg_unfold [THEN trans]) -apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: posDivAlg_termination) -done - -lemma posDivAlg_induct_lemma [rule_format]: - assumes prem: - "!!a b. [| a \<in> int; b \<in> int; - ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] ==> P(<a,b>)" - shows "<u,v> \<in> int*int --> P(<u,v>)" -apply (rule_tac a = "<u,v>" in wf_induct) -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)" - in wf_measure) -apply clarify -apply (rule prem) -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) -apply auto -apply (simp add: not_zle_iff_zless posDivAlg_termination) -done - - -lemma posDivAlg_induct [consumes 2]: - assumes u_int: "u \<in> int" - and v_int: "v \<in> int" - and ih: "!!a b. [| a \<in> int; b \<in> int; - ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)") -apply simp -apply (rule posDivAlg_induct_lemma) -apply (simp (no_asm_use)) -apply (rule ih) -apply (auto simp add: u_int v_int) -done - -(*FIXME: use intify in integ_of so that we always have integ_of w \<in> int. - then this rewrite can work for ALL constants!!*) -lemma intify_eq_0_iff_zle: "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)" -apply (simp (no_asm) add: int_eq_iff_zle) -done - - -subsection{* Some convenient biconditionals for products of signs *} - -lemma zmult_pos: "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j" -apply (drule zmult_zless_mono1) -apply auto -done - -lemma zmult_neg: "[| i $< #0; j $< #0 |] ==> #0 $< i $* j" -apply (drule zmult_zless_mono1_neg) -apply auto -done - -lemma zmult_pos_neg: "[| #0 $< i; j $< #0 |] ==> i $* j $< #0" -apply (drule zmult_zless_mono1_neg) -apply auto -done - -(** Inequality reasoning **) - -lemma int_0_less_lemma: - "[| x \<in> int; y \<in> int |] - ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" -apply (auto simp add: zle_def not_zless_iff_zle zmult_pos zmult_neg) -apply (rule ccontr) -apply (rule_tac [2] ccontr) -apply (auto simp add: zle_def not_zless_iff_zle) -apply (erule_tac P = "#0$< x$* y" in rev_mp) -apply (erule_tac [2] P = "#0$< x$* y" in rev_mp) -apply (drule zmult_pos_neg, assumption) - prefer 2 - apply (drule zmult_pos_neg, assumption) -apply (auto dest: zless_not_sym simp add: zmult_commute) -done - -lemma int_0_less_mult_iff: - "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)" -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_less_lemma) -apply auto -done - -lemma int_0_le_lemma: - "[| x \<in> int; y \<in> int |] - ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)" -by (auto simp add: zle_def not_zless_iff_zle int_0_less_mult_iff) - -lemma int_0_le_mult_iff: - "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))" -apply (cut_tac x = "intify (x)" and y = "intify (y)" in int_0_le_lemma) -apply auto -done - -lemma zmult_less_0_iff: - "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)" -apply (auto simp add: int_0_le_mult_iff not_zle_iff_zless [THEN iff_sym]) -apply (auto dest: zless_not_sym simp add: not_zle_iff_zless) -done - -lemma zmult_le_0_iff: - "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)" -by (auto dest: zless_not_sym - simp add: int_0_less_mult_iff not_zless_iff_zle [THEN iff_sym]) - - -(*Typechecking for posDivAlg*) -lemma posDivAlg_type [rule_format]: - "[| a \<in> int; b \<in> int |] ==> posDivAlg(<a,b>) \<in> int * int" -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) -apply assumption+ -apply (case_tac "#0 $< ba") - apply (simp add: posDivAlg_eqn adjust_def integ_of_type - split add: split_if_asm) - apply clarify - apply (simp add: int_0_less_mult_iff not_zle_iff_zless) -apply (simp add: not_zless_iff_zle) -apply (subst posDivAlg_unfold) -apply simp -done - -(*Correctness of posDivAlg: it computes quotients correctly*) -lemma posDivAlg_correct [rule_format]: - "[| a \<in> int; b \<in> int |] - ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))" -apply (rule_tac u = "a" and v = "b" in posDivAlg_induct) -apply auto - apply (simp_all add: quorem_def) - txt{*base case: a<b*} - apply (simp add: posDivAlg_eqn) - apply (simp add: not_zless_iff_zle [THEN iff_sym]) - apply (simp add: int_0_less_mult_iff) -txt{*main argument*} -apply (subst posDivAlg_eqn) -apply (simp_all (no_asm_simp)) -apply (erule splitE) -apply (rule posDivAlg_type) -apply (simp_all add: int_0_less_mult_iff) -apply (auto simp add: zadd_zmult_distrib2 Let_def) -txt{*now just linear arithmetic*} -apply (simp add: not_zle_iff_zless zdiff_zless_iff) -done - - -subsection{*Correctness of negDivAlg, the division algorithm for a<0 and b>0*} - -lemma negDivAlg_termination: - "[| #0 $< b; a $+ b $< #0 |] - ==> nat_of($- a $- #2 $* b) < nat_of($- a $- b)" -apply (simp (no_asm) add: zless_nat_conj) -apply (simp add: zcompare_rls not_zle_iff_zless zless_zdiff_iff [THEN iff_sym] - zless_zminus) -done - -lemmas negDivAlg_unfold = def_wfrec [OF negDivAlg_def wf_measure] - -lemma negDivAlg_eqn: - "[| #0 $< b; a : int; b : int |] ==> - negDivAlg(<a,b>) = - (if #0 $<= a$+b then <#-1,a$+b> - else adjust(b, negDivAlg (<a, #2$*b>)))" -apply (rule negDivAlg_unfold [THEN trans]) -apply (simp (no_asm_simp) add: vimage_iff not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: negDivAlg_termination) -done - -lemma negDivAlg_induct_lemma [rule_format]: - assumes prem: - "!!a b. [| a \<in> int; b \<in> int; - ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] - ==> P(<a,b>)" - shows "<u,v> \<in> int*int --> P(<u,v>)" -apply (rule_tac a = "<u,v>" in wf_induct) -apply (rule_tac A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)" - in wf_measure) -apply clarify -apply (rule prem) -apply (drule_tac [3] x = "<xa, #2 $\<times> y>" in spec) -apply auto -apply (simp add: not_zle_iff_zless negDivAlg_termination) -done - -lemma negDivAlg_induct [consumes 2]: - assumes u_int: "u \<in> int" - and v_int: "v \<in> int" - and ih: "!!a b. [| a \<in> int; b \<in> int; - ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] - ==> P(a,b)" - shows "P(u,v)" -apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)") -apply simp -apply (rule negDivAlg_induct_lemma) -apply (simp (no_asm_use)) -apply (rule ih) -apply (auto simp add: u_int v_int) -done - - -(*Typechecking for negDivAlg*) -lemma negDivAlg_type: - "[| a \<in> int; b \<in> int |] ==> negDivAlg(<a,b>) \<in> int * int" -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) -apply assumption+ -apply (case_tac "#0 $< ba") - apply (simp add: negDivAlg_eqn adjust_def integ_of_type - split add: split_if_asm) - apply clarify - apply (simp add: int_0_less_mult_iff not_zle_iff_zless) -apply (simp add: not_zless_iff_zle) -apply (subst negDivAlg_unfold) -apply simp -done - - -(*Correctness of negDivAlg: it computes quotients correctly - It doesn't work if a=0 because the 0/b=0 rather than -1*) -lemma negDivAlg_correct [rule_format]: - "[| a \<in> int; b \<in> int |] - ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))" -apply (rule_tac u = "a" and v = "b" in negDivAlg_induct) - apply auto - apply (simp_all add: quorem_def) - txt{*base case: @{term "0$<=a$+b"}*} - apply (simp add: negDivAlg_eqn) - apply (simp add: not_zless_iff_zle [THEN iff_sym]) - apply (simp add: int_0_less_mult_iff) -txt{*main argument*} -apply (subst negDivAlg_eqn) -apply (simp_all (no_asm_simp)) -apply (erule splitE) -apply (rule negDivAlg_type) -apply (simp_all add: int_0_less_mult_iff) -apply (auto simp add: zadd_zmult_distrib2 Let_def) -txt{*now just linear arithmetic*} -apply (simp add: not_zle_iff_zless zdiff_zless_iff) -done - - -subsection{* Existence shown by proving the division algorithm to be correct *} - -(*the case a=0*) -lemma quorem_0: "[|b \<noteq> #0; b \<in> int|] ==> quorem (<#0,b>, <#0,#0>)" -by (force simp add: quorem_def neq_iff_zless) - -lemma posDivAlg_zero_divisor: "posDivAlg(<a,#0>) = <#0,a>" -apply (subst posDivAlg_unfold) -apply simp -done - -lemma posDivAlg_0 [simp]: "posDivAlg (<#0,b>) = <#0,#0>" -apply (subst posDivAlg_unfold) -apply (simp add: not_zle_iff_zless) -done - - -(*Needed below. Actually it's an equivalence.*) -lemma linear_arith_lemma: "~ (#0 $<= #-1 $+ b) ==> (b $<= #0)" -apply (simp add: not_zle_iff_zless) -apply (drule zminus_zless_zminus [THEN iffD2]) -apply (simp add: zadd_commute zless_add1_iff_zle zle_zminus) -done - -lemma negDivAlg_minus1 [simp]: "negDivAlg (<#-1,b>) = <#-1, b$-#1>" -apply (subst negDivAlg_unfold) -apply (simp add: linear_arith_lemma integ_of_type vimage_iff) -done - -lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>" -apply (unfold negateSnd_def) -apply auto -done - -lemma negateSnd_type: "qr \<in> int * int ==> negateSnd (qr) \<in> int * int" -apply (unfold negateSnd_def) -apply auto -done - -lemma quorem_neg: - "[|quorem (<$-a,$-b>, qr); a \<in> int; b \<in> int; qr \<in> int * int|] - ==> quorem (<a,b>, negateSnd(qr))" -apply clarify -apply (auto elim: zless_asym simp add: quorem_def zless_zminus) -txt{*linear arithmetic from here on*} -apply (simp_all add: zminus_equation [of a] zminus_zless) -apply (cut_tac [2] z = "b" and w = "#0" in zless_linear) -apply (cut_tac [1] z = "b" and w = "#0" in zless_linear) -apply auto -apply (blast dest: zle_zless_trans)+ -done - -lemma divAlg_correct: - "[|b \<noteq> #0; a \<in> int; b \<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))" -apply (auto simp add: quorem_0 divAlg_def) -apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct - posDivAlg_type negDivAlg_type) -apply (auto simp add: quorem_def neq_iff_zless) -txt{*linear arithmetic from here on*} -apply (auto simp add: zle_def) -done - -lemma divAlg_type: "[|a \<in> int; b \<in> int|] ==> divAlg(<a,b>) \<in> int * int" -apply (auto simp add: divAlg_def) -apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type) -done - - -(** intify cancellation **) - -lemma zdiv_intify1 [simp]: "intify(x) zdiv y = x zdiv y" -apply (simp (no_asm) add: zdiv_def) -done - -lemma zdiv_intify2 [simp]: "x zdiv intify(y) = x zdiv y" -apply (simp (no_asm) add: zdiv_def) -done - -lemma zdiv_type [iff,TC]: "z zdiv w \<in> int" -apply (unfold zdiv_def) -apply (blast intro: fst_type divAlg_type) -done - -lemma zmod_intify1 [simp]: "intify(x) zmod y = x zmod y" -apply (simp (no_asm) add: zmod_def) -done - -lemma zmod_intify2 [simp]: "x zmod intify(y) = x zmod y" -apply (simp (no_asm) add: zmod_def) -done - -lemma zmod_type [iff,TC]: "z zmod w \<in> int" -apply (unfold zmod_def) -apply (rule snd_type) -apply (blast intro: divAlg_type) -done - - -(** Arbitrary definitions for division by zero. Useful to simplify - certain equations **) - -lemma DIVISION_BY_ZERO_ZDIV: "a zdiv #0 = #0" -apply (simp (no_asm) add: zdiv_def divAlg_def posDivAlg_zero_divisor) -done (*NOT for adding to default simpset*) - -lemma DIVISION_BY_ZERO_ZMOD: "a zmod #0 = intify(a)" -apply (simp (no_asm) add: zmod_def divAlg_def posDivAlg_zero_divisor) -done (*NOT for adding to default simpset*) - - - -(** Basic laws about division and remainder **) - -lemma raw_zmod_zdiv_equality: - "[| a \<in> int; b \<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (cut_tac a = "a" and b = "b" in divAlg_correct) -apply (auto simp add: quorem_def zdiv_def zmod_def split_def) -done - -lemma zmod_zdiv_equality: "intify(a) = b $* (a zdiv b) $+ (a zmod b)" -apply (rule trans) -apply (rule_tac b = "intify (b)" in raw_zmod_zdiv_equality) -apply auto -done - -lemma pos_mod: "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) -apply (blast dest: zle_zless_trans)+ -done - -lemmas pos_mod_sign = pos_mod [THEN conjunct1, standard] -and pos_mod_bound = pos_mod [THEN conjunct2, standard] - -lemma neg_mod: "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in divAlg_correct) -apply (auto simp add: intify_eq_0_iff_zle quorem_def zmod_def split_def) -apply (blast dest: zle_zless_trans) -apply (blast dest: zless_trans)+ -done - -lemmas neg_mod_sign = neg_mod [THEN conjunct1, standard] -and neg_mod_bound = neg_mod [THEN conjunct2, standard] - - -(** proving general properties of zdiv and zmod **) - -lemma quorem_div_mod: - "[|b \<noteq> #0; a \<in> int; b \<in> int |] - ==> quorem (<a,b>, <a zdiv b, a zmod b>)" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound - neg_mod_sign neg_mod_bound) -done - -(*Surely quorem(<a,b>,<q,r>) implies a \<in> int, but it doesn't matter*) -lemma quorem_div: - "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int |] - ==> a zdiv b = q" -by (blast intro: quorem_div_mod [THEN unique_quotient]) - -lemma quorem_mod: - "[| quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int |] - ==> a zmod b = r" -by (blast intro: quorem_div_mod [THEN unique_remainder]) - -lemma zdiv_pos_pos_trivial_raw: - "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zdiv b = #0" -apply (rule quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans)+ -done - -lemma zdiv_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_pos_pos_trivial_raw) -apply auto -done - -lemma zdiv_neg_neg_trivial_raw: - "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zdiv b = #0" -apply (rule_tac r = "a" in quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans zless_trans)+ -done - -lemma zdiv_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_neg_neg_trivial_raw) -apply auto -done - -lemma zadd_le_0_lemma: "[| a$+b $<= #0; #0 $< a; #0 $< b |] ==> False" -apply (drule_tac z' = "#0" and z = "b" in zadd_zless_mono) -apply (auto simp add: zle_def) -apply (blast dest: zless_trans) -done - -lemma zdiv_pos_neg_trivial_raw: - "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" -apply (rule_tac r = "a $+ b" in quorem_div) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ -done - -lemma zdiv_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zdiv b = #-1" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zdiv_pos_neg_trivial_raw) -apply auto -done - -(*There is no zdiv_neg_pos_trivial because #0 zdiv b = #0 would supersede it*) - - -lemma zmod_pos_pos_trivial_raw: - "[| a \<in> int; b \<in> int; #0 $<= a; a $< b |] ==> a zmod b = a" -apply (rule_tac q = "#0" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans)+ -done - -lemma zmod_pos_pos_trivial: "[| #0 $<= a; a $< b |] ==> a zmod b = intify(a)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_pos_pos_trivial_raw) -apply auto -done - -lemma zmod_neg_neg_trivial_raw: - "[| a \<in> int; b \<in> int; a $<= #0; b $< a |] ==> a zmod b = a" -apply (rule_tac q = "#0" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zle_zless_trans zless_trans)+ -done - -lemma zmod_neg_neg_trivial: "[| a $<= #0; b $< a |] ==> a zmod b = intify(a)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_neg_neg_trivial_raw) -apply auto -done - -lemma zmod_pos_neg_trivial_raw: - "[| a \<in> int; b \<in> int; #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" -apply (rule_tac q = "#-1" in quorem_mod) -apply (auto simp add: quorem_def) -(*linear arithmetic*) -apply (blast dest: zadd_le_0_lemma zle_zless_trans)+ -done - -lemma zmod_pos_neg_trivial: "[| #0 $< a; a$+b $<= #0 |] ==> a zmod b = a$+b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" - in zmod_pos_neg_trivial_raw) -apply auto -done - -(*There is no zmod_neg_pos_trivial...*) - - -(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*) - -lemma zdiv_zminus_zminus_raw: - "[|a \<in> int; b \<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_div]) -apply auto -done - -lemma zdiv_zminus_zminus [simp]: "($-a) zdiv ($-b) = a zdiv b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zminus_zminus_raw) -apply auto -done - -(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*) -lemma zmod_zminus_zminus_raw: - "[|a \<in> int; b \<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (subst quorem_div_mod [THEN quorem_neg, simplified, THEN quorem_mod]) -apply auto -done - -lemma zmod_zminus_zminus [simp]: "($-a) zmod ($-b) = $- (a zmod b)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zminus_zminus_raw) -apply auto -done - - -subsection{* division of a number by itself *} - -lemma self_quotient_aux1: "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q" -apply (subgoal_tac "#0 $< a$*q") -apply (cut_tac w = "#0" and z = "q" in add1_zle_iff) -apply (simp add: int_0_less_mult_iff) -apply (blast dest: zless_trans) -(*linear arithmetic...*) -apply (drule_tac t = "%x. x $- r" in subst_context) -apply (drule sym) -apply (simp add: zcompare_rls) -done - -lemma self_quotient_aux2: "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1" -apply (subgoal_tac "#0 $<= a$* (#1$-q)") - apply (simp add: int_0_le_mult_iff zcompare_rls) - apply (blast dest: zle_zless_trans) -apply (simp add: zdiff_zmult_distrib2) -apply (drule_tac t = "%x. x $- a $* q" in subst_context) -apply (simp add: zcompare_rls) -done - -lemma self_quotient: - "[| quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; a \<noteq> #0|] ==> q = #1" -apply (simp add: split_ifs quorem_def neq_iff_zless) -apply (rule zle_anti_sym) -apply safe -apply auto -prefer 4 apply (blast dest: zless_trans) -apply (blast dest: zless_trans) -apply (rule_tac [3] a = "$-a" and r = "$-r" in self_quotient_aux1) -apply (rule_tac a = "$-a" and r = "$-r" in self_quotient_aux2) -apply (rule_tac [6] zminus_equation [THEN iffD1]) -apply (rule_tac [2] zminus_equation [THEN iffD1]) -apply (force intro: self_quotient_aux1 self_quotient_aux2 - simp add: zadd_commute zmult_zminus)+ -done - -lemma self_remainder: - "[|quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0|] ==> r = #0" -apply (frule self_quotient) -apply (auto simp add: quorem_def) -done - -lemma zdiv_self_raw: "[|a \<noteq> #0; a \<in> int|] ==> a zdiv a = #1" -apply (blast intro: quorem_div_mod [THEN self_quotient]) -done - -lemma zdiv_self [simp]: "intify(a) \<noteq> #0 ==> a zdiv a = #1" -apply (drule zdiv_self_raw) -apply auto -done - -(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *) -lemma zmod_self_raw: "a \<in> int ==> a zmod a = #0" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: quorem_div_mod [THEN self_remainder]) -done - -lemma zmod_self [simp]: "a zmod a = #0" -apply (cut_tac a = "intify (a)" in zmod_self_raw) -apply auto -done - - -subsection{* Computation of division and remainder *} - -lemma zdiv_zero [simp]: "#0 zdiv b = #0" -apply (simp (no_asm) add: zdiv_def divAlg_def) -done - -lemma zdiv_eq_minus1: "#0 $< b ==> #-1 zdiv b = #-1" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -done - -lemma zmod_zero [simp]: "#0 zmod b = #0" -apply (simp (no_asm) add: zmod_def divAlg_def) -done - -lemma zdiv_minus1: "#0 $< b ==> #-1 zdiv b = #-1" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -done - -lemma zmod_minus1: "#0 $< b ==> #-1 zmod b = b $- #1" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -done - -(** a positive, b positive **) - -lemma zdiv_pos_pos: "[| #0 $< a; #0 $<= b |] - ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply (auto simp add: zle_def) -done - -lemma zmod_pos_pos: - "[| #0 $< a; #0 $<= b |] - ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply (auto simp add: zle_def) -done - -(** a negative, b positive **) - -lemma zdiv_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply (blast dest: zle_zless_trans) -done - -lemma zmod_neg_pos: - "[| a $< #0; #0 $< b |] - ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply (blast dest: zle_zless_trans) -done - -(** a positive, b negative **) - -lemma zdiv_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def intify_eq_0_iff_zle) -apply auto -apply (blast dest: zle_zless_trans)+ -apply (blast dest: zless_trans) -apply (blast intro: zless_imp_zle) -done - -lemma zmod_pos_neg: - "[| #0 $< a; b $< #0 |] - ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def intify_eq_0_iff_zle) -apply auto -apply (blast dest: zle_zless_trans)+ -apply (blast dest: zless_trans) -apply (blast intro: zless_imp_zle) -done - -(** a negative, b negative **) - -lemma zdiv_neg_neg: - "[| a $< #0; b $<= #0 |] - ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zdiv_def divAlg_def) -apply auto -apply (blast dest!: zle_zless_trans)+ -done - -lemma zmod_neg_neg: - "[| a $< #0; b $<= #0 |] - ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))" -apply (simp (no_asm_simp) add: zmod_def divAlg_def) -apply auto -apply (blast dest!: zle_zless_trans)+ -done - -declare zdiv_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zdiv_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zdiv_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zdiv_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_pos_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_neg_pos [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_pos_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare zmod_neg_neg [of "integ_of (v)" "integ_of (w)", standard, simp] -declare posDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] -declare negDivAlg_eqn [of concl: "integ_of (v)" "integ_of (w)", standard, simp] - - -(** Special-case simplification **) - -lemma zmod_1 [simp]: "a zmod #1 = #0" -apply (cut_tac a = "a" and b = "#1" in pos_mod_sign) -apply (cut_tac [2] a = "a" and b = "#1" in pos_mod_bound) -apply auto -(*arithmetic*) -apply (drule add1_zle_iff [THEN iffD2]) -apply (rule zle_anti_sym) -apply auto -done - -lemma zdiv_1 [simp]: "a zdiv #1 = intify(a)" -apply (cut_tac a = "a" and b = "#1" in zmod_zdiv_equality) -apply auto -done - -lemma zmod_minus1_right [simp]: "a zmod #-1 = #0" -apply (cut_tac a = "a" and b = "#-1" in neg_mod_sign) -apply (cut_tac [2] a = "a" and b = "#-1" in neg_mod_bound) -apply auto -(*arithmetic*) -apply (drule add1_zle_iff [THEN iffD2]) -apply (rule zle_anti_sym) -apply auto -done - -lemma zdiv_minus1_right_raw: "a \<in> int ==> a zdiv #-1 = $-a" -apply (cut_tac a = "a" and b = "#-1" in zmod_zdiv_equality) -apply auto -apply (rule equation_zminus [THEN iffD2]) -apply auto -done - -lemma zdiv_minus1_right: "a zdiv #-1 = $-a" -apply (cut_tac a = "intify (a)" in zdiv_minus1_right_raw) -apply auto -done -declare zdiv_minus1_right [simp] - - -subsection{* Monotonicity in the first argument (divisor) *} - -lemma zdiv_mono1: "[| a $<= a'; #0 $< b |] ==> a zdiv b $<= a' zdiv b" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) -apply (rule unique_quotient_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all (no_asm_simp) add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono1_neg: "[| a $<= a'; b $< #0 |] ==> a' zdiv b $<= a zdiv b" -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a'" and b = "b" in zmod_zdiv_equality) -apply (rule unique_quotient_lemma_neg) -apply (erule subst) -apply (erule subst) -apply (simp_all (no_asm_simp) add: neg_mod_sign neg_mod_bound) -done - - -subsection{* Monotonicity in the second argument (dividend) *} - -lemma q_pos_lemma: - "[| #0 $<= b'$*q' $+ r'; r' $< b'; #0 $< b' |] ==> #0 $<= q'" -apply (subgoal_tac "#0 $< b'$* (q' $+ #1)") - apply (simp add: int_0_less_mult_iff) - apply (blast dest: zless_trans intro: zless_add1_iff_zle [THEN iffD1]) -apply (simp add: zadd_zmult_distrib2) -apply (erule zle_zless_trans) -apply (erule zadd_zless_mono2) -done - -lemma zdiv_mono2_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; #0 $<= b'$*q' $+ r'; - r' $< b'; #0 $<= r; #0 $< b'; b' $<= b |] - ==> q $<= q'" -apply (frule q_pos_lemma, assumption+) -apply (subgoal_tac "b$*q $< b$* (q' $+ #1)") - apply (simp add: zmult_zless_cancel1) - apply (force dest: zless_add1_iff_zle [THEN iffD1] zless_trans zless_zle_trans) -apply (subgoal_tac "b$*q = r' $- r $+ b'$*q'") - prefer 2 apply (simp add: zcompare_rls) -apply (simp (no_asm_simp) add: zadd_zmult_distrib2) -apply (subst zadd_commute [of "b $\<times> q'"], rule zadd_zless_mono) - prefer 2 apply (blast intro: zmult_zle_mono1) -apply (subgoal_tac "r' $+ #0 $< b $+ r") - apply (simp add: zcompare_rls) -apply (rule zadd_zless_mono) - apply auto -apply (blast dest: zless_zle_trans) -done - - -lemma zdiv_mono2_raw: - "[| #0 $<= a; #0 $< b'; b' $<= b; a \<in> int |] - ==> a zdiv b $<= a zdiv b'" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) -apply (rule zdiv_mono2_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono2: - "[| #0 $<= a; #0 $< b'; b' $<= b |] - ==> a zdiv b $<= a zdiv b'" -apply (cut_tac a = "intify (a)" in zdiv_mono2_raw) -apply auto -done - -lemma q_neg_lemma: - "[| b'$*q' $+ r' $< #0; #0 $<= r'; #0 $< b' |] ==> q' $< #0" -apply (subgoal_tac "b'$*q' $< #0") - prefer 2 apply (force intro: zle_zless_trans) -apply (simp add: zmult_less_0_iff) -apply (blast dest: zless_trans) -done - - - -lemma zdiv_mono2_neg_lemma: - "[| b$*q $+ r = b'$*q' $+ r'; b'$*q' $+ r' $< #0; - r $< b; #0 $<= r'; #0 $< b'; b' $<= b |] - ==> q' $<= q" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (frule q_neg_lemma, assumption+) -apply (subgoal_tac "b$*q' $< b$* (q $+ #1)") - apply (simp add: zmult_zless_cancel1) - apply (blast dest: zless_trans zless_add1_iff_zle [THEN iffD1]) -apply (simp (no_asm_simp) add: zadd_zmult_distrib2) -apply (subgoal_tac "b$*q' $<= b'$*q'") - prefer 2 - apply (simp add: zmult_zle_cancel2) - apply (blast dest: zless_trans) -apply (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)") - prefer 2 - apply (erule ssubst) - apply simp - apply (drule_tac w' = "r" and z' = "#0" in zadd_zless_mono) - apply (assumption) - apply simp -apply (simp (no_asm_use) add: zadd_commute) -apply (rule zle_zless_trans) - prefer 2 apply (assumption) -apply (simp (no_asm_simp) add: zmult_zle_cancel2) -apply (blast dest: zless_trans) -done - -lemma zdiv_mono2_neg_raw: - "[| a $< #0; #0 $< b'; b' $<= b; a \<in> int |] - ==> a zdiv b' $<= a zdiv b" -apply (subgoal_tac "#0 $< b") - prefer 2 apply (blast dest: zless_zle_trans) -apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality) -apply (cut_tac a = "a" and b = "b'" in zmod_zdiv_equality) -apply (rule zdiv_mono2_neg_lemma) -apply (erule subst) -apply (erule subst) -apply (simp_all add: pos_mod_sign pos_mod_bound) -done - -lemma zdiv_mono2_neg: "[| a $< #0; #0 $< b'; b' $<= b |] - ==> a zdiv b' $<= a zdiv b" -apply (cut_tac a = "intify (a)" in zdiv_mono2_neg_raw) -apply auto -done - - - -subsection{* More algebraic laws for zdiv and zmod *} - -(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **) - -lemma zmult1_lemma: - "[| quorem(<b,c>, <q,r>); c \<in> int; c \<noteq> #0 |] - ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)" -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 - pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) -apply (auto intro: raw_zmod_zdiv_equality) -done - -lemma zdiv_zmult1_eq_raw: - "[|b \<in> int; c \<in> int|] - ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_div]) -apply auto -done - -lemma zdiv_zmult1_eq: "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zdiv_zmult1_eq_raw) -apply auto -done - -lemma zmod_zmult1_eq_raw: - "[|b \<in> int; c \<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zmult1_lemma, THEN quorem_mod]) -apply auto -done - -lemma zmod_zmult1_eq: "(a$*b) zmod c = a$*(b zmod c) zmod c" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult1_eq_raw) -apply auto -done - -lemma zmod_zmult1_eq': "(a$*b) zmod c = ((a zmod c) $* b) zmod c" -apply (rule trans) -apply (rule_tac b = " (b $* a) zmod c" in trans) -apply (rule_tac [2] zmod_zmult1_eq) -apply (simp_all (no_asm) add: zmult_commute) -done - -lemma zmod_zmult_distrib: "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c" -apply (rule zmod_zmult1_eq' [THEN trans]) -apply (rule zmod_zmult1_eq) -done - -lemma zdiv_zmult_self1 [simp]: "intify(b) \<noteq> #0 ==> (a$*b) zdiv b = intify(a)" -apply (simp (no_asm_simp) add: zdiv_zmult1_eq) -done - -lemma zdiv_zmult_self2 [simp]: "intify(b) \<noteq> #0 ==> (b$*a) zdiv b = intify(a)" -apply (subst zmult_commute , erule zdiv_zmult_self1) -done - -lemma zmod_zmult_self1 [simp]: "(a$*b) zmod b = #0" -apply (simp (no_asm) add: zmod_zmult1_eq) -done - -lemma zmod_zmult_self2 [simp]: "(b$*a) zmod b = #0" -apply (simp (no_asm) add: zmult_commute zmod_zmult1_eq) -done - - -(** proving (a$+b) zdiv c = - a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **) - -lemma zadd1_lemma: - "[| quorem(<a,c>, <aq,ar>); quorem(<b,c>, <bq,br>); - c \<in> int; c \<noteq> #0 |] - ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)" -apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2 - pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound) -apply (auto intro: raw_zmod_zdiv_equality) -done - -(*NOT suitable for rewriting: the RHS has an instance of the LHS*) -lemma zdiv_zadd1_eq_raw: - "[|a \<in> int; b \<in> int; c \<in> int|] ==> - (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, - THEN quorem_div]) -done - -lemma zdiv_zadd1_eq: - "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)" -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" - in zdiv_zadd1_eq_raw) -apply auto -done - -lemma zmod_zadd1_eq_raw: - "[|a \<in> int; b \<in> int; c \<in> int|] - ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (blast intro: zadd1_lemma [OF quorem_div_mod quorem_div_mod, - THEN quorem_mod]) -done - -lemma zmod_zadd1_eq: "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c" -apply (cut_tac a = "intify (a)" and b = "intify (b)" and c = "intify (c)" - in zmod_zadd1_eq_raw) -apply auto -done - -lemma zmod_div_trivial_raw: - "[|a \<in> int; b \<in> int|] ==> (a zmod b) zdiv b = #0" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound - zdiv_pos_pos_trivial neg_mod_sign neg_mod_bound zdiv_neg_neg_trivial) -done - -lemma zmod_div_trivial [simp]: "(a zmod b) zdiv b = #0" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_div_trivial_raw) -apply auto -done - -lemma zmod_mod_trivial_raw: - "[|a \<in> int; b \<in> int|] ==> (a zmod b) zmod b = a zmod b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless pos_mod_sign pos_mod_bound - zmod_pos_pos_trivial neg_mod_sign neg_mod_bound zmod_neg_neg_trivial) -done - -lemma zmod_mod_trivial [simp]: "(a zmod b) zmod b = a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_mod_trivial_raw) -apply auto -done - -lemma zmod_zadd_left_eq: "(a$+b) zmod c = ((a zmod c) $+ b) zmod c" -apply (rule trans [symmetric]) -apply (rule zmod_zadd1_eq) -apply (simp (no_asm)) -apply (rule zmod_zadd1_eq [symmetric]) -done - -lemma zmod_zadd_right_eq: "(a$+b) zmod c = (a $+ (b zmod c)) zmod c" -apply (rule trans [symmetric]) -apply (rule zmod_zadd1_eq) -apply (simp (no_asm)) -apply (rule zmod_zadd1_eq [symmetric]) -done - - -lemma zdiv_zadd_self1 [simp]: - "intify(a) \<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1" -by (simp (no_asm_simp) add: zdiv_zadd1_eq) - -lemma zdiv_zadd_self2 [simp]: - "intify(a) \<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1" -by (simp (no_asm_simp) add: zdiv_zadd1_eq) - -lemma zmod_zadd_self1 [simp]: "(a$+b) zmod a = b zmod a" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (simp (no_asm_simp) add: zmod_zadd1_eq) -done - -lemma zmod_zadd_self2 [simp]: "(b$+a) zmod a = b zmod a" -apply (case_tac "a = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (simp (no_asm_simp) add: zmod_zadd1_eq) -done - - -subsection{* proving a zdiv (b*c) = (a zdiv b) zdiv c *} - -(*The condition c>0 seems necessary. Consider that 7 zdiv ~6 = ~2 but - 7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1. The subcase (a zdiv b) zmod c = 0 seems - to cause particular problems.*) - -(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **) - -lemma zdiv_zmult2_aux1: - "[| #0 $< c; b $< r; r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r" -apply (subgoal_tac "b $* (c $- q zmod c) $< r $* #1") -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) -apply (rule zle_zless_trans) -apply (erule_tac [2] zmult_zless_mono1) -apply (rule zmult_zle_mono2_neg) -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) -apply (blast intro: zless_imp_zle dest: zless_zle_trans) -done - -lemma zdiv_zmult2_aux2: - "[| #0 $< c; b $< r; r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0" -apply (subgoal_tac "b $* (q zmod c) $<= #0") - prefer 2 - apply (simp add: zmult_le_0_iff pos_mod_sign) - apply (blast intro: zless_imp_zle dest: zless_zle_trans) -(*arithmetic*) -apply (drule zadd_zle_mono) -apply assumption -apply (simp add: zadd_commute) -done - -lemma zdiv_zmult2_aux3: - "[| #0 $< c; #0 $<= r; r $< b |] ==> #0 $<= b $* (q zmod c) $+ r" -apply (subgoal_tac "#0 $<= b $* (q zmod c)") - prefer 2 - apply (simp add: int_0_le_mult_iff pos_mod_sign) - apply (blast intro: zless_imp_zle dest: zle_zless_trans) -(*arithmetic*) -apply (drule zadd_zle_mono) -apply assumption -apply (simp add: zadd_commute) -done - -lemma zdiv_zmult2_aux4: - "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c" -apply (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)") -apply (simp add: zdiff_zmult_distrib2 zadd_commute zcompare_rls) -apply (rule zless_zle_trans) -apply (erule zmult_zless_mono1) -apply (rule_tac [2] zmult_zle_mono2) -apply (auto simp add: zcompare_rls zadd_commute add1_zle_iff pos_mod_bound) -apply (blast intro: zless_imp_zle dest: zle_zless_trans) -done - -lemma zdiv_zmult2_lemma: - "[| quorem (<a,b>, <q,r>); a \<in> int; b \<in> int; b \<noteq> #0; #0 $< c |] - ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)" -apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def - neq_iff_zless int_0_less_mult_iff - zadd_zmult_distrib2 [symmetric] zdiv_zmult2_aux1 zdiv_zmult2_aux2 - zdiv_zmult2_aux3 zdiv_zmult2_aux4) -apply (blast dest: zless_trans)+ -done - -lemma zdiv_zmult2_eq_raw: - "[|#0 $< c; a \<in> int; b \<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_div]) -apply (auto simp add: intify_eq_0_iff_zle) -apply (blast dest: zle_zless_trans) -done - -lemma zdiv_zmult2_eq: "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zdiv_zmult2_eq_raw) -apply auto -done - -lemma zmod_zmult2_eq_raw: - "[|#0 $< c; a \<in> int; b \<in> int|] - ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (rule quorem_div_mod [THEN zdiv_zmult2_lemma, THEN quorem_mod]) -apply (auto simp add: intify_eq_0_iff_zle) -apply (blast dest: zle_zless_trans) -done - -lemma zmod_zmult2_eq: - "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b" -apply (cut_tac a = "intify (a)" and b = "intify (b)" in zmod_zmult2_eq_raw) -apply auto -done - -subsection{* Cancellation of common factors in "zdiv" *} - -lemma zdiv_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (subst zdiv_zmult2_eq) -apply auto -done - -lemma zdiv_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (subgoal_tac " (c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)") -apply (rule_tac [2] zdiv_zmult_zmult1_aux1) -apply auto -done - -lemma zdiv_zmult_zmult1_raw: - "[|intify(c) \<noteq> #0; b \<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless [of b] - zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) -done - -lemma zdiv_zmult_zmult1: "intify(c) \<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b" -apply (cut_tac b = "intify (b)" in zdiv_zmult_zmult1_raw) -apply auto -done - -lemma zdiv_zmult_zmult2: "intify(c) \<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b" -apply (drule zdiv_zmult_zmult1) -apply (auto simp add: zmult_commute) -done - - -subsection{* Distribution of factors over "zmod" *} - -lemma zmod_zmult_zmult1_aux1: - "[| #0 $< b; intify(c) \<noteq> #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (subst zmod_zmult2_eq) -apply auto -done - -lemma zmod_zmult_zmult1_aux2: - "[| b $< #0; intify(c) \<noteq> #0 |] - ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (subgoal_tac " (c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))") -apply (rule_tac [2] zmod_zmult_zmult1_aux1) -apply auto -done - -lemma zmod_zmult_zmult1_raw: - "[|b \<in> int; c \<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (case_tac "b = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (case_tac "c = #0") - apply (simp add: DIVISION_BY_ZERO_ZDIV DIVISION_BY_ZERO_ZMOD) -apply (auto simp add: neq_iff_zless [of b] - zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) -done - -lemma zmod_zmult_zmult1: "(c$*a) zmod (c$*b) = c $* (a zmod b)" -apply (cut_tac b = "intify (b)" and c = "intify (c)" in zmod_zmult_zmult1_raw) -apply auto -done - -lemma zmod_zmult_zmult2: "(a$*c) zmod (b$*c) = (a zmod b) $* c" -apply (cut_tac c = "c" in zmod_zmult_zmult1) -apply (auto simp add: zmult_commute) -done - - -(** Quotients of signs **) - -lemma zdiv_neg_pos_less0: "[| a $< #0; #0 $< b |] ==> a zdiv b $< #0" -apply (subgoal_tac "a zdiv b $<= #-1") -apply (erule zle_zless_trans) -apply (simp (no_asm)) -apply (rule zle_trans) -apply (rule_tac a' = "#-1" in zdiv_mono1) -apply (rule zless_add1_iff_zle [THEN iffD1]) -apply (simp (no_asm)) -apply (auto simp add: zdiv_minus1) -done - -lemma zdiv_nonneg_neg_le0: "[| #0 $<= a; b $< #0 |] ==> a zdiv b $<= #0" -apply (drule zdiv_mono1_neg) -apply auto -done - -lemma pos_imp_zdiv_nonneg_iff: "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)" -apply auto -apply (drule_tac [2] zdiv_mono1) -apply (auto simp add: neq_iff_zless) -apply (simp (no_asm_use) add: not_zless_iff_zle [THEN iff_sym]) -apply (blast intro: zdiv_neg_pos_less0) -done - -lemma neg_imp_zdiv_nonneg_iff: "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)" -apply (subst zdiv_zminus_zminus [symmetric]) -apply (rule iff_trans) -apply (rule pos_imp_zdiv_nonneg_iff) -apply auto -done - -(*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*) -lemma pos_imp_zdiv_neg_iff: "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)" -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) -apply (erule pos_imp_zdiv_nonneg_iff) -done - -(*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*) -lemma neg_imp_zdiv_neg_iff: "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)" -apply (simp (no_asm_simp) add: not_zle_iff_zless [THEN iff_sym]) -apply (erule neg_imp_zdiv_nonneg_iff) -done - -(* - THESE REMAIN TO BE CONVERTED -- but aren't that useful! - - subsection{* Speeding up the division algorithm with shifting *} - - (** computing "zdiv" by shifting **) - - lemma pos_zdiv_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a" - apply (case_tac "a = #0") - apply (subgoal_tac "#1 $<= a") - apply (arith_tac 2) - apply (subgoal_tac "#1 $< a $* #2") - apply (arith_tac 2) - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") - apply (rule_tac [2] zmult_zle_mono2) - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) - apply (subst zdiv_zadd1_eq) - apply (simp (no_asm_simp) add: zdiv_zmult_zmult2 zmod_zmult_zmult2 zdiv_pos_pos_trivial) - apply (subst zdiv_pos_pos_trivial) - apply (simp (no_asm_simp) add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) - apply (auto simp add: zmod_pos_pos_trivial) - apply (subgoal_tac "#0 $<= b zmod a") - apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) - apply arith - done - - - lemma neg_zdiv_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a" - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)") - apply (rule_tac [2] pos_zdiv_mult_2) - apply (auto simp add: zmult_zminus_right) - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") - apply (Simp_tac 2) - apply (asm_full_simp_tac (HOL_ss add: zdiv_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) - done - - - (*Not clear why this must be proved separately; probably integ_of causes - simplification problems*) - lemma lemma: "~ #0 $<= x ==> x $<= #0" - apply auto - done - - lemma zdiv_integ_of_BIT: "integ_of (v BIT b) zdiv integ_of (w BIT False) = - (if ~b | #0 $<= integ_of w - then integ_of v zdiv (integ_of w) - else (integ_of v $+ #1) zdiv (integ_of w))" - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) - apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) - done - - declare zdiv_integ_of_BIT [simp] - - - (** computing "zmod" by shifting (proofs resemble those for "zdiv") **) - - lemma pos_zmod_mult_2: "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)" - apply (case_tac "a = #0") - apply (subgoal_tac "#1 $<= a") - apply (arith_tac 2) - apply (subgoal_tac "#1 $< a $* #2") - apply (arith_tac 2) - apply (subgoal_tac "#2$* (#1 $+ b zmod a) $<= #2$*a") - apply (rule_tac [2] zmult_zle_mono2) - apply (auto simp add: zadd_commute zmult_commute add1_zle_iff pos_mod_bound) - apply (subst zmod_zadd1_eq) - apply (simp (no_asm_simp) add: zmod_zmult_zmult2 zmod_pos_pos_trivial) - apply (rule zmod_pos_pos_trivial) - apply (simp (no_asm_simp) # add: [zmod_pos_pos_trivial pos_mod_sign [THEN zadd_zle_mono1] RSN (2,zle_trans) ]) - apply (auto simp add: zmod_pos_pos_trivial) - apply (subgoal_tac "#0 $<= b zmod a") - apply (asm_simp_tac (simpset () add: pos_mod_sign) 2) - apply arith - done - - - lemma neg_zmod_mult_2: "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1" - apply (subgoal_tac " (#1 $+ #2$* ($-b-#1)) zmod (#2$* ($-a)) = #1 $+ #2$* (($-b-#1) zmod ($-a))") - apply (rule_tac [2] pos_zmod_mult_2) - apply (auto simp add: zmult_zminus_right) - apply (subgoal_tac " (#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))") - apply (Simp_tac 2) - apply (asm_full_simp_tac (HOL_ss add: zmod_zminus_zminus zdiff_def zminus_zadd_distrib [symmetric]) - apply (dtac (zminus_equation [THEN iffD1, symmetric]) - apply auto - done - - lemma zmod_integ_of_BIT: "integ_of (v BIT b) zmod integ_of (w BIT False) = - (if b then - if #0 $<= integ_of w - then #2 $* (integ_of v zmod integ_of w) $+ #1 - else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 - else #2 $* (integ_of v zmod integ_of w))" - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) - apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) - done - - declare zmod_integ_of_BIT [simp] -*) - -ML{* -val zspos_add_zspos_imp_zspos = thm "zspos_add_zspos_imp_zspos"; -val zpos_add_zpos_imp_zpos = thm "zpos_add_zpos_imp_zpos"; -val zneg_add_zneg_imp_zneg = thm "zneg_add_zneg_imp_zneg"; -val zneg_or_0_add_zneg_or_0_imp_zneg_or_0 = thm "zneg_or_0_add_zneg_or_0_imp_zneg_or_0"; -val zero_lt_zmagnitude = thm "zero_lt_zmagnitude"; -val zless_add_succ_iff = thm "zless_add_succ_iff"; -val zadd_succ_zle_iff = thm "zadd_succ_zle_iff"; -val zless_add1_iff_zle = thm "zless_add1_iff_zle"; -val add1_zle_iff = thm "add1_zle_iff"; -val add1_left_zle_iff = thm "add1_left_zle_iff"; -val zmult_zle_mono1 = thm "zmult_zle_mono1"; -val zmult_zle_mono1_neg = thm "zmult_zle_mono1_neg"; -val zmult_zle_mono2 = thm "zmult_zle_mono2"; -val zmult_zle_mono2_neg = thm "zmult_zle_mono2_neg"; -val zmult_zle_mono = thm "zmult_zle_mono"; -val zmult_zless_mono2 = thm "zmult_zless_mono2"; -val zmult_zless_mono1 = thm "zmult_zless_mono1"; -val zmult_zless_mono = thm "zmult_zless_mono"; -val zmult_zless_mono1_neg = thm "zmult_zless_mono1_neg"; -val zmult_zless_mono2_neg = thm "zmult_zless_mono2_neg"; -val zmult_eq_0_iff = thm "zmult_eq_0_iff"; -val zmult_zless_cancel2 = thm "zmult_zless_cancel2"; -val zmult_zless_cancel1 = thm "zmult_zless_cancel1"; -val zmult_zle_cancel2 = thm "zmult_zle_cancel2"; -val zmult_zle_cancel1 = thm "zmult_zle_cancel1"; -val int_eq_iff_zle = thm "int_eq_iff_zle"; -val zmult_cancel2 = thm "zmult_cancel2"; -val zmult_cancel1 = thm "zmult_cancel1"; -val unique_quotient = thm "unique_quotient"; -val unique_remainder = thm "unique_remainder"; -val adjust_eq = thm "adjust_eq"; -val posDivAlg_termination = thm "posDivAlg_termination"; -val posDivAlg_unfold = thm "posDivAlg_unfold"; -val posDivAlg_eqn = thm "posDivAlg_eqn"; -val posDivAlg_induct = thm "posDivAlg_induct"; -val intify_eq_0_iff_zle = thm "intify_eq_0_iff_zle"; -val zmult_pos = thm "zmult_pos"; -val zmult_neg = thm "zmult_neg"; -val zmult_pos_neg = thm "zmult_pos_neg"; -val int_0_less_mult_iff = thm "int_0_less_mult_iff"; -val int_0_le_mult_iff = thm "int_0_le_mult_iff"; -val zmult_less_0_iff = thm "zmult_less_0_iff"; -val zmult_le_0_iff = thm "zmult_le_0_iff"; -val posDivAlg_type = thm "posDivAlg_type"; -val posDivAlg_correct = thm "posDivAlg_correct"; -val negDivAlg_termination = thm "negDivAlg_termination"; -val negDivAlg_unfold = thm "negDivAlg_unfold"; -val negDivAlg_eqn = thm "negDivAlg_eqn"; -val negDivAlg_induct = thm "negDivAlg_induct"; -val negDivAlg_type = thm "negDivAlg_type"; -val negDivAlg_correct = thm "negDivAlg_correct"; -val quorem_0 = thm "quorem_0"; -val posDivAlg_zero_divisor = thm "posDivAlg_zero_divisor"; -val posDivAlg_0 = thm "posDivAlg_0"; -val negDivAlg_minus1 = thm "negDivAlg_minus1"; -val negateSnd_eq = thm "negateSnd_eq"; -val negateSnd_type = thm "negateSnd_type"; -val quorem_neg = thm "quorem_neg"; -val divAlg_correct = thm "divAlg_correct"; -val divAlg_type = thm "divAlg_type"; -val zdiv_intify1 = thm "zdiv_intify1"; -val zdiv_intify2 = thm "zdiv_intify2"; -val zdiv_type = thm "zdiv_type"; -val zmod_intify1 = thm "zmod_intify1"; -val zmod_intify2 = thm "zmod_intify2"; -val zmod_type = thm "zmod_type"; -val DIVISION_BY_ZERO_ZDIV = thm "DIVISION_BY_ZERO_ZDIV"; -val DIVISION_BY_ZERO_ZMOD = thm "DIVISION_BY_ZERO_ZMOD"; -val zmod_zdiv_equality = thm "zmod_zdiv_equality"; -val pos_mod = thm "pos_mod"; -val pos_mod_sign = thm "pos_mod_sign"; -val neg_mod = thm "neg_mod"; -val neg_mod_sign = thm "neg_mod_sign"; -val quorem_div_mod = thm "quorem_div_mod"; -val quorem_div = thm "quorem_div"; -val quorem_mod = thm "quorem_mod"; -val zdiv_pos_pos_trivial = thm "zdiv_pos_pos_trivial"; -val zdiv_neg_neg_trivial = thm "zdiv_neg_neg_trivial"; -val zdiv_pos_neg_trivial = thm "zdiv_pos_neg_trivial"; -val zmod_pos_pos_trivial = thm "zmod_pos_pos_trivial"; -val zmod_neg_neg_trivial = thm "zmod_neg_neg_trivial"; -val zmod_pos_neg_trivial = thm "zmod_pos_neg_trivial"; -val zdiv_zminus_zminus = thm "zdiv_zminus_zminus"; -val zmod_zminus_zminus = thm "zmod_zminus_zminus"; -val self_quotient = thm "self_quotient"; -val self_remainder = thm "self_remainder"; -val zdiv_self = thm "zdiv_self"; -val zmod_self = thm "zmod_self"; -val zdiv_zero = thm "zdiv_zero"; -val zdiv_eq_minus1 = thm "zdiv_eq_minus1"; -val zmod_zero = thm "zmod_zero"; -val zdiv_minus1 = thm "zdiv_minus1"; -val zmod_minus1 = thm "zmod_minus1"; -val zdiv_pos_pos = thm "zdiv_pos_pos"; -val zmod_pos_pos = thm "zmod_pos_pos"; -val zdiv_neg_pos = thm "zdiv_neg_pos"; -val zmod_neg_pos = thm "zmod_neg_pos"; -val zdiv_pos_neg = thm "zdiv_pos_neg"; -val zmod_pos_neg = thm "zmod_pos_neg"; -val zdiv_neg_neg = thm "zdiv_neg_neg"; -val zmod_neg_neg = thm "zmod_neg_neg"; -val zmod_1 = thm "zmod_1"; -val zdiv_1 = thm "zdiv_1"; -val zmod_minus1_right = thm "zmod_minus1_right"; -val zdiv_minus1_right = thm "zdiv_minus1_right"; -val zdiv_mono1 = thm "zdiv_mono1"; -val zdiv_mono1_neg = thm "zdiv_mono1_neg"; -val zdiv_mono2 = thm "zdiv_mono2"; -val zdiv_mono2_neg = thm "zdiv_mono2_neg"; -val zdiv_zmult1_eq = thm "zdiv_zmult1_eq"; -val zmod_zmult1_eq = thm "zmod_zmult1_eq"; -val zmod_zmult1_eq' = thm "zmod_zmult1_eq'"; -val zmod_zmult_distrib = thm "zmod_zmult_distrib"; -val zdiv_zmult_self1 = thm "zdiv_zmult_self1"; -val zdiv_zmult_self2 = thm "zdiv_zmult_self2"; -val zmod_zmult_self1 = thm "zmod_zmult_self1"; -val zmod_zmult_self2 = thm "zmod_zmult_self2"; -val zdiv_zadd1_eq = thm "zdiv_zadd1_eq"; -val zmod_zadd1_eq = thm "zmod_zadd1_eq"; -val zmod_div_trivial = thm "zmod_div_trivial"; -val zmod_mod_trivial = thm "zmod_mod_trivial"; -val zmod_zadd_left_eq = thm "zmod_zadd_left_eq"; -val zmod_zadd_right_eq = thm "zmod_zadd_right_eq"; -val zdiv_zadd_self1 = thm "zdiv_zadd_self1"; -val zdiv_zadd_self2 = thm "zdiv_zadd_self2"; -val zmod_zadd_self1 = thm "zmod_zadd_self1"; -val zmod_zadd_self2 = thm "zmod_zadd_self2"; -val zdiv_zmult2_eq = thm "zdiv_zmult2_eq"; -val zmod_zmult2_eq = thm "zmod_zmult2_eq"; -val zdiv_zmult_zmult1 = thm "zdiv_zmult_zmult1"; -val zdiv_zmult_zmult2 = thm "zdiv_zmult_zmult2"; -val zmod_zmult_zmult1 = thm "zmod_zmult_zmult1"; -val zmod_zmult_zmult2 = thm "zmod_zmult_zmult2"; -val zdiv_neg_pos_less0 = thm "zdiv_neg_pos_less0"; -val zdiv_nonneg_neg_le0 = thm "zdiv_nonneg_neg_le0"; -val pos_imp_zdiv_nonneg_iff = thm "pos_imp_zdiv_nonneg_iff"; -val neg_imp_zdiv_nonneg_iff = thm "neg_imp_zdiv_nonneg_iff"; -val pos_imp_zdiv_neg_iff = thm "pos_imp_zdiv_neg_iff"; -val neg_imp_zdiv_neg_iff = thm "neg_imp_zdiv_neg_iff"; -*} - -end -
--- a/src/ZF/Integ/int_arith.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,443 +0,0 @@ -(* Title: ZF/Integ/int_arith.ML - ID: $Id$ - Author: Larry Paulson - Copyright 2000 University of Cambridge - -Simprocs for linear arithmetic. -*) - - -(** To simplify inequalities involving integer negation and literals, - such as -x = #3 -**) - -Addsimps [inst "y" "integ_of(?w)" zminus_equation, - inst "x" "integ_of(?w)" equation_zminus]; - -AddIffs [inst "y" "integ_of(?w)" zminus_zless, - inst "x" "integ_of(?w)" zless_zminus]; - -AddIffs [inst "y" "integ_of(?w)" zminus_zle, - inst "x" "integ_of(?w)" zle_zminus]; - -Addsimps [inst "s" "integ_of(?w)" (thm "Let_def")]; - -(*** Simprocs for numeric literals ***) - -(** Combining of literal coefficients in sums of products **) - -Goal "(x $< y) <-> (x$-y $< #0)"; -by (simp_tac (simpset() addsimps zcompare_rls) 1); -qed "zless_iff_zdiff_zless_0"; - -Goal "[| x: int; y: int |] ==> (x = y) <-> (x$-y = #0)"; -by (asm_simp_tac (simpset() addsimps zcompare_rls) 1); -qed "eq_iff_zdiff_eq_0"; - -Goal "(x $<= y) <-> (x$-y $<= #0)"; -by (asm_simp_tac (simpset() addsimps zcompare_rls) 1); -qed "zle_iff_zdiff_zle_0"; - - -(** For combine_numerals **) - -Goal "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k"; -by (simp_tac (simpset() addsimps [zadd_zmult_distrib]@zadd_ac) 1); -qed "left_zadd_zmult_distrib"; - - -(** For cancel_numerals **) - -val rel_iff_rel_0_rls = map (inst "y" "?u$+?v") - [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0, - zle_iff_zdiff_zle_0] @ - map (inst "y" "n") - [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0, - zle_iff_zdiff_zle_0]; - -Goal "(i$*u $+ m = j$*u $+ n) <-> ((i$-j)$*u $+ m = intify(n))"; -by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); -by (simp_tac (simpset() addsimps zcompare_rls) 1); -by (simp_tac (simpset() addsimps zadd_ac) 1); -qed "eq_add_iff1"; - -Goal "(i$*u $+ m = j$*u $+ n) <-> (intify(m) = (j$-i)$*u $+ n)"; -by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); -by (simp_tac (simpset() addsimps zcompare_rls) 1); -by (simp_tac (simpset() addsimps zadd_ac) 1); -qed "eq_add_iff2"; - -Goal "(i$*u $+ m $< j$*u $+ n) <-> ((i$-j)$*u $+ m $< n)"; -by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@ - zadd_ac@rel_iff_rel_0_rls) 1); -qed "less_add_iff1"; - -Goal "(i$*u $+ m $< j$*u $+ n) <-> (m $< (j$-i)$*u $+ n)"; -by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@ - zadd_ac@rel_iff_rel_0_rls) 1); -qed "less_add_iff2"; - -Goal "(i$*u $+ m $<= j$*u $+ n) <-> ((i$-j)$*u $+ m $<= n)"; -by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); -by (simp_tac (simpset() addsimps zcompare_rls) 1); -by (simp_tac (simpset() addsimps zadd_ac) 1); -qed "le_add_iff1"; - -Goal "(i$*u $+ m $<= j$*u $+ n) <-> (m $<= (j$-i)$*u $+ n)"; -by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); -by (simp_tac (simpset() addsimps zcompare_rls) 1); -by (simp_tac (simpset() addsimps zadd_ac) 1); -qed "le_add_iff2"; - - -structure Int_Numeral_Simprocs = -struct - -(*Utilities*) - -val integ_of_const = Const ("Bin.integ_of", iT --> iT); - -fun mk_numeral n = integ_of_const $ NumeralSyntax.mk_bin n; - -(*Decodes a binary INTEGER*) -fun dest_numeral (Const("Bin.integ_of", _) $ w) = - (NumeralSyntax.dest_bin w - handle Match => raise TERM("Int_Numeral_Simprocs.dest_numeral:1", [w])) - | dest_numeral t = raise TERM("Int_Numeral_Simprocs.dest_numeral:2", [t]); - -fun find_first_numeral past (t::terms) = - ((dest_numeral t, rev past @ terms) - handle TERM _ => find_first_numeral (t::past) terms) - | find_first_numeral past [] = raise TERM("find_first_numeral", []); - -val zero = mk_numeral 0; -val mk_plus = FOLogic.mk_binop "Int.zadd"; - -val iT = Ind_Syntax.iT; - -val zminus_const = Const ("Int.zminus", iT --> iT); - -(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*) -fun mk_sum [] = zero - | mk_sum [t,u] = mk_plus (t, u) - | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); - -(*this version ALWAYS includes a trailing zero*) -fun long_mk_sum [] = zero - | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); - -val dest_plus = FOLogic.dest_bin "Int.zadd" iT; - -(*decompose additions AND subtractions as a sum*) -fun dest_summing (pos, Const ("Int.zadd", _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (pos, u, ts)) - | dest_summing (pos, Const ("Int.zdiff", _) $ t $ u, ts) = - dest_summing (pos, t, dest_summing (not pos, u, ts)) - | dest_summing (pos, t, ts) = - if pos then t::ts else zminus_const$t :: ts; - -fun dest_sum t = dest_summing (true, t, []); - -val mk_diff = FOLogic.mk_binop "Int.zdiff"; -val dest_diff = FOLogic.dest_bin "Int.zdiff" iT; - -val one = mk_numeral 1; -val mk_times = FOLogic.mk_binop "Int.zmult"; - -fun mk_prod [] = one - | mk_prod [t] = t - | mk_prod (t :: ts) = if t = one then mk_prod ts - else mk_times (t, mk_prod ts); - -val dest_times = FOLogic.dest_bin "Int.zmult" iT; - -fun dest_prod t = - let val (t,u) = dest_times t - in dest_prod t @ dest_prod u end - handle TERM _ => [t]; - -(*DON'T do the obvious simplifications; that would create special cases*) -fun mk_coeff (k, t) = mk_times (mk_numeral k, t); - -(*Express t as a product of (possibly) a numeral with other sorted terms*) -fun dest_coeff sign (Const ("Int.zminus", _) $ t) = dest_coeff (~sign) t - | dest_coeff sign t = - let val ts = sort Term.term_ord (dest_prod t) - val (n, ts') = find_first_numeral [] ts - handle TERM _ => (1, ts) - in (sign*n, mk_prod ts') end; - -(*Find first coefficient-term THAT MATCHES u*) -fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) - | find_first_coeff past u (t::terms) = - let val (n,u') = dest_coeff 1 t - in if u aconv u' then (n, rev past @ terms) - else find_first_coeff (t::past) u terms - end - handle TERM _ => find_first_coeff (t::past) u terms; - - -(*Simplify #1*n and n*#1 to n*) -val add_0s = [zadd_0_intify, zadd_0_right_intify]; - -val mult_1s = [zmult_1_intify, zmult_1_right_intify, - zmult_minus1, zmult_minus1_right]; - -val tc_rules = [integ_of_type, intify_in_int, - int_of_type, zadd_type, zdiff_type, zmult_type] @ - thms "bin.intros"; -val intifys = [intify_ident, zadd_intify1, zadd_intify2, - zdiff_intify1, zdiff_intify2, zmult_intify1, zmult_intify2, - zless_intify1, zless_intify2, zle_intify1, zle_intify2]; - -(*To perform binary arithmetic*) -val bin_simps = [add_integ_of_left] @ bin_arith_simps @ bin_rel_simps; - -(*To evaluate binary negations of coefficients*) -val zminus_simps = NCons_simps @ - [integ_of_minus RS sym, - bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, - bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; - -(*To let us treat subtraction as addition*) -val diff_simps = [zdiff_def, zminus_zadd_distrib, zminus_zminus]; - -(*push the unary minus down: - x * y = x * - y *) -val int_minus_mult_eq_1_to_2 = - [zmult_zminus, zmult_zminus_right RS sym] MRS trans |> standard; - -(*to extract again any uncancelled minuses*) -val int_minus_from_mult_simps = - [zminus_zminus, zmult_zminus, zmult_zminus_right]; - -(*combine unary minus with numeric literals, however nested within a product*) -val int_mult_minus_simps = - [zmult_assoc, zmult_zminus RS sym, int_minus_mult_eq_1_to_2]; - -fun prep_simproc (name, pats, proc) = - Simplifier.simproc (the_context ()) name pats proc; - -structure CancelNumeralsCommon = - struct - val mk_sum = (fn T:typ => mk_sum) - val dest_sum = dest_sum - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff 1 - val find_first_coeff = find_first_coeff [] - fun trans_tac _ = ArithData.gen_trans_tac iff_trans - - val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac - val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys - val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - THEN ALLGOALS (SIMPSET' (fn simpset => asm_simp_tac (Simplifier.inherit_context ss simpset))) - val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) - end; - - -structure EqCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = ArithData.prove_conv "inteq_cancel_numerals" - val mk_bal = FOLogic.mk_eq - val dest_bal = FOLogic.dest_eq - val bal_add1 = eq_add_iff1 RS iff_trans - val bal_add2 = eq_add_iff2 RS iff_trans -); - -structure LessCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = ArithData.prove_conv "intless_cancel_numerals" - val mk_bal = FOLogic.mk_binrel "Int.zless" - val dest_bal = FOLogic.dest_bin "Int.zless" iT - val bal_add1 = less_add_iff1 RS iff_trans - val bal_add2 = less_add_iff2 RS iff_trans -); - -structure LeCancelNumerals = CancelNumeralsFun - (open CancelNumeralsCommon - val prove_conv = ArithData.prove_conv "intle_cancel_numerals" - val mk_bal = FOLogic.mk_binrel "Int.zle" - val dest_bal = FOLogic.dest_bin "Int.zle" iT - val bal_add1 = le_add_iff1 RS iff_trans - val bal_add2 = le_add_iff2 RS iff_trans -); - -val cancel_numerals = - map prep_simproc - [("inteq_cancel_numerals", - ["l $+ m = n", "l = m $+ n", - "l $- m = n", "l = m $- n", - "l $* m = n", "l = m $* n"], - K EqCancelNumerals.proc), - ("intless_cancel_numerals", - ["l $+ m $< n", "l $< m $+ n", - "l $- m $< n", "l $< m $- n", - "l $* m $< n", "l $< m $* n"], - K LessCancelNumerals.proc), - ("intle_cancel_numerals", - ["l $+ m $<= n", "l $<= m $+ n", - "l $- m $<= n", "l $<= m $- n", - "l $* m $<= n", "l $<= m $* n"], - K LeCancelNumerals.proc)]; - - -(*version without the hyps argument*) -fun prove_conv_nohyps name tacs sg = ArithData.prove_conv name tacs sg []; - -structure CombineNumeralsData = - struct - type coeff = IntInf.int - val iszero = (fn x : IntInf.int => x = 0) - val add = IntInf.+ - val mk_sum = (fn T:typ => long_mk_sum) (*to work for #2*x $+ #3*x *) - val dest_sum = dest_sum - val mk_coeff = mk_coeff - val dest_coeff = dest_coeff 1 - val left_distrib = left_zadd_zmult_distrib RS trans - val prove_conv = prove_conv_nohyps "int_combine_numerals" - fun trans_tac _ = ArithData.gen_trans_tac trans - - val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac @ intifys - val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys - val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) - - val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) - end; - -structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); - -val combine_numerals = - prep_simproc ("int_combine_numerals", ["i $+ j", "i $- j"], K CombineNumerals.proc); - - - -(** Constant folding for integer multiplication **) - -(*The trick is to regard products as sums, e.g. #3 $* x $* #4 as - the "sum" of #3, x, #4; the literals are then multiplied*) - - -structure CombineNumeralsProdData = - struct - type coeff = IntInf.int - val iszero = (fn x : IntInf.int => x = 0) - val add = IntInf.* - val mk_sum = (fn T:typ => mk_prod) - val dest_sum = dest_prod - fun mk_coeff(k,t) = if t=one then mk_numeral k - else raise TERM("mk_coeff", []) - fun dest_coeff t = (dest_numeral t, one) (*We ONLY want pure numerals.*) - val left_distrib = zmult_assoc RS sym RS trans - val prove_conv = prove_conv_nohyps "int_combine_numerals_prod" - fun trans_tac _ = ArithData.gen_trans_tac trans - - - -val norm_ss1 = ZF_ss addsimps mult_1s @ diff_simps @ zminus_simps - val norm_ss2 = ZF_ss addsimps [zmult_zminus_right RS sym] @ - bin_simps @ zmult_ac @ tc_rules @ intifys - fun norm_tac ss = - ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) - THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) - - val numeral_simp_ss = ZF_ss addsimps bin_simps @ tc_rules @ intifys - fun numeral_simp_tac ss = - ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - val simplify_meta_eq = ArithData.simplify_meta_eq (mult_1s); - end; - - -structure CombineNumeralsProd = CombineNumeralsFun(CombineNumeralsProdData); - -val combine_numerals_prod = - prep_simproc ("int_combine_numerals_prod", ["i $* j"], K CombineNumeralsProd.proc); - -end; - - -Addsimprocs Int_Numeral_Simprocs.cancel_numerals; -Addsimprocs [Int_Numeral_Simprocs.combine_numerals, - Int_Numeral_Simprocs.combine_numerals_prod]; - - -(*examples:*) -(* -print_depth 22; -set timing; -set trace_simp; -fun test s = (Goal s; by (Asm_simp_tac 1)); -val sg = #sign (rep_thm (topthm())); -val t = FOLogic.dest_Trueprop (Logic.strip_assums_concl(getgoal 1)); -val (t,_) = FOLogic.dest_eq t; - -(*combine_numerals_prod (products of separate literals) *) -test "#5 $* x $* #3 = y"; - -test "y2 $+ ?x42 = y $+ y2"; - -test "oo : int ==> l $+ (l $+ #2) $+ oo = oo"; - -test "#9$*x $+ y = x$*#23 $+ z"; -test "y $+ x = x $+ z"; - -test "x : int ==> x $+ y $+ z = x $+ z"; -test "x : int ==> y $+ (z $+ x) = z $+ x"; -test "z : int ==> x $+ y $+ z = (z $+ y) $+ (x $+ w)"; -test "z : int ==> x$*y $+ z = (z $+ y) $+ (y$*x $+ w)"; - -test "#-3 $* x $+ y $<= x $* #2 $+ z"; -test "y $+ x $<= x $+ z"; -test "x $+ y $+ z $<= x $+ z"; - -test "y $+ (z $+ x) $< z $+ x"; -test "x $+ y $+ z $< (z $+ y) $+ (x $+ w)"; -test "x$*y $+ z $< (z $+ y) $+ (y$*x $+ w)"; - -test "l $+ #2 $+ #2 $+ #2 $+ (l $+ #2) $+ (oo $+ #2) = uu"; -test "u : int ==> #2 $* u = u"; -test "(i $+ j $+ #12 $+ k) $- #15 = y"; -test "(i $+ j $+ #12 $+ k) $- #5 = y"; - -test "y $- b $< b"; -test "y $- (#3 $* b $+ c) $< b $- #2 $* c"; - -test "(#2 $* x $- (u $* v) $+ y) $- v $* #3 $* u = w"; -test "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u $* #4 = w"; -test "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u = w"; -test "u $* v $- (x $* u $* v $+ (u $* v) $* #4 $+ y) = w"; - -test "(i $+ j $+ #12 $+ k) = u $+ #15 $+ y"; -test "(i $+ j $* #2 $+ #12 $+ k) = j $+ #5 $+ y"; - -test "#2 $* y $+ #3 $* z $+ #6 $* w $+ #2 $* y $+ #3 $* z $+ #2 $* u = #2 $* y' $+ #3 $* z' $+ #6 $* w' $+ #2 $* y' $+ #3 $* z' $+ u $+ vv"; - -test "a $+ $-(b$+c) $+ b = d"; -test "a $+ $-(b$+c) $- b = d"; - -(*negative numerals*) -test "(i $+ j $+ #-2 $+ k) $- (u $+ #5 $+ y) = zz"; -test "(i $+ j $+ #-3 $+ k) $< u $+ #5 $+ y"; -test "(i $+ j $+ #3 $+ k) $< u $+ #-6 $+ y"; -test "(i $+ j $+ #-12 $+ k) $- #15 = y"; -test "(i $+ j $+ #12 $+ k) $- #-15 = y"; -test "(i $+ j $+ #-12 $+ k) $- #-15 = y"; - -(*Multiplying separated numerals*) -Goal "#6 $* ($# x $* #2) = uu"; -Goal "#4 $* ($# x $* $# x) $* (#2 $* $# x) = uu"; -*) -
--- a/src/ZF/Integ/twos_compl.ML Thu May 31 11:00:06 2007 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,128 +0,0 @@ -(* Title: ZF/ex/twos-compl.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge - -ML code for Arithmetic on binary integers; the model for theory Bin - - The sign Pls stands for an infinite string of leading 0s. - The sign Min stands for an infinite string of leading 1s. - -See int_of_binary for the numerical interpretation. A number can have -multiple representations, namely leading 0s with sign Pls and leading 1s with -sign Min. A number is in NORMAL FORM if it has no such extra bits. - -The representation expects that (m mod 2) is 0 or 1, even if m is negative; -For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 - -Still needs division! - -print_depth 40; -System.Control.Print.printDepth := 350; -*) - -infix 5 $$ $$$ - -(*Recursive datatype of binary integers*) -datatype bin = Pls | Min | $$ of bin * int; - -(** Conversions between bin and int **) - -fun int_of_binary Pls = 0 - | int_of_binary Min = ~1 - | int_of_binary (w$$b) = 2 * int_of_binary w + b; - -fun binary_of_int 0 = Pls - | binary_of_int ~1 = Min - | binary_of_int n = binary_of_int (n div 2) $$ (n mod 2); - -(*** Addition ***) - -(*Attach a bit while preserving the normal form. Cases left as default - are Pls$$$1 and Min$$$0. *) -fun Pls $$$ 0 = Pls - | Min $$$ 1 = Min - | v $$$ x = v$$x; - -(*Successor of an integer, assumed to be in normal form. - If w$$1 is normal then w is not Min, so bin_succ(w) $$ 0 is normal. - But Min$$0 is normal while Min$$1 is not.*) -fun bin_succ Pls = Pls$$1 - | bin_succ Min = Pls - | bin_succ (w$$1) = bin_succ(w) $$ 0 - | bin_succ (w$$0) = w $$$ 1; - -(*Predecessor of an integer, assumed to be in normal form. - If w$$0 is normal then w is not Pls, so bin_pred(w) $$ 1 is normal. - But Pls$$1 is normal while Pls$$0 is not.*) -fun bin_pred Pls = Min - | bin_pred Min = Min$$0 - | bin_pred (w$$1) = w $$$ 0 - | bin_pred (w$$0) = bin_pred(w) $$ 1; - -(*Sum of two binary integers in normal form. - Ensure last $$ preserves normal form! *) -fun bin_add (Pls, w) = w - | bin_add (Min, w) = bin_pred w - | bin_add (v$$x, Pls) = v$$x - | bin_add (v$$x, Min) = bin_pred (v$$x) - | bin_add (v$$x, w$$y) = - bin_add(v, if x+y=2 then bin_succ w else w) $$$ ((x+y) mod 2); - -(*** Subtraction ***) - -(*Unary minus*) -fun bin_minus Pls = Pls - | bin_minus Min = Pls$$1 - | bin_minus (w$$1) = bin_pred (bin_minus(w) $$$ 0) - | bin_minus (w$$0) = bin_minus(w) $$ 0; - -(*** Multiplication ***) - -(*product of two bins; a factor of 0 might cause leading 0s in result*) -fun bin_mult (Pls, _) = Pls - | bin_mult (Min, v) = bin_minus v - | bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$$ 0, v) - | bin_mult (w$$0, v) = bin_mult(w,v) $$$ 0; - -(*** Testing ***) - -(*tests addition*) -fun checksum m n = - let val wm = binary_of_int m - and wn = binary_of_int n - val wsum = bin_add(wm,wn) - in if m+n = int_of_binary wsum then (wm, wn, wsum, m+n) - else raise Match - end; - -fun bfact n = if n=0 then Pls$$1 - else bin_mult(binary_of_int n, bfact(n-1)); - -(*Examples... -bfact 5; -int_of_binary it; -bfact 69; -int_of_binary it; - -(*For {HOL,ZF}/ex/BinEx.ML*) -bin_add(binary_of_int 13, binary_of_int 19); -bin_add(binary_of_int 1234, binary_of_int 5678); -bin_add(binary_of_int 1359, binary_of_int ~2468); -bin_add(binary_of_int 93746, binary_of_int ~46375); -bin_minus(binary_of_int 65745); -bin_minus(binary_of_int ~54321); -bin_mult(binary_of_int 13, binary_of_int 19); -bin_mult(binary_of_int ~84, binary_of_int 51); -bin_mult(binary_of_int 255, binary_of_int 255); -bin_mult(binary_of_int 1359, binary_of_int ~2468); - - -(*leading zeros??*) -bin_add(binary_of_int 1234, binary_of_int ~1234); -bin_mult(binary_of_int 1234, Pls); - -(*leading ones??*) -bin_add(binary_of_int 1, binary_of_int ~2); -bin_add(binary_of_int 1234, binary_of_int ~1235); -*)
--- a/src/ZF/IsaMakefile Thu May 31 11:00:06 2007 +0200 +++ b/src/ZF/IsaMakefile Thu May 31 12:06:31 2007 +0200 @@ -28,18 +28,18 @@ FOL: @cd $(SRC)/FOL; $(ISATOOL) make FOL -$(OUT)/ZF$(ML_SUFFIX): $(OUT)/FOL$(ML_SUFFIX) AC.thy Arith.thy \ - ArithSimp.thy Bool.thy Cardinal.thy CardinalArith.thy Cardinal_AC.thy \ - Datatype.thy Epsilon.thy Finite.thy Fixedpt.thy Inductive.thy \ - InfDatatype.thy Integ/Bin.thy Integ/EquivClass.thy Integ/Int.thy \ - Integ/IntArith.thy Integ/IntDiv.thy Integ/int_arith.ML List.thy \ - Main.thy Main_ZFC.thy Nat.thy OrdQuant.thy Order.thy OrderArith.thy \ - OrderType.thy Ordinal.thy Perm.thy QPair.thy QUniv.thy ROOT.ML Sum.thy \ - Tools/cartprod.ML Tools/datatype_package.ML Tools/ind_cases.ML \ - Tools/induct_tacs.ML Tools/inductive_package.ML \ - Tools/numeral_syntax.ML Tools/primrec_package.ML Tools/typechk.ML \ - Trancl.thy Univ.thy WF.thy ZF.thy Zorn.thy arith_data.ML \ - equalities.thy func.thy ind_syntax.ML pair.thy simpdata.ML upair.thy +$(OUT)/ZF$(ML_SUFFIX): $(OUT)/FOL$(ML_SUFFIX) AC.thy Arith.thy \ + ArithSimp.thy Bin.thy Bool.thy Cardinal.thy CardinalArith.thy \ + Cardinal_AC.thy Datatype.thy Epsilon.thy EquivClass.thy Finite.thy \ + Fixedpt.thy Inductive.thy InfDatatype.thy Int.thy IntArith.thy \ + IntDiv.thy List.thy Main.thy Main_ZFC.thy Nat.thy OrdQuant.thy \ + Order.thy OrderArith.thy OrderType.thy Ordinal.thy Perm.thy QPair.thy \ + QUniv.thy ROOT.ML Sum.thy Tools/cartprod.ML Tools/datatype_package.ML \ + Tools/ind_cases.ML Tools/induct_tacs.ML Tools/inductive_package.ML \ + Tools/numeral_syntax.ML Tools/primrec_package.ML Tools/typechk.ML \ + Trancl.thy Univ.thy WF.thy ZF.thy Zorn.thy arith_data.ML \ + equalities.thy func.thy ind_syntax.ML int_arith.ML pair.thy \ + simpdata.ML upair.thy @$(ISATOOL) usedir -b -r $(OUT)/FOL ZF
--- a/src/ZF/ROOT.ML Thu May 31 11:00:06 2007 +0200 +++ b/src/ZF/ROOT.ML Thu May 31 12:06:31 2007 +0200 @@ -13,6 +13,6 @@ reset eta_contract; -with_path "Integ" use_thy "Main_ZFC"; +use_thy "Main_ZFC"; Goal "True"; (*leave subgoal package empty*)
--- a/src/ZF/Tools/numeral_syntax.ML Thu May 31 11:00:06 2007 +0200 +++ b/src/ZF/Tools/numeral_syntax.ML Thu May 31 12:06:31 2007 +0200 @@ -3,7 +3,7 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Concrete syntax for generic numerals. Assumes consts and syntax of -theory ZF/Integ/Bin. +theory Bin. *) signature NUMERAL_SYNTAX =
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/Tools/twos_compl.ML Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,128 @@ +(* Title: ZF/ex/twos-compl.ML + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1993 University of Cambridge + +ML code for Arithmetic on binary integers; the model for theory Bin + + The sign Pls stands for an infinite string of leading 0s. + The sign Min stands for an infinite string of leading 1s. + +See int_of_binary for the numerical interpretation. A number can have +multiple representations, namely leading 0s with sign Pls and leading 1s with +sign Min. A number is in NORMAL FORM if it has no such extra bits. + +The representation expects that (m mod 2) is 0 or 1, even if m is negative; +For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1 + +Still needs division! + +print_depth 40; +System.Control.Print.printDepth := 350; +*) + +infix 5 $$ $$$ + +(*Recursive datatype of binary integers*) +datatype bin = Pls | Min | $$ of bin * int; + +(** Conversions between bin and int **) + +fun int_of_binary Pls = 0 + | int_of_binary Min = ~1 + | int_of_binary (w$$b) = 2 * int_of_binary w + b; + +fun binary_of_int 0 = Pls + | binary_of_int ~1 = Min + | binary_of_int n = binary_of_int (n div 2) $$ (n mod 2); + +(*** Addition ***) + +(*Attach a bit while preserving the normal form. Cases left as default + are Pls$$$1 and Min$$$0. *) +fun Pls $$$ 0 = Pls + | Min $$$ 1 = Min + | v $$$ x = v$$x; + +(*Successor of an integer, assumed to be in normal form. + If w$$1 is normal then w is not Min, so bin_succ(w) $$ 0 is normal. + But Min$$0 is normal while Min$$1 is not.*) +fun bin_succ Pls = Pls$$1 + | bin_succ Min = Pls + | bin_succ (w$$1) = bin_succ(w) $$ 0 + | bin_succ (w$$0) = w $$$ 1; + +(*Predecessor of an integer, assumed to be in normal form. + If w$$0 is normal then w is not Pls, so bin_pred(w) $$ 1 is normal. + But Pls$$1 is normal while Pls$$0 is not.*) +fun bin_pred Pls = Min + | bin_pred Min = Min$$0 + | bin_pred (w$$1) = w $$$ 0 + | bin_pred (w$$0) = bin_pred(w) $$ 1; + +(*Sum of two binary integers in normal form. + Ensure last $$ preserves normal form! *) +fun bin_add (Pls, w) = w + | bin_add (Min, w) = bin_pred w + | bin_add (v$$x, Pls) = v$$x + | bin_add (v$$x, Min) = bin_pred (v$$x) + | bin_add (v$$x, w$$y) = + bin_add(v, if x+y=2 then bin_succ w else w) $$$ ((x+y) mod 2); + +(*** Subtraction ***) + +(*Unary minus*) +fun bin_minus Pls = Pls + | bin_minus Min = Pls$$1 + | bin_minus (w$$1) = bin_pred (bin_minus(w) $$$ 0) + | bin_minus (w$$0) = bin_minus(w) $$ 0; + +(*** Multiplication ***) + +(*product of two bins; a factor of 0 might cause leading 0s in result*) +fun bin_mult (Pls, _) = Pls + | bin_mult (Min, v) = bin_minus v + | bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$$ 0, v) + | bin_mult (w$$0, v) = bin_mult(w,v) $$$ 0; + +(*** Testing ***) + +(*tests addition*) +fun checksum m n = + let val wm = binary_of_int m + and wn = binary_of_int n + val wsum = bin_add(wm,wn) + in if m+n = int_of_binary wsum then (wm, wn, wsum, m+n) + else raise Match + end; + +fun bfact n = if n=0 then Pls$$1 + else bin_mult(binary_of_int n, bfact(n-1)); + +(*Examples... +bfact 5; +int_of_binary it; +bfact 69; +int_of_binary it; + +(*For {HOL,ZF}/ex/BinEx.ML*) +bin_add(binary_of_int 13, binary_of_int 19); +bin_add(binary_of_int 1234, binary_of_int 5678); +bin_add(binary_of_int 1359, binary_of_int ~2468); +bin_add(binary_of_int 93746, binary_of_int ~46375); +bin_minus(binary_of_int 65745); +bin_minus(binary_of_int ~54321); +bin_mult(binary_of_int 13, binary_of_int 19); +bin_mult(binary_of_int ~84, binary_of_int 51); +bin_mult(binary_of_int 255, binary_of_int 255); +bin_mult(binary_of_int 1359, binary_of_int ~2468); + + +(*leading zeros??*) +bin_add(binary_of_int 1234, binary_of_int ~1234); +bin_mult(binary_of_int 1234, Pls); + +(*leading ones??*) +bin_add(binary_of_int 1, binary_of_int ~2); +bin_add(binary_of_int 1234, binary_of_int ~1235); +*)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/ZF/int_arith.ML Thu May 31 12:06:31 2007 +0200 @@ -0,0 +1,443 @@ +(* Title: ZF/int_arith.ML + ID: $Id$ + Author: Larry Paulson + Copyright 2000 University of Cambridge + +Simprocs for linear arithmetic. +*) + + +(** To simplify inequalities involving integer negation and literals, + such as -x = #3 +**) + +Addsimps [inst "y" "integ_of(?w)" zminus_equation, + inst "x" "integ_of(?w)" equation_zminus]; + +AddIffs [inst "y" "integ_of(?w)" zminus_zless, + inst "x" "integ_of(?w)" zless_zminus]; + +AddIffs [inst "y" "integ_of(?w)" zminus_zle, + inst "x" "integ_of(?w)" zle_zminus]; + +Addsimps [inst "s" "integ_of(?w)" (thm "Let_def")]; + +(*** Simprocs for numeric literals ***) + +(** Combining of literal coefficients in sums of products **) + +Goal "(x $< y) <-> (x$-y $< #0)"; +by (simp_tac (simpset() addsimps zcompare_rls) 1); +qed "zless_iff_zdiff_zless_0"; + +Goal "[| x: int; y: int |] ==> (x = y) <-> (x$-y = #0)"; +by (asm_simp_tac (simpset() addsimps zcompare_rls) 1); +qed "eq_iff_zdiff_eq_0"; + +Goal "(x $<= y) <-> (x$-y $<= #0)"; +by (asm_simp_tac (simpset() addsimps zcompare_rls) 1); +qed "zle_iff_zdiff_zle_0"; + + +(** For combine_numerals **) + +Goal "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k"; +by (simp_tac (simpset() addsimps [zadd_zmult_distrib]@zadd_ac) 1); +qed "left_zadd_zmult_distrib"; + + +(** For cancel_numerals **) + +val rel_iff_rel_0_rls = map (inst "y" "?u$+?v") + [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0, + zle_iff_zdiff_zle_0] @ + map (inst "y" "n") + [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0, + zle_iff_zdiff_zle_0]; + +Goal "(i$*u $+ m = j$*u $+ n) <-> ((i$-j)$*u $+ m = intify(n))"; +by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); +by (simp_tac (simpset() addsimps zcompare_rls) 1); +by (simp_tac (simpset() addsimps zadd_ac) 1); +qed "eq_add_iff1"; + +Goal "(i$*u $+ m = j$*u $+ n) <-> (intify(m) = (j$-i)$*u $+ n)"; +by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); +by (simp_tac (simpset() addsimps zcompare_rls) 1); +by (simp_tac (simpset() addsimps zadd_ac) 1); +qed "eq_add_iff2"; + +Goal "(i$*u $+ m $< j$*u $+ n) <-> ((i$-j)$*u $+ m $< n)"; +by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@ + zadd_ac@rel_iff_rel_0_rls) 1); +qed "less_add_iff1"; + +Goal "(i$*u $+ m $< j$*u $+ n) <-> (m $< (j$-i)$*u $+ n)"; +by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@ + zadd_ac@rel_iff_rel_0_rls) 1); +qed "less_add_iff2"; + +Goal "(i$*u $+ m $<= j$*u $+ n) <-> ((i$-j)$*u $+ m $<= n)"; +by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); +by (simp_tac (simpset() addsimps zcompare_rls) 1); +by (simp_tac (simpset() addsimps zadd_ac) 1); +qed "le_add_iff1"; + +Goal "(i$*u $+ m $<= j$*u $+ n) <-> (m $<= (j$-i)$*u $+ n)"; +by (simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1); +by (simp_tac (simpset() addsimps zcompare_rls) 1); +by (simp_tac (simpset() addsimps zadd_ac) 1); +qed "le_add_iff2"; + + +structure Int_Numeral_Simprocs = +struct + +(*Utilities*) + +val integ_of_const = Const ("Bin.integ_of", iT --> iT); + +fun mk_numeral n = integ_of_const $ NumeralSyntax.mk_bin n; + +(*Decodes a binary INTEGER*) +fun dest_numeral (Const("Bin.integ_of", _) $ w) = + (NumeralSyntax.dest_bin w + handle Match => raise TERM("Int_Numeral_Simprocs.dest_numeral:1", [w])) + | dest_numeral t = raise TERM("Int_Numeral_Simprocs.dest_numeral:2", [t]); + +fun find_first_numeral past (t::terms) = + ((dest_numeral t, rev past @ terms) + handle TERM _ => find_first_numeral (t::past) terms) + | find_first_numeral past [] = raise TERM("find_first_numeral", []); + +val zero = mk_numeral 0; +val mk_plus = FOLogic.mk_binop "Int.zadd"; + +val iT = Ind_Syntax.iT; + +val zminus_const = Const ("Int.zminus", iT --> iT); + +(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*) +fun mk_sum [] = zero + | mk_sum [t,u] = mk_plus (t, u) + | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +(*this version ALWAYS includes a trailing zero*) +fun long_mk_sum [] = zero + | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); + +val dest_plus = FOLogic.dest_bin "Int.zadd" iT; + +(*decompose additions AND subtractions as a sum*) +fun dest_summing (pos, Const ("Int.zadd", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (pos, u, ts)) + | dest_summing (pos, Const ("Int.zdiff", _) $ t $ u, ts) = + dest_summing (pos, t, dest_summing (not pos, u, ts)) + | dest_summing (pos, t, ts) = + if pos then t::ts else zminus_const$t :: ts; + +fun dest_sum t = dest_summing (true, t, []); + +val mk_diff = FOLogic.mk_binop "Int.zdiff"; +val dest_diff = FOLogic.dest_bin "Int.zdiff" iT; + +val one = mk_numeral 1; +val mk_times = FOLogic.mk_binop "Int.zmult"; + +fun mk_prod [] = one + | mk_prod [t] = t + | mk_prod (t :: ts) = if t = one then mk_prod ts + else mk_times (t, mk_prod ts); + +val dest_times = FOLogic.dest_bin "Int.zmult" iT; + +fun dest_prod t = + let val (t,u) = dest_times t + in dest_prod t @ dest_prod u end + handle TERM _ => [t]; + +(*DON'T do the obvious simplifications; that would create special cases*) +fun mk_coeff (k, t) = mk_times (mk_numeral k, t); + +(*Express t as a product of (possibly) a numeral with other sorted terms*) +fun dest_coeff sign (Const ("Int.zminus", _) $ t) = dest_coeff (~sign) t + | dest_coeff sign t = + let val ts = sort Term.term_ord (dest_prod t) + val (n, ts') = find_first_numeral [] ts + handle TERM _ => (1, ts) + in (sign*n, mk_prod ts') end; + +(*Find first coefficient-term THAT MATCHES u*) +fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) + | find_first_coeff past u (t::terms) = + let val (n,u') = dest_coeff 1 t + in if u aconv u' then (n, rev past @ terms) + else find_first_coeff (t::past) u terms + end + handle TERM _ => find_first_coeff (t::past) u terms; + + +(*Simplify #1*n and n*#1 to n*) +val add_0s = [zadd_0_intify, zadd_0_right_intify]; + +val mult_1s = [zmult_1_intify, zmult_1_right_intify, + zmult_minus1, zmult_minus1_right]; + +val tc_rules = [integ_of_type, intify_in_int, + int_of_type, zadd_type, zdiff_type, zmult_type] @ + thms "bin.intros"; +val intifys = [intify_ident, zadd_intify1, zadd_intify2, + zdiff_intify1, zdiff_intify2, zmult_intify1, zmult_intify2, + zless_intify1, zless_intify2, zle_intify1, zle_intify2]; + +(*To perform binary arithmetic*) +val bin_simps = [add_integ_of_left] @ bin_arith_simps @ bin_rel_simps; + +(*To evaluate binary negations of coefficients*) +val zminus_simps = NCons_simps @ + [integ_of_minus RS sym, + bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, + bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; + +(*To let us treat subtraction as addition*) +val diff_simps = [zdiff_def, zminus_zadd_distrib, zminus_zminus]; + +(*push the unary minus down: - x * y = x * - y *) +val int_minus_mult_eq_1_to_2 = + [zmult_zminus, zmult_zminus_right RS sym] MRS trans |> standard; + +(*to extract again any uncancelled minuses*) +val int_minus_from_mult_simps = + [zminus_zminus, zmult_zminus, zmult_zminus_right]; + +(*combine unary minus with numeric literals, however nested within a product*) +val int_mult_minus_simps = + [zmult_assoc, zmult_zminus RS sym, int_minus_mult_eq_1_to_2]; + +fun prep_simproc (name, pats, proc) = + Simplifier.simproc (the_context ()) name pats proc; + +structure CancelNumeralsCommon = + struct + val mk_sum = (fn T:typ => mk_sum) + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val find_first_coeff = find_first_coeff [] + fun trans_tac _ = ArithData.gen_trans_tac iff_trans + + val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac + val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys + val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys + fun norm_tac ss = + ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys + fun numeral_simp_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + THEN ALLGOALS (SIMPSET' (fn simpset => asm_simp_tac (Simplifier.inherit_context ss simpset))) + val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) + end; + + +structure EqCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = ArithData.prove_conv "inteq_cancel_numerals" + val mk_bal = FOLogic.mk_eq + val dest_bal = FOLogic.dest_eq + val bal_add1 = eq_add_iff1 RS iff_trans + val bal_add2 = eq_add_iff2 RS iff_trans +); + +structure LessCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = ArithData.prove_conv "intless_cancel_numerals" + val mk_bal = FOLogic.mk_binrel "Int.zless" + val dest_bal = FOLogic.dest_bin "Int.zless" iT + val bal_add1 = less_add_iff1 RS iff_trans + val bal_add2 = less_add_iff2 RS iff_trans +); + +structure LeCancelNumerals = CancelNumeralsFun + (open CancelNumeralsCommon + val prove_conv = ArithData.prove_conv "intle_cancel_numerals" + val mk_bal = FOLogic.mk_binrel "Int.zle" + val dest_bal = FOLogic.dest_bin "Int.zle" iT + val bal_add1 = le_add_iff1 RS iff_trans + val bal_add2 = le_add_iff2 RS iff_trans +); + +val cancel_numerals = + map prep_simproc + [("inteq_cancel_numerals", + ["l $+ m = n", "l = m $+ n", + "l $- m = n", "l = m $- n", + "l $* m = n", "l = m $* n"], + K EqCancelNumerals.proc), + ("intless_cancel_numerals", + ["l $+ m $< n", "l $< m $+ n", + "l $- m $< n", "l $< m $- n", + "l $* m $< n", "l $< m $* n"], + K LessCancelNumerals.proc), + ("intle_cancel_numerals", + ["l $+ m $<= n", "l $<= m $+ n", + "l $- m $<= n", "l $<= m $- n", + "l $* m $<= n", "l $<= m $* n"], + K LeCancelNumerals.proc)]; + + +(*version without the hyps argument*) +fun prove_conv_nohyps name tacs sg = ArithData.prove_conv name tacs sg []; + +structure CombineNumeralsData = + struct + type coeff = IntInf.int + val iszero = (fn x : IntInf.int => x = 0) + val add = IntInf.+ + val mk_sum = (fn T:typ => long_mk_sum) (*to work for #2*x $+ #3*x *) + val dest_sum = dest_sum + val mk_coeff = mk_coeff + val dest_coeff = dest_coeff 1 + val left_distrib = left_zadd_zmult_distrib RS trans + val prove_conv = prove_conv_nohyps "int_combine_numerals" + fun trans_tac _ = ArithData.gen_trans_tac trans + + val norm_ss1 = ZF_ss addsimps add_0s @ mult_1s @ diff_simps @ zminus_simps @ zadd_ac @ intifys + val norm_ss2 = ZF_ss addsimps bin_simps @ int_mult_minus_simps @ intifys + val norm_ss3 = ZF_ss addsimps int_minus_from_mult_simps @ zadd_ac @ zmult_ac @ tc_rules @ intifys + fun norm_tac ss = + ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) + THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss3)) + + val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys + fun numeral_simp_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) + end; + +structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); + +val combine_numerals = + prep_simproc ("int_combine_numerals", ["i $+ j", "i $- j"], K CombineNumerals.proc); + + + +(** Constant folding for integer multiplication **) + +(*The trick is to regard products as sums, e.g. #3 $* x $* #4 as + the "sum" of #3, x, #4; the literals are then multiplied*) + + +structure CombineNumeralsProdData = + struct + type coeff = IntInf.int + val iszero = (fn x : IntInf.int => x = 0) + val add = IntInf.* + val mk_sum = (fn T:typ => mk_prod) + val dest_sum = dest_prod + fun mk_coeff(k,t) = if t=one then mk_numeral k + else raise TERM("mk_coeff", []) + fun dest_coeff t = (dest_numeral t, one) (*We ONLY want pure numerals.*) + val left_distrib = zmult_assoc RS sym RS trans + val prove_conv = prove_conv_nohyps "int_combine_numerals_prod" + fun trans_tac _ = ArithData.gen_trans_tac trans + + + +val norm_ss1 = ZF_ss addsimps mult_1s @ diff_simps @ zminus_simps + val norm_ss2 = ZF_ss addsimps [zmult_zminus_right RS sym] @ + bin_simps @ zmult_ac @ tc_rules @ intifys + fun norm_tac ss = + ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss1)) + THEN ALLGOALS (asm_simp_tac (Simplifier.inherit_context ss norm_ss2)) + + val numeral_simp_ss = ZF_ss addsimps bin_simps @ tc_rules @ intifys + fun numeral_simp_tac ss = + ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) + val simplify_meta_eq = ArithData.simplify_meta_eq (mult_1s); + end; + + +structure CombineNumeralsProd = CombineNumeralsFun(CombineNumeralsProdData); + +val combine_numerals_prod = + prep_simproc ("int_combine_numerals_prod", ["i $* j"], K CombineNumeralsProd.proc); + +end; + + +Addsimprocs Int_Numeral_Simprocs.cancel_numerals; +Addsimprocs [Int_Numeral_Simprocs.combine_numerals, + Int_Numeral_Simprocs.combine_numerals_prod]; + + +(*examples:*) +(* +print_depth 22; +set timing; +set trace_simp; +fun test s = (Goal s; by (Asm_simp_tac 1)); +val sg = #sign (rep_thm (topthm())); +val t = FOLogic.dest_Trueprop (Logic.strip_assums_concl(getgoal 1)); +val (t,_) = FOLogic.dest_eq t; + +(*combine_numerals_prod (products of separate literals) *) +test "#5 $* x $* #3 = y"; + +test "y2 $+ ?x42 = y $+ y2"; + +test "oo : int ==> l $+ (l $+ #2) $+ oo = oo"; + +test "#9$*x $+ y = x$*#23 $+ z"; +test "y $+ x = x $+ z"; + +test "x : int ==> x $+ y $+ z = x $+ z"; +test "x : int ==> y $+ (z $+ x) = z $+ x"; +test "z : int ==> x $+ y $+ z = (z $+ y) $+ (x $+ w)"; +test "z : int ==> x$*y $+ z = (z $+ y) $+ (y$*x $+ w)"; + +test "#-3 $* x $+ y $<= x $* #2 $+ z"; +test "y $+ x $<= x $+ z"; +test "x $+ y $+ z $<= x $+ z"; + +test "y $+ (z $+ x) $< z $+ x"; +test "x $+ y $+ z $< (z $+ y) $+ (x $+ w)"; +test "x$*y $+ z $< (z $+ y) $+ (y$*x $+ w)"; + +test "l $+ #2 $+ #2 $+ #2 $+ (l $+ #2) $+ (oo $+ #2) = uu"; +test "u : int ==> #2 $* u = u"; +test "(i $+ j $+ #12 $+ k) $- #15 = y"; +test "(i $+ j $+ #12 $+ k) $- #5 = y"; + +test "y $- b $< b"; +test "y $- (#3 $* b $+ c) $< b $- #2 $* c"; + +test "(#2 $* x $- (u $* v) $+ y) $- v $* #3 $* u = w"; +test "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u $* #4 = w"; +test "(#2 $* x $* u $* v $+ (u $* v) $* #4 $+ y) $- v $* u = w"; +test "u $* v $- (x $* u $* v $+ (u $* v) $* #4 $+ y) = w"; + +test "(i $+ j $+ #12 $+ k) = u $+ #15 $+ y"; +test "(i $+ j $* #2 $+ #12 $+ k) = j $+ #5 $+ y"; + +test "#2 $* y $+ #3 $* z $+ #6 $* w $+ #2 $* y $+ #3 $* z $+ #2 $* u = #2 $* y' $+ #3 $* z' $+ #6 $* w' $+ #2 $* y' $+ #3 $* z' $+ u $+ vv"; + +test "a $+ $-(b$+c) $+ b = d"; +test "a $+ $-(b$+c) $- b = d"; + +(*negative numerals*) +test "(i $+ j $+ #-2 $+ k) $- (u $+ #5 $+ y) = zz"; +test "(i $+ j $+ #-3 $+ k) $< u $+ #5 $+ y"; +test "(i $+ j $+ #3 $+ k) $< u $+ #-6 $+ y"; +test "(i $+ j $+ #-12 $+ k) $- #15 = y"; +test "(i $+ j $+ #12 $+ k) $- #-15 = y"; +test "(i $+ j $+ #-12 $+ k) $- #-15 = y"; + +(*Multiplying separated numerals*) +Goal "#6 $* ($# x $* #2) = uu"; +Goal "#4 $* ($# x $* $# x) $* (#2 $* $# x) = uu"; +*) +