]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/num/word32.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / num / word32.ma
old mode 100755 (executable)
new mode 100644 (file)
index 09cd0ad..e9b99ef
@@ -144,9 +144,40 @@ ndefinition compl_w32 ≝
  [ true ⇒ succ_w32 (not_w32 w)
  | false ⇒ not_w32 (pred_w32 w) ].
 
+(* operatore abs *)
+ndefinition abs_w32 ≝
+λw:word32.match getMSB_w32 w with
+ [ true ⇒ compl_w32 w | false ⇒ w ].
+
 (* operatore x in [inf,sup] o in sup],[inf *)
 ndefinition inrange_w32 ≝
 λx,inf,sup:word32.
  match le_w32 inf sup with
   [ true ⇒ and_bool | false ⇒ or_bool ]
  (le_w32 inf x) (le_w32 x sup).
+
+(* operatore moltiplicazione senza segno *)
+(* 〈a1,a2〉 * 〈b1,b2〉 = (a1*b1) x0 x0 + x0 (a1*b2) x0 + x0 (a2*b1) x0 + x0 x0 (a2*b2) *)
+ndefinition mulu_w16_aux ≝
+λw.nat_it ? rol_w32 w nat8.
+
+ndefinition mulu_w16 ≝
+λw1,w2:word16.
+ plus_w32_d_d 〈(mulu_b8 (cnH ? w1) (cnH ? w2)).〈〈x0,x0〉:〈x0,x0〉〉〉
+ (plus_w32_d_d (mulu_w16_aux (extu_w32 (mulu_b8 (cnH ? w1) (cnL ? w2))))
+  (plus_w32_d_d (mulu_w16_aux (extu_w32 (mulu_b8 (cnL ? w1) (cnH ? w2))))
+                (extu_w32 (mulu_b8 (cnL ? w1) (cnL ? w2))))).
+
+(* operatore moltiplicazione con segno *)
+(* x * y = sgn(x) * sgn(y) * |x| * |y| *)
+ndefinition muls_w16 ≝
+λw1,w2:word16.
+(* ++/-- → +, +-/-+ → - *)
+ match (getMSB_w16 w1) ⊙ (getMSB_w16 w2) with
+  (* +- -+ → - *)
+  [ true ⇒ compl_w32
+  (* ++/-- → + *)
+  | false ⇒ λx.x ] (mulu_w16 (abs_w16 w1) (abs_w16 w2)).
+
+nlemma pippo : ∃b.b=(muls_w16 〈〈xC,x3〉:〈x0,x4〉〉 〈〈x7,xE〉:〈xF,x9〉〉). nnormalize;
+