]> matita.cs.unibo.it Git - helm.git/blob - matita/library/assembly/byte.ma
fixed includes and added notation for bytes
[helm.git] / matita / library / assembly / byte.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/assembly/byte".
16
17 include "assembly/exadecimal.ma".
18
19 record byte : Type ≝ {
20  bh: exadecimal;
21  bl: exadecimal
22 }.
23
24 notation "〈 x, y 〉" non associative with precedence 80 for @{ 'mk_byte $x $y }.
25 interpretation "mk_byte" 'mk_byte x y = 
26  (cic:/matita/assembly/byte/byte.ind#xpointer(1/1/1) x y).
27
28 definition eqbyte ≝
29  λb,b'. eqex (bh b) (bh b') ∧ eqex (bl b) (bl b').
30
31 definition plusbyte ≝
32  λb1,b2,c.
33   match plusex (bl b1) (bl b2) c with
34    [ couple l c' ⇒
35       match plusex (bh b1) (bh b2) c' with
36        [ couple h c'' ⇒ couple ? ? (mk_byte h l) c'' ]].
37
38 definition nat_of_byte ≝ λb:byte. 16*(bh b) + (bl b).
39
40 coercion cic:/matita/assembly/byte/nat_of_byte.con.
41
42 definition byte_of_nat ≝
43  λn. mk_byte (exadecimal_of_nat (n / 16)) (exadecimal_of_nat n).
44
45 lemma byte_of_nat_nat_of_byte: ∀b. byte_of_nat (nat_of_byte b) = b.
46  intros;
47  elim b;
48  elim e;
49  elim e1;
50  reflexivity.
51 qed.
52
53 lemma lt_nat_of_byte_256: ∀b. nat_of_byte b < 256.
54  intro;
55  unfold nat_of_byte;
56  letin H ≝ (lt_nat_of_exadecimal_16 (bh b)); clearbody H;
57  letin K ≝ (lt_nat_of_exadecimal_16 (bl b)); clearbody K;
58  unfold lt in H K ⊢ %;
59  letin H' ≝ (le_S_S_to_le ? ? H); clearbody H'; clear H;
60  letin K' ≝ (le_S_S_to_le ? ? K); clearbody K'; clear K;
61  apply le_S_S;
62  cut (16*bh b ≤ 16*15);
63   [ letin Hf ≝ (le_plus ? ? ? ? Hcut K'); clearbody Hf;
64     simplify in Hf:(? ? %);
65     assumption
66   | autobatch
67   ]
68 qed.
69
70 lemma nat_of_byte_byte_of_nat: ∀n. nat_of_byte (byte_of_nat n) = n \mod 256.
71  intro;
72  letin H ≝ (lt_nat_of_byte_256 (byte_of_nat n)); clearbody H;
73  rewrite < (lt_to_eq_mod ? ? H); clear H;
74  unfold byte_of_nat;
75  unfold nat_of_byte;
76  change with ((16*(exadecimal_of_nat (n/16)) + exadecimal_of_nat n) \mod 256 = n \mod 256);
77  letin H ≝ (div_mod n 16 ?); clearbody H; [ autobatch | ];
78  rewrite > symmetric_times in H;
79  rewrite > nat_of_exadecimal_exadecimal_of_nat in ⊢ (? ? (? (? % ?) ?) ?);
80  rewrite > nat_of_exadecimal_exadecimal_of_nat in ⊢ (? ? (? (? ? %) ?) ?);
81  rewrite > H in ⊢ (? ? ? (? % ?)); clear H;
82  rewrite > mod_plus in ⊢ (? ? % ?);
83  rewrite > mod_plus in ⊢ (? ? ? %);
84  apply eq_mod_to_eq_plus_mod;
85  rewrite < mod_mod in ⊢ (? ? ? %); [ | autobatch];
86  rewrite < mod_mod in ⊢ (? ? % ?); [ | autobatch];
87  rewrite < (eq_mod_times_times_mod ? ? 16 256) in ⊢ (? ? (? % ?) ?); [2: reflexivity | ];
88  rewrite < mod_mod in ⊢ (? ? % ?);
89   [ reflexivity
90   | autobatch
91   ].
92 qed.
93
94 lemma eq_nat_of_byte_n_nat_of_byte_mod_n_256:
95  ∀n. byte_of_nat n = byte_of_nat (n \mod 256).
96  intro;
97  unfold byte_of_nat;
98  apply eq_f2;
99   [ rewrite > exadecimal_of_nat_mod in ⊢ (? ? % ?);
100     rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %);
101     apply eq_f;
102     elim daemon
103   | rewrite > exadecimal_of_nat_mod;
104     rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %);
105     rewrite > divides_to_eq_mod_mod_mod;
106      [ reflexivity
107      | autobatch
108      ]
109   ]
110 qed.
111
112 lemma plusbyte_ok:
113  ∀b1,b2,c.
114   match plusbyte b1 b2 c with
115    [ couple r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_byte r + nat_of_bool c' * 256
116    ].
117  intros;
118  unfold plusbyte;
119  generalize in match (plusex_ok (bl b1) (bl b2) c);
120  elim (plusex (bl b1) (bl b2) c);
121  simplify in H ⊢ %;
122  generalize in match (plusex_ok (bh b1) (bh b2) t1);
123  elim (plusex (bh b1) (bh b2) t1);
124  simplify in H1 ⊢ %;
125  change in ⊢ (? ? ? (? (? % ?) ?)) with (16 * t2);
126  unfold nat_of_byte;
127  letin K ≝ (eq_f ? ? (λy.16*y) ? ? H1); clearbody K; clear H1;
128  rewrite > distr_times_plus in K:(? ? ? %);
129  rewrite > symmetric_times in K:(? ? ? (? ? (? ? %)));
130  rewrite < associative_times in K:(? ? ? (? ? %));
131  normalize in K:(? ? ? (? ? (? % ?)));
132  rewrite > symmetric_times in K:(? ? ? (? ? %));
133  rewrite > sym_plus in ⊢ (? ? ? (? % ?));
134  rewrite > associative_plus in ⊢ (? ? ? %);
135  letin K' ≝ (eq_f ? ? (plus t) ? ? K); clearbody K'; clear K;
136   apply transitive_eq; [3: apply K' | skip | ];
137  clear K';
138  rewrite > sym_plus in ⊢ (? ? (? (? ? %) ?) ?);
139  rewrite > associative_plus in ⊢ (? ? (? % ?) ?);
140  rewrite > associative_plus in ⊢ (? ? % ?);
141  rewrite > associative_plus in ⊢ (? ? (? ? %) ?);
142  rewrite > associative_plus in ⊢ (? ? (? ? (? ? %)) ?);
143  rewrite > sym_plus in ⊢ (? ? (? ? (? ? (? ? %))) ?);
144  rewrite < associative_plus in ⊢ (? ? (? ? (? ? %)) ?);
145  rewrite < associative_plus in ⊢ (? ? (? ? %) ?);
146  rewrite < associative_plus in ⊢ (? ? (? ? (? % ?)) ?);
147  rewrite > H; clear H;
148  autobatch paramodulation.
149 qed.
150
151 definition bpred ≝
152  λb.
153   match eqex (bl b) x0 with
154    [ true ⇒ mk_byte (xpred (bh b)) (xpred (bl b))
155    | false ⇒ mk_byte (bh b) (xpred (bl b))
156    ]. 
157
158 lemma plusbyte_O_x:
159  ∀b. plusbyte (mk_byte x0 x0) b false = couple ? ? b false.
160  intros;
161  elim b;
162  elim e;
163  elim e1;
164  reflexivity.
165 qed.
166
167 definition plusbytenc ≝
168  λx,y.
169   match plusbyte x y false with
170    [couple res _ ⇒ res].
171
172 definition plusbytec ≝
173  λx,y.
174   match plusbyte x y false with
175    [couple _ c ⇒ c].
176
177 lemma plusbytenc_O_x:
178  ∀x. plusbytenc (mk_byte x0 x0) x = x.
179  intros;
180  unfold plusbytenc;
181  rewrite > plusbyte_O_x;
182  reflexivity.
183 qed.
184
185 lemma eq_nat_of_byte_mod: ∀b. nat_of_byte b = nat_of_byte b \mod 256.
186  intro;
187  lapply (lt_nat_of_byte_256 b);
188  rewrite > (lt_to_eq_mod ? ? Hletin) in ⊢ (? ? ? %);
189  reflexivity.
190 qed.
191
192 theorem plusbytenc_ok:
193  ∀b1,b2:byte. nat_of_byte (plusbytenc b1 b2) = (b1 + b2) \mod 256.
194  intros;
195  unfold plusbytenc;
196  generalize in match (plusbyte_ok b1 b2 false);
197  elim (plusbyte b1 b2 false);
198  simplify in H ⊢ %;
199  change with (nat_of_byte t = (b1 + b2) \mod 256);
200  rewrite < plus_n_O in H;
201  rewrite > H; clear H;
202  rewrite > mod_plus;
203  letin K ≝ (eq_nat_of_byte_mod t); clearbody K;
204  letin K' ≝ (eq_mod_times_n_m_m_O (nat_of_bool t1) 256 ?); clearbody K';
205   [ autobatch | ];
206  autobatch paramodulation.
207 qed.
208
209
210
211 lemma eq_eqbyte_x0_x0_byte_of_nat_S_false:
212  ∀b. b < 255 → eqbyte (mk_byte x0 x0) (byte_of_nat (S b))) = false.
213  intros;
214  unfold byte_of_nat;
215  cut (b < 15 ∨ b ≥ 15);
216   [ elim Hcut;
217     [ unfold eqbyte;
218       change in ⊢ (? ? (? ? %) ?) with (eqex x0 (exadecimal_of_nat (S b))); 
219       rewrite > eq_eqex_S_x0_false;
220        [ elim (eqex (bh (mk_byte x0 x0))
221 (bh (mk_byte (exadecimal_of_nat (S b/16)) (exadecimal_of_nat (S b)))));simplify;
222 (*
223  alias id "andb_sym" = "cic:/matita/nat/propr_div_mod_lt_le_totient1_aux/andb_sym.con".
224          rewrite > andb_sym;
225 *)
226          reflexivity
227        | assumption
228        ]
229     | unfold eqbyte;
230       change in ⊢ (? ? (? % ?) ?) with (eqex x0 (exadecimal_of_nat (S b/16)));
231       letin K ≝ (leq_m_n_to_eq_div_n_m_S (S b) 16 ? ?);
232        [ autobatch
233        | unfold in H1;
234          apply le_S_S;
235          assumption
236        | clearbody K;
237          elim K; clear K;
238          rewrite > H2;
239          rewrite > eq_eqex_S_x0_false;
240           [ reflexivity
241           | unfold lt;
242             unfold lt in H;
243             rewrite < H2;
244             clear H2; clear a; clear H1; clear Hcut;
245             apply (le_times_to_le 16) [ autobatch | ] ;
246             rewrite > (div_mod (S b) 16) in H;[2:autobatch|]
247             rewrite > (div_mod 255 16) in H:(? ? %);[2:autobatch|]
248             lapply (le_to_le_plus_to_le ? ? ? ? ? H);
249             [apply lt_S_to_le;
250              apply lt_mod_m_m;autobatch
251             |rewrite > sym_times;
252              rewrite > sym_times in ⊢ (? ? %); (* just to speed up qed *)
253              normalize in \vdash (? ? %);apply Hletin;
254             ]
255           ] 
256        ]
257     ]
258   | elim (or_lt_le b 15);unfold ge;autobatch
259   ].
260 qed.
261
262 lemma eq_bpred_S_a_a:
263  ∀a. a < 255 → bpred (byte_of_nat (S a)) = byte_of_nat a.
264 (*
265  intros;
266  unfold bpred;
267  apply (bool_elim ? (eqex (bl (byte_of_nat (S a))) x0)); intros;
268   [ change with (mk_byte (xpred (bh (byte_of_nat (S a)))) (xpred (bl (byte_of_nat (S a))))
269      = byte_of_nat a);
270     rewrite > (eqex_true_to_eq ? ? H1);
271     normalize in ⊢ (? ? (? ? %) ?);
272     unfold byte_of_nat;
273     change with (mk_byte (xpred (exadecimal_of_nat (S a/16))) xF =
274                  mk_byte (exadecimal_of_nat (a/16)) (exadecimal_of_nat a));
275     
276     
277   |
278  change in ⊢ (? ? match ? % ? in bool return ? with [true\rArr ?|false\rArr ?] ?);
279  unfold byte_of_nat;
280  unfold bpred;
281  simplify;
282 *)
283 elim daemon. (*
284  intros;
285  unfold byte_of_nat;
286  cut (a \mod 16 = 15 ∨ a \mod 16 < 15);
287   [ elim Hcut;
288      [ 
289      |
290      ]
291   | autobatch
292   ].*)
293 qed.
294  
295 lemma plusbytenc_S:
296  ∀x:byte.∀n.plusbytenc (byte_of_nat (x*n)) x = byte_of_nat (x * S n).
297  intros;
298  rewrite < byte_of_nat_nat_of_byte;
299  rewrite > (plusbytenc_ok (byte_of_nat (x*n)) x);
300  rewrite < times_n_Sm;
301  rewrite > nat_of_byte_byte_of_nat in ⊢ (? ? (? (? (? % ?) ?)) ?);
302  rewrite > eq_nat_of_byte_n_nat_of_byte_mod_n_256 in ⊢ (? ? ? %);
303  rewrite > mod_plus in ⊢ (? ? (? %) ?);
304  rewrite > mod_plus in ⊢ (? ? ? (? %));
305  rewrite < mod_mod in ⊢ (? ? (? (? (? % ?) ?)) ?); [2: autobatch | ];
306  rewrite > sym_plus in ⊢ (? ? (? (? % ?)) ?);
307  reflexivity.
308 qed. 
309
310 lemma eq_plusbytec_x0_x0_x_false:
311  ∀x.plusbytec (mk_byte x0 x0) x = false.
312  intro;
313  elim x;
314  elim e;
315  elim e1;
316  reflexivity.
317 qed.