X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fdiv_and_mod.ma;h=523f7431403a14071e2154d661eb9f22ef552332;hb=8da75ffeb862139e8607c5a3a3aa782458da82cb;hp=658b07b686eb5e6c23f02f3247c64f02f8f25381;hpb=6d49a181a1b771f797d37b02661b5743aee86ac1;p=helm.git diff --git a/helm/software/matita/library/nat/div_and_mod.ma b/helm/software/matita/library/nat/div_and_mod.ma index 658b07b68..523f74314 100644 --- a/helm/software/matita/library/nat/div_and_mod.ma +++ b/helm/software/matita/library/nat/div_and_mod.ma @@ -223,6 +223,7 @@ apply (div_mod_spec_to_eq2 (q*m+r) m ((q*m+r)/ m) ((q*m+r) \mod m) q r) |apply div_mod_spec_intro[assumption|reflexivity] ] qed. + (* some properties of div and mod *) theorem div_times: \forall n,m:nat. ((S n)*m) / (S n) = m. intros. @@ -293,6 +294,50 @@ constructor 1. assumption.reflexivity. qed. +theorem mod_SO: \forall n:nat. mod n (S O) = O. +intro. +apply sym_eq. +apply le_n_O_to_eq. +apply le_S_S_to_le. +apply lt_mod_m_m. +apply le_n. +qed. + +theorem div_SO: \forall n:nat. div n (S O) = n. +intro. +rewrite > (div_mod ? (S O)) in \vdash (? ? ? %) + [rewrite > mod_SO. + rewrite < plus_n_O. + apply times_n_SO + |apply le_n + ] +qed. + +theorem or_div_mod: \forall n,q. O < q \to +((S (n \mod q)=q) \land S n = (S (div n q)) * q \lor +((S (n \mod q) sym_plus. + rewrite < H1 in ⊢ (? ? ? (? ? %)). + rewrite < plus_n_Sm. + apply eq_f. + apply div_mod. + assumption + ] + ] +qed. + (* injectivity *) theorem injective_times_r: \forall n:nat.injective nat nat (\lambda m:nat.(S n)*m). change with (\forall n,p,q:nat.(S n)*p = (S n)*q \to p=q). @@ -338,6 +383,7 @@ qed. variant inj_times_l1:\forall n. O < n \to \forall p,q:nat.p*n = q*n \to p=q \def lt_O_to_injective_times_l. + (* n_divides computes the pair (div,mod) *) (* p is just an upper bound, acc is an accumulator *)