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