]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/word16.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / num / word16.ma
index 965fb4b38c1921758449ef6acdd3b601c18f54a3..e0c59bab3869df2166db9a80123aaf63353ebf57 100755 (executable)
@@ -15,8 +15,8 @@
 (* ********************************************************************** *)
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
-(*   Sviluppato da: Cosimo Oliboni, oliboni@cs.unibo.it                   *)
-(*     Cosimo Oliboni, oliboni@cs.unibo.it                                *)
+(*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
+(*   Ultima modifica: 05/08/2009                                          *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
@@ -482,14 +482,12 @@ ndefinition w16_to_recw16_aux4 ≝
  | xF ⇒ w16_to_recw16_aux4_15 … recw
  ].
 
-nlemma w16_to_recw16_aux5 : ∀b.rec_byte8 (〈b8h b,b8l b〉) → rec_byte8 b.
- #b; nelim b; #e1; #e2; nnormalize; #input; napply input. nqed.
-
-ndefinition w16_to_recw16 ≝
-λn.w16_to_recw16_aux4 (w16h n) (b8h (w16l n)) (b8l (w16l n))
-    (w16_to_recw16_aux3 (w16h n) (b8h (w16l n)) (w16_to_recw16_aux2 (w16h n) ?)).
- nelim n; #b1; #b2;
- nchange with (rec_byte8 b1);
- napply (w16_to_recw16_aux5 b1);
- napply (b8_to_recb8 b1).
+ndefinition w16_to_recw16 : Πw.rec_word16 w.
+ #w;
+ nletin K ≝ (w16_to_recw16_aux4 (w16h w) (b8h (w16l w)) (b8l (w16l w))
+              (w16_to_recw16_aux3 (w16h w) (b8h (w16l w)) (w16_to_recw16_aux2 (w16h w) (b8_to_recb8 (w16h w)))));
+ nelim w in K:(%) ⊢ %;
+ #b1; #b2;
+ nelim b2;
+ #e1; #e2; nnormalize; #H; napply H.
 nqed.