]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/nat/congruence.ma
4266d3f6ffb22412d869bd4a357237937e3dcc67
[helm.git] / helm / software / 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 include "nat/relevant_equations.ma".
16 include "nat/primes.ma".
17
18 definition S_mod: nat \to nat \to nat \def
19 \lambda n,m:nat. (S m) \mod n.
20
21 definition congruent: nat \to nat \to nat \to Prop \def
22 \lambda n,m,p:nat. mod n p = mod m p.
23
24 interpretation "congruent" 'congruent n m p = (congruent n m p).
25
26 notation < "hvbox(n break \cong\sub p m)"
27   (*non associative*) with precedence 45
28 for @{ 'congruent $n $m $p }.
29
30 theorem congruent_n_n: \forall n,p:nat.congruent n n p.
31 intros.unfold congruent.reflexivity.
32 qed.
33
34 theorem transitive_congruent: \forall p:nat. transitive nat 
35 (\lambda n,m. congruent n m p).
36 intros.unfold transitive.unfold congruent.intros.
37 whd.apply (trans_eq ? ? (y \mod p)).
38 apply H.apply H1.
39 qed.
40
41 theorem le_to_mod: \forall n,m:nat. n \lt m \to n = n \mod m.
42 intros.
43 apply (div_mod_spec_to_eq2 n m O n (n/m) (n \mod m)).
44 constructor 1.assumption.simplify.reflexivity.
45 apply div_mod_spec_div_mod.
46 apply (le_to_lt_to_lt O n m).apply le_O_n.assumption.
47 qed.
48
49 theorem mod_mod : \forall n,p:nat. O<p \to n \mod p = (n \mod p) \mod p.
50 intros.
51 rewrite > (div_mod (n \mod p) p) in \vdash (? ? % ?).
52 rewrite > (eq_div_O ? p).reflexivity.
53 (* uffa: hint non lo trova lt vs. le*)
54 apply lt_mod_m_m.
55 assumption.
56 assumption.
57 qed.
58
59 theorem mod_times_mod : \forall n,m,p:nat. O<p \to O<m \to n \mod p = (n \mod (m*p)) \mod p.
60 intros.
61 apply (div_mod_spec_to_eq2 n p (n/p) (n \mod p) 
62 (n/(m*p)*m + (n \mod (m*p)/p))).
63 apply div_mod_spec_div_mod.assumption.
64 constructor 1.
65 apply lt_mod_m_m.assumption.
66 rewrite > times_plus_l.
67 rewrite > assoc_plus.
68 rewrite < div_mod.
69 rewrite > assoc_times.
70 rewrite < div_mod.
71 reflexivity.
72 rewrite > (times_n_O O).
73 apply lt_times.
74 assumption.assumption.assumption.
75 qed.
76
77 theorem congruent_n_mod_n : 
78 \forall n,p:nat. O < p \to congruent n (n \mod p) p.
79 intros.unfold congruent.
80 apply mod_mod.assumption.
81 qed.
82
83 theorem congruent_n_mod_times : 
84 \forall n,m,p:nat. O < p \to O < m \to congruent n (n \mod (m*p)) p.
85 intros.unfold congruent.
86 apply mod_times_mod.assumption.assumption.
87 qed.
88
89 theorem eq_times_plus_to_congruent: \forall n,m,p,r:nat. O< p \to 
90 n = r*p+m \to congruent n m p.
91 intros.unfold congruent.
92 apply (div_mod_spec_to_eq2 n p (div n p) (mod n p) (r +(div m p)) (mod m p)).
93 apply div_mod_spec_div_mod.assumption.
94 constructor 1.
95 apply lt_mod_m_m.assumption.
96 (*cut (n = r * p + (m / p * p + m \mod p)).*)
97 (*lapply (div_mod m p H). 
98 rewrite > sym_times.
99 rewrite > distr_times_plus.
100 (*rewrite > (sym_times p (m/p)).*)
101 (*rewrite > sym_times.*)
102 rewrite > assoc_plus.
103 autobatch paramodulation.
104 rewrite < div_mod.
105 assumption.
106 assumption.
107 *)
108 rewrite > sym_times.
109 rewrite > distr_times_plus.
110 rewrite > sym_times.
111 rewrite > (sym_times p).
112 rewrite > assoc_plus.
113 rewrite < div_mod.
114 assumption.assumption.
115 qed.
116
117 theorem divides_to_congruent: \forall n,m,p:nat. O < p \to m \le n \to 
118 divides p (n - m) \to congruent n m p.
119 intros.elim H2.
120 apply (eq_times_plus_to_congruent n m p n1).
121 assumption.
122 rewrite < sym_plus.
123 apply minus_to_plus.assumption.
124 rewrite > sym_times. assumption.
125 qed.
126
127 theorem congruent_to_divides: \forall n,m,p:nat.
128 O < p \to congruent n m p \to divides p (n - m).
129 intros.unfold congruent in H1.
130 apply (witness ? ? ((n / p)-(m / p))).
131 rewrite > sym_times.
132 rewrite > (div_mod n p) in \vdash (? ? % ?).
133 rewrite > (div_mod m p) in \vdash (? ? % ?).
134 rewrite < (sym_plus (m \mod p)).
135 rewrite < H1.
136 rewrite < (eq_minus_minus_minus_plus ? (n \mod p)).
137 rewrite < minus_plus_m_m.
138 apply sym_eq.
139 apply times_minus_l.
140 assumption.assumption.
141 qed.
142
143 theorem mod_times: \forall n,m,p:nat. 
144 O < p \to mod (n*m) p = mod ((mod n p)*(mod m p)) p.
145 intros.
146 change with (congruent (n*m) ((mod n p)*(mod m p)) p).
147 apply (eq_times_plus_to_congruent ? ? p 
148 ((n / p)*p*(m / p) + (n / p)*(m \mod p) + (n \mod p)*(m / p))).
149 assumption.
150 apply (trans_eq ? ? (((n/p)*p+(n \mod p))*((m/p)*p+(m \mod p)))).
151 apply eq_f2.
152 apply div_mod.assumption.
153 apply div_mod.assumption.
154 apply (trans_eq ? ? (((n/p)*p)*((m/p)*p) + (n/p)*p*(m \mod p) +
155 (n \mod p)*((m / p)*p) + (n \mod p)*(m \mod p))).
156 apply times_plus_plus.
157 apply eq_f2.
158 rewrite < assoc_times.
159 rewrite > (assoc_times (n/p) p (m \mod p)).
160 rewrite > (sym_times p (m \mod p)).
161 rewrite < (assoc_times (n/p) (m \mod p) p).
162 rewrite < times_plus_l.
163 rewrite < (assoc_times (n \mod p)).
164 rewrite < times_plus_l.
165 apply eq_f2.
166 apply eq_f2.reflexivity.
167 reflexivity.reflexivity.
168 reflexivity.
169 qed.
170
171 theorem congruent_times: \forall n,m,n1,m1,p. O < p \to congruent n n1 p \to 
172 congruent m m1 p \to congruent (n*m) (n1*m1) p.
173 unfold congruent. 
174 intros. 
175 rewrite > (mod_times n m p H).
176 rewrite > H1.
177 rewrite > H2.
178 apply sym_eq.
179 apply mod_times.assumption.
180 qed.
181
182 theorem congruent_pi: \forall f:nat \to nat. \forall n,m,p:nat.O < p \to
183 congruent (pi n f m) (pi n (\lambda m. mod (f m) p) m) p.
184 intros.
185 elim n. simplify.
186 apply congruent_n_mod_n.assumption.
187 simplify.
188 apply congruent_times.
189 assumption.
190 apply congruent_n_mod_n.assumption.
191 assumption.
192 qed.