]> matita.cs.unibo.it Git - helm.git/blob - helm/matita/library/nat/congruence.ma
ocaml 3.09 transition
[helm.git] / helm / matita / library / nat / congruence.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        Matita is distributed under the terms of the          *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/congruence".
16
17 include "nat/relevant_equations.ma".
18 include "nat/primes.ma".
19
20 definition S_mod: nat \to nat \to nat \def
21 \lambda n,m:nat. (S m) \mod n.
22
23 definition congruent: nat \to nat \to nat \to Prop \def
24 \lambda n,m,p:nat. mod n p = mod m p.
25
26 theorem congruent_n_n: \forall n,p:nat.congruent n n p.
27 intros.unfold congruent.reflexivity.
28 qed.
29
30 theorem transitive_congruent: \forall p:nat. transitive nat 
31 (\lambda n,m. congruent n m p).
32 intros.unfold transitive.unfold congruent.intros.
33 whd.apply (trans_eq ? ? (y \mod p)).
34 apply H.apply H1.
35 qed.
36
37 theorem le_to_mod: \forall n,m:nat. n \lt m \to n = n \mod m.
38 intros.
39 apply (div_mod_spec_to_eq2 n m O n (n/m) (n \mod m)).
40 constructor 1.assumption.simplify.reflexivity.
41 apply div_mod_spec_div_mod.
42 apply (le_to_lt_to_lt O n m).apply le_O_n.assumption.
43 qed.
44
45 theorem mod_mod : \forall n,p:nat. O<p \to n \mod p = (n \mod p) \mod p.
46 intros.
47 rewrite > (div_mod (n \mod p) p) in \vdash (? ? % ?).
48 rewrite > (eq_div_O ? p).reflexivity.
49 (* uffa: hint non lo trova lt vs. le*)
50 apply lt_mod_m_m.
51 assumption.
52 assumption.
53 qed.
54
55 theorem mod_times_mod : \forall n,m,p:nat. O<p \to O<m \to n \mod p = (n \mod (m*p)) \mod p.
56 intros.
57 apply (div_mod_spec_to_eq2 n p (n/p) (n \mod p) 
58 (n/(m*p)*m + (n \mod (m*p)/p))).
59 apply div_mod_spec_div_mod.assumption.
60 constructor 1.
61 apply lt_mod_m_m.assumption.
62 rewrite > times_plus_l.
63 rewrite > assoc_plus.
64 rewrite < div_mod.
65 rewrite > assoc_times.
66 rewrite < div_mod.
67 reflexivity.
68 rewrite > (times_n_O O).
69 apply lt_times.
70 assumption.assumption.assumption.
71 qed.
72
73 theorem congruent_n_mod_n : 
74 \forall n,p:nat. O < p \to congruent n (n \mod p) p.
75 intros.unfold congruent.
76 apply mod_mod.assumption.
77 qed.
78
79 theorem congruent_n_mod_times : 
80 \forall n,m,p:nat. O < p \to O < m \to congruent n (n \mod (m*p)) p.
81 intros.unfold congruent.
82 apply mod_times_mod.assumption.assumption.
83 qed.
84
85 theorem eq_times_plus_to_congruent: \forall n,m,p,r:nat. O< p \to 
86 n = r*p+m \to congruent n m p.
87 intros.unfold congruent.
88 apply (div_mod_spec_to_eq2 n p (div n p) (mod n p) (r +(div m p)) (mod m p)).
89 apply div_mod_spec_div_mod.assumption.
90 constructor 1.
91 apply lt_mod_m_m.assumption.
92 rewrite > sym_times.
93 rewrite > distr_times_plus.
94 rewrite > sym_times.
95 rewrite > (sym_times p).
96 rewrite > assoc_plus.
97 rewrite < div_mod.
98 assumption.assumption.
99 qed.
100
101 theorem divides_to_congruent: \forall n,m,p:nat. O < p \to m \le n \to 
102 divides p (n - m) \to congruent n m p.
103 intros.elim H2.
104 apply (eq_times_plus_to_congruent n m p n2).
105 assumption.
106 rewrite < sym_plus.
107 apply minus_to_plus.assumption.
108 rewrite > sym_times. assumption.
109 qed.
110
111 theorem congruent_to_divides: \forall n,m,p:nat.
112 O < p \to congruent n m p \to divides p (n - m).
113 intros.unfold congruent in H1.
114 apply (witness ? ? ((n / p)-(m / p))).
115 rewrite > sym_times.
116 rewrite > (div_mod n p) in \vdash (? ? % ?).
117 rewrite > (div_mod m p) in \vdash (? ? % ?).
118 rewrite < (sym_plus (m \mod p)).
119 rewrite < H1.
120 rewrite < (eq_minus_minus_minus_plus ? (n \mod p)).
121 rewrite < minus_plus_m_m.
122 apply sym_eq.
123 apply times_minus_l.
124 assumption.assumption.
125 qed.
126
127 theorem mod_times: \forall n,m,p:nat. 
128 O < p \to mod (n*m) p = mod ((mod n p)*(mod m p)) p.
129 intros.
130 change with (congruent (n*m) ((mod n p)*(mod m p)) p).
131 apply (eq_times_plus_to_congruent ? ? p 
132 ((n / p)*p*(m / p) + (n / p)*(m \mod p) + (n \mod p)*(m / p))).
133 assumption.
134 apply (trans_eq ? ? (((n/p)*p+(n \mod p))*((m/p)*p+(m \mod p)))).
135 apply eq_f2.
136 apply div_mod.assumption.
137 apply div_mod.assumption.
138 apply (trans_eq ? ? (((n/p)*p)*((m/p)*p) + (n/p)*p*(m \mod p) +
139 (n \mod p)*((m / p)*p) + (n \mod p)*(m \mod p))).
140 apply times_plus_plus.
141 apply eq_f2.
142 rewrite < assoc_times.
143 rewrite > (assoc_times (n/p) p (m \mod p)).
144 rewrite > (sym_times p (m \mod p)).
145 rewrite < (assoc_times (n/p) (m \mod p) p).
146 rewrite < times_plus_l.
147 rewrite < (assoc_times (n \mod p)).
148 rewrite < times_plus_l.
149 apply eq_f2.
150 apply eq_f2.reflexivity.
151 reflexivity.reflexivity.
152 reflexivity.
153 qed.
154
155 theorem congruent_times: \forall n,m,n1,m1,p. O < p \to congruent n n1 p \to 
156 congruent m m1 p \to congruent (n*m) (n1*m1) p.
157 unfold congruent. 
158 intros. 
159 rewrite > (mod_times n m p H).
160 rewrite > H1.
161 rewrite > H2.
162 apply sym_eq.
163 apply mod_times.assumption.
164 qed.
165
166 theorem congruent_pi: \forall f:nat \to nat. \forall n,m,p:nat.O < p \to
167 congruent (pi n f m) (pi n (\lambda m. mod (f m) p) m) p.
168 intros.
169 elim n.change with (congruent (f m) (f m \mod p) p).
170 apply congruent_n_mod_n.assumption.
171 change with (congruent ((f (S n1+m))*(pi n1 f m)) 
172 (((f (S n1+m))\mod p)*(pi n1 (\lambda m.(f m) \mod p) m)) p).
173 apply congruent_times.
174 assumption.
175 apply congruent_n_mod_n.assumption.
176 assumption.
177 qed.