X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fcongruence.ma;h=0a635ff78d9dc98ce0e68bae63b028a218ef7b64;hb=73a66cce6e72c654fdcd0ce760c405a74af70d08;hp=af744cf3475087290d2100638f3a6af3b729c732;hpb=55b82bd235d82ff7f0a40d980effe1efde1f5073;p=helm.git diff --git a/helm/software/matita/library/nat/congruence.ma b/helm/software/matita/library/nat/congruence.ma index af744cf34..0a635ff78 100644 --- a/helm/software/matita/library/nat/congruence.ma +++ b/helm/software/matita/library/nat/congruence.ma @@ -1,5 +1,5 @@ (**************************************************************************) -(* ___ *) +(* ___ *) (* ||M|| *) (* ||A|| A project by Andrea Asperti *) (* ||T|| *) @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/congruence". - include "nat/relevant_equations.ma". include "nat/primes.ma". @@ -23,6 +21,8 @@ definition S_mod: nat \to nat \to nat \def definition congruent: nat \to nat \to nat \to Prop \def \lambda n,m,p:nat. mod n p = mod m p. +interpretation "congruent" 'congruent n m p = (congruent n m p). + theorem congruent_n_n: \forall n,p:nat.congruent n n p. intros.unfold congruent.reflexivity. qed. @@ -89,6 +89,18 @@ apply (div_mod_spec_to_eq2 n p (div n p) (mod n p) (r +(div m p)) (mod m p)). apply div_mod_spec_div_mod.assumption. constructor 1. apply lt_mod_m_m.assumption. +(*cut (n = r * p + (m / p * p + m \mod p)).*) +(*lapply (div_mod m p H). +rewrite > sym_times. +rewrite > distr_times_plus. +(*rewrite > (sym_times p (m/p)).*) +(*rewrite > sym_times.*) +rewrite > assoc_plus. +autobatch paramodulation. +rewrite < div_mod. +assumption. +assumption. +*) rewrite > sym_times. rewrite > distr_times_plus. rewrite > sym_times. @@ -101,7 +113,7 @@ qed. theorem divides_to_congruent: \forall n,m,p:nat. O < p \to m \le n \to divides p (n - m) \to congruent n m p. intros.elim H2. -apply (eq_times_plus_to_congruent n m p n2). +apply (eq_times_plus_to_congruent n m p n1). assumption. rewrite < sym_plus. apply minus_to_plus.assumption. @@ -166,10 +178,9 @@ qed. theorem congruent_pi: \forall f:nat \to nat. \forall n,m,p:nat.O < p \to congruent (pi n f m) (pi n (\lambda m. mod (f m) p) m) p. intros. -elim n.change with (congruent (f m) (f m \mod p) p). +elim n. simplify. apply congruent_n_mod_n.assumption. -change with (congruent ((f (S n1+m))*(pi n1 f m)) -(((f (S n1+m))\mod p)*(pi n1 (\lambda m.(f m) \mod p) m)) p). +simplify. apply congruent_times. assumption. apply congruent_n_mod_n.assumption.