]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/div_and_mod.ma
Notation for equality used everywhere.
[helm.git] / helm / matita / library / nat / div_and_mod.ma
index 2b00a97abcf4c867d7b24bb5704beaa6d2ee562e..4c43d33bd9d1d57ba4b283016c850f0cb0bf7a29 100644 (file)
@@ -71,7 +71,7 @@ apply le_n.
 qed.
 
 theorem div_aux_mod_aux: \forall p,n,m:nat. 
-(eq nat n (plus (times (div_aux p n m) (S m)) (mod_aux p n m) )).
+(n=plus (times (div_aux p n m) (S m)) (mod_aux p n m)).
 intro.elim p.
 simplify.elim leb n m.
 simplify.apply refl_eq.
@@ -82,7 +82,7 @@ simplify.intro.apply refl_eq.
 simplify.intro.
 rewrite > assoc_plus. 
 elim (H (minus n1 (S m)) m).
-change with (eq nat n1 (plus (S m) (minus n1 (S m)))).
+change with (n1=plus (S m) (minus n1 (S m))).
 rewrite < sym_plus.
 apply plus_minus_m_m.
 change with lt m n1.
@@ -90,7 +90,7 @@ apply not_le_to_lt.exact H1.
 qed.
 
 theorem div_mod: \forall n,m:nat. 
-(lt O m) \to (eq nat n (plus (times (div n m) m) (mod n m))).
+(lt O m) \to n=plus (times (div n m) m) (mod n m).
 intros 2.elim m.elim (not_le_Sn_O O H).
 simplify.
 apply div_aux_mod_aux.
@@ -98,14 +98,14 @@ qed.
 
 inductive div_mod_spec (n,m,q,r:nat) : Prop \def
 div_mod_spec_intro: 
-(lt r m) \to (eq nat n (plus (times q m) r)) \to (div_mod_spec n m q r).
+(lt r m) \to n=plus (times q m) r \to (div_mod_spec n m q r).
 
 (* 
 definition div_mod_spec : nat \to nat \to nat \to nat \to Prop \def
-\lambda n,m,q,r:nat.(And (lt r m) (eq nat n (plus (times q m) r))).
+\lambda n,m,q,r:nat.(And (lt r m) n=plus (times q m) r).
 *)
 
-theorem div_mod_spec_to_not_eq_O: \forall n,m,q,r.(div_mod_spec n m q r) \to Not (eq nat m O).
+theorem div_mod_spec_to_not_eq_O: \forall n,m,q,r.(div_mod_spec n m q r) \to Not (m=O).
 intros 4.simplify.intros.elim H.absurd le (S r) O.
 rewrite < H1.assumption.
 exact not_le_Sn_O r.
@@ -175,4 +175,4 @@ apply H5.
 apply le_times_r.
 apply lt_to_le.
 apply H6.
-qed.
\ No newline at end of file
+qed.