]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/div_and_mod.ma
fixed propagation under Fix/Lambda/Case of coercions, better names are
[helm.git] / matita / library / nat / div_and_mod.ma
index 2f186dd31ae1ff0cf31eb5971c59c388d14ba349..d7750e39ad7c03e4424d1d221e50710843d7e87a 100644 (file)
 
 set "baseuri" "cic:/matita/nat/div_and_mod".
 
+include "datatypes/constructors.ma".
 include "nat/minus.ma".
 
+
 let rec mod_aux p m n: nat \def
 match (leb m n) with
 [ true \Rightarrow m
@@ -27,7 +29,7 @@ match (leb m n) with
 definition mod : nat \to nat \to nat \def
 \lambda n,m.
 match m with 
-[O \Rightarrow m
+[O \Rightarrow n
 | (S p) \Rightarrow mod_aux n n p]. 
 
 interpretation "natural remainder" 'module x y =
@@ -142,7 +144,7 @@ rewrite > distr_times_minus.
 rewrite > plus_minus.
 rewrite > sym_times.
 rewrite < H5.
-rewrite < sym_times.
+rewrite < sym_times. 
 apply plus_to_minus.
 apply H3.
 apply le_times_r.
@@ -192,14 +194,34 @@ unfold lt.apply le_S_S.apply le_O_n.
 rewrite < plus_n_O.rewrite < sym_times.reflexivity.
 qed.
 
+lemma div_plus_times: \forall m,q,r:nat. r < m \to  (q*m+r)/ m = q. 
+intros.
+apply (div_mod_spec_to_eq (q*m+r) m ? ((q*m+r) \mod m) ? r)
+  [apply div_mod_spec_div_mod.
+   apply (le_to_lt_to_lt ? r)
+    [apply le_O_n|assumption]
+  |apply div_mod_spec_intro[assumption|reflexivity]
+  ]
+qed.
+
+lemma mod_plus_times: \forall m,q,r:nat. r < m \to  (q*m+r) \mod m = r. 
+intros.
+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_div_mod.
+   apply (le_to_lt_to_lt ? r)
+    [apply le_O_n|assumption]
+  |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.
-apply (div_mod_spec_to_eq ((S n)*m) (S n) ? ? ? O).
-goal 15. (* ?11 is closed with the following tactics *)
-apply div_mod_spec_div_mod.
-unfold lt.apply le_S_S.apply le_O_n.
-apply div_mod_spec_times.
+apply (div_mod_spec_to_eq ((S n)*m) (S n) ? ? ? O);
+[2: apply div_mod_spec_div_mod.
+    unfold lt.apply le_S_S.apply le_O_n.
+|   skip
+|   apply div_mod_spec_times
+]
 qed.
 
 theorem div_n_n: \forall n:nat. O < n \to n / n = S O.
@@ -296,3 +318,28 @@ 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 *)
+let rec n_divides_aux p n m acc \def
+  match n \mod m with
+  [ O \Rightarrow 
+    match p with
+      [ O \Rightarrow pair nat nat acc n
+      | (S p) \Rightarrow n_divides_aux p (n / m) m (S acc)]
+  | (S a) \Rightarrow pair nat nat acc n].
+
+(* n_divides n m = <q,r> if m divides n q times, with remainder r *)
+definition n_divides \def \lambda n,m:nat.n_divides_aux n n m O.
+
+
+(*a simple variant of div_times theorem *)
+theorem div_times_ltO: \forall a,b:nat. O \lt b \to
+a*b/b = a.
+intros.
+rewrite > sym_times.
+rewrite > (S_pred b H).
+apply div_times.
+qed.
+