set "baseuri" "cic:/matita/assembly/test/".
-include "vm.ma".
-
-notation "hvbox(# break a)"
- non associative with precedence 80
-for @{ 'byte_of_opcode $a }.
-interpretation "byte_of_opcode" 'byte_of_opcode a =
- (cic:/matita/assembly/vm/byte_of_opcode.con a).
+include "assembly/vm.ma".
definition mult_source : list byte ≝
- [#LDAi; mk_byte x0 x0; (* A := 0 *)
- #STAd; mk_byte x2 x0; (* Z := A *)
- #LDAd; mk_byte x1 xF; (* (l1) A := Y *)
- #BEQ; mk_byte x0 xA; (* if A == 0 then goto l2 *)
- #LDAd; mk_byte x2 x0; (* A := Z *)
- #DECd; mk_byte x1 xF; (* Y := Y - 1 *)
- #ADDd; mk_byte x1 xE; (* A += X *)
- #STAd; mk_byte x2 x0; (* Z := A *)
- #BRA; mk_byte xF x2; (* goto l1 *)
- #LDAd; mk_byte x2 x0].(* (l2) *)
+ [#LDAi; 〈x0, x0〉; (* A := 0 *)
+ #STAd; 〈x2, x0〉; (* Z := A *)
+ #LDAd; 〈x1, xF〉; (* (l1) A := Y *)
+ #BEQ; 〈x0, xA〉; (* if A == 0 then goto l2 *)
+ #LDAd; 〈x2, x0〉; (* A := Z *)
+ #DECd; 〈x1, xF〉; (* Y := Y - 1 *)
+ #ADDd; 〈x1, xE〉; (* A += X *)
+ #STAd; 〈x2, x0〉; (* Z := A *)
+ #BRA; 〈xF, x2〉; (* goto l1 *)
+ #LDAd; 〈x2, x0〉].(* (l2) *)
definition mult_memory ≝
λx,y.λa:addr.
match leb a 29 with
- [ true ⇒ nth ? mult_source (mk_byte x0 x0) a
+ [ true ⇒ nth ? mult_source 〈x0, x0〉 a
| false ⇒
match eqb a 30 with
[ true ⇒ x
definition mult_status ≝
λx,y.
- mk_status (mk_byte x0 x0) 0 0 false false (mult_memory x y) 0.
+ mk_status 〈x0, x0〉 0 0 false false (mult_memory x y) 0.
lemma test_O_O:
let i ≝ 14 in
- let s ≝ execute (mult_status (mk_byte x0 x0) (mk_byte x0 x0)) i in
+ let s ≝ execute (mult_status 〈x0, x0〉 〈x0, x0〉) i in
pc s = 20 ∧ mem s 32 = byte_of_nat 0.
normalize;
split;
qed.
lemma test_0_2:
- let x ≝ mk_byte x0 x0 in
- let y ≝ mk_byte x0 x2 in
+ let x ≝ 〈x0, x0〉 in
+ let y ≝ 〈x0, x2〉 in
let i ≝ 14 + 23 * nat_of_byte y in
let s ≝ execute (mult_status x y) i in
pc s = 20 ∧ mem s 32 = plusbytenc x x.
lemma test_x_1:
∀x.
- let y ≝ mk_byte x0 x1 in
+ let y ≝ 〈x0, x1〉 in
let i ≝ 14 + 23 * nat_of_byte y in
let s ≝ execute (mult_status x y) i in
pc s = 20 ∧ mem s 32 = x.
intros;
split;
[ reflexivity
- | change in ⊢ (? ? % ?) with (plusbytenc (mk_byte x0 x0) x);
+ | change in ⊢ (? ? % ?) with (plusbytenc 〈x0, x0〉 x);
rewrite > plusbytenc_O_x;
reflexivity
].
lemma test_x_2:
∀x.
- let y ≝ mk_byte x0 x2 in
+ let y ≝ 〈x0, x2〉 in
let i ≝ 14 + 23 * nat_of_byte y in
let s ≝ execute (mult_status x y) i in
pc s = 20 ∧ mem s 32 = plusbytenc x x.
split;
[ reflexivity
| change in ⊢ (? ? % ?) with
- (plusbytenc (plusbytenc (mk_byte x0 x0) x) x);
+ (plusbytenc (plusbytenc 〈x0, x0〉 x) x);
rewrite > plusbytenc_O_x;
reflexivity
].
∀x,y:byte.∀j:nat. j ≤ y →
execute (mult_status x y) (5 + 23*j)
=
- mk_status (byte_of_nat (x * j)) 4 0 (eqbyte (mk_byte x0 x0) (byte_of_nat (x*j)))
+ mk_status (byte_of_nat (x * j)) 4 0 (eqbyte 〈x0, x0〉 (byte_of_nat (x*j)))
(plusbytec (byte_of_nat (x*pred j)) x)
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (y - j))) 32
(byte_of_nat (x * j)))
rewrite < minus_n_O;
normalize in ⊢ (? ? (? (? ? %) ?) ?);
whd in ⊢ (? ? % ?);
- change in ⊢ (? ? % ?) with (update (mult_memory x y) 32 (mk_byte x0 x0) a);
+ change in ⊢ (? ? % ?) with (update (mult_memory x y) 32 〈x0, x0〉 a);
change in ⊢ (? ? ? %) with (update (update (update (mult_memory x y) 30 x) 31
(byte_of_nat y)) 32 (byte_of_nat 0) a);
change in ⊢ (? ? ? (? (? (? ? ? %) ? ?) ? ? ?)) with (mult_memory x y 30);
letin K ≝
(breakpoint
(mk_status (byte_of_nat (x*n)) 4 O
- (eqbyte (mk_byte x0 x0) (byte_of_nat (x*n)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (x*n)))
(plusbytec (byte_of_nat (x*pred n)) x)
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S a))) 32
(byte_of_nat (x*n))) O)
letin K ≝
(breakpoint
(mk_status (byte_of_nat (S a)) 6 O
- (eqbyte (mk_byte x0 x0) (byte_of_nat (S a)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (S a)))
(plusbytec (byte_of_nat (x*pred n)) x)
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S a))) 32
(byte_of_nat (x*n))) O)
letin K ≝
(breakpoint
(mk_status (byte_of_nat (S a)) 8 O
- (eqbyte (mk_byte x0 x0) (byte_of_nat (S a)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (S a)))
(plusbytec (byte_of_nat (x*pred n)) x)
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S a))) 32
(byte_of_nat (x*n))) O)
whd in ⊢ (? ? (? % ?) ?);
change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with (byte_of_nat (x*n));
normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?);
- change in ⊢ (? ? (? (? ? ? ? % ? ? ?) ?) ?) with (eqbyte (mk_byte x0 x0) (byte_of_nat (x*n)));
+ change in ⊢ (? ? (? (? ? ? ? % ? ? ?) ?) ?) with (eqbyte 〈x0, x0〉 (byte_of_nat (x*n)));
(* instruction DECd *)
letin K ≝
(breakpoint
(mk_status (byte_of_nat (x*n)) 10 O
- (eqbyte (mk_byte x0 x0) (byte_of_nat (x*n)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (x*n)))
(plusbytec (byte_of_nat (x*pred n)) x)
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S a))) 32
(byte_of_nat (x*n))) O)
letin K ≝
(breakpoint
(mk_status (byte_of_nat (x*n)) 12
- O (eqbyte (mk_byte x0 x0) (byte_of_nat (y-S n)))
+ O (eqbyte 〈x0, x0〉 (byte_of_nat (y-S n)))
(plusbytec (byte_of_nat (x*pred n)) x)
(update
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S (y-S n))))
letin K ≝
(breakpoint
(mk_status (byte_of_nat (x*S n)) 14 O
- (eqbyte (mk_byte x0 x0) (byte_of_nat (x*S n)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (x*S n)))
(plusbytec (byte_of_nat (x*n)) x)
(update
(update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat (S (y-S n))))
let i ≝ 14 + 23 * y in
execute (mult_status x y) i =
mk_status (byte_of_nat (x*y)) 20 0
- (eqbyte (mk_byte x0 x0) (byte_of_nat (x*y)))
+ (eqbyte 〈x0, x0〉 (byte_of_nat (x*y)))
(plusbytec (byte_of_nat (x*pred y)) x)
(update
- (update (mult_memory x y) 31 (mk_byte x0 x0))
+ (update (mult_memory x y) 31 〈x0, x0〉)
32 (byte_of_nat (x*y)))
0.
intros;
| intro;
change with (update (update (update (mult_memory x y) 30 x) 31 (byte_of_nat O)) 32
(byte_of_nat (nat_of_byte x*nat_of_byte y)) a =
- update (update (mult_memory x y) 31 (mk_byte x0 x0)) 32
+ update (update (mult_memory x y) 31 〈x0, x0〉) 32
(byte_of_nat (nat_of_byte x*nat_of_byte y)) a);
apply inj_update; intro;
apply inj_update; intro;