]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/congruence.ma
Dummy dependent types are no longer cleaned in inductive type arities.
[helm.git] / helm / software / matita / library / nat / congruence.ma
index 56a9f2ab59cd569e2078b8beaf06cf9f268ff9c7..753745d4540d23f1378fcd3f4ac9cbc2791f14e0 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/congruence".
-
 include "nat/relevant_equations.ma".
 include "nat/primes.ma".
 
@@ -23,6 +21,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 +94,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.