X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fnat%2Fcongruence.ma;h=f418c1b8578a0a7de9e983a96c7291fd28fc73c4;hb=8ee0e6f729105eaf1907de0baef22e170b0d17b3;hp=af744cf3475087290d2100638f3a6af3b729c732;hpb=7f2444c2670cadafddd8785b687ef312158376b0;p=helm.git diff --git a/matita/library/nat/congruence.ma b/matita/library/nat/congruence.ma index af744cf34..f418c1b85 100644 --- a/matita/library/nat/congruence.ma +++ b/matita/library/nat/congruence.ma @@ -23,6 +23,13 @@ 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 = + (cic:/matita/nat/congruence/congruent.con n m p). + +notation < "hvbox(n break \cong\sub p m)" + (*non associative*) with precedence 45 +for @{ 'congruent $n $m $p }. + theorem congruent_n_n: \forall n,p:nat.congruent n n p. intros.unfold congruent.reflexivity. qed. @@ -89,6 +96,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. @@ -166,10 +185,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.