λx,y.
mk_status 〈x0, x0〉 0 0 false false (mult_memory x y) 0.
+notation " 'M' \sub (x y)" non associative with precedence 80 for
+ @{ 'memory $x $y }.
+
+interpretation "mult_memory" 'memory x y =
+ (cic:/matita/assembly/test/mult_memory.con x y).
+
+notation " 'M' \sub (x y) \nbsp a" non associative with precedence 80 for
+ @{ 'memory4 $x $y $a }.
+
+interpretation "mult_memory4" 'memory4 x y a =
+ (cic:/matita/assembly/test/mult_memory.con x y a).
+
+notation " \Sigma \sub (x y)" non associative with precedence 80 for
+ @{ 'status $x $y }.
+
+interpretation "mult_status" 'status x y =
+ (cic:/matita/assembly/test/mult_status.con x y).
+
lemma test_O_O:
let i ≝ 14 in
let s ≝ execute (mult_status 〈x0, x0〉 〈x0, x0〉) i in
reflexivity.
qed.
+(*
lemma test_0_2:
let x ≝ 〈x0, x0〉 in
let y ≝ 〈x0, x2〉 in
split;
reflexivity.
qed.
+*)
lemma test_x_1:
∀x.
].
qed.
+(*
lemma test_x_2:
∀x.
let y ≝ 〈x0, x2〉 in
reflexivity
].
qed.
+*)
lemma loop_invariant':
∀x,y:byte.∀j:nat. j ≤ y →
| intro;
rewrite < minus_n_O;
normalize in ⊢ (? ? (? (? ? %) ?) ?);
- whd in ⊢ (? ? % ?);
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);
[ apply (y - S n)
| split;
[ rewrite < (minus_S_S y n);
- autobatch
+ apply (minus_Sn_m (nat_of_byte y) (S n) H1)
| letin K ≝ (lt_nat_of_byte_256 y); clearbody K;
letin K' ≝ (lt_minus_m y (S n) ? ?); clearbody K';
autobatch
]
qed.
+
theorem test_x_y:
∀x,y:byte.
let i ≝ 14 + 23 * y in
execute (mult_status x y) i =
- mk_status (byte_of_nat (x*y)) 20 0
- (eqbyte 〈x0, x0〉 (byte_of_nat (x*y)))
+ mk_status (#(x*y)) 20 0
+ (eqbyte 〈x0, x0〉 (#(x*y)))
(plusbytec (byte_of_nat (x*pred y)) x)
(update
(update (mult_memory x y) 31 〈x0, x0〉)
apply status_eq;
[1,2,3,4,5,7: normalize; reflexivity
| intro;
+ letin xxx \def ((mult_memory x y) { a ↦ x }).
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 〈x0, x0〉) 32
clk : nat
}.
+notation "{hvbox('Acc' ≝ acc ; break 'Pc' ≝ pc ; break 'Spc' ≝ spc ; break 'Fz' ≝ fz ; break 'Fc' ≝ fc ; break 'M' ≝ m ; break 'Clk' ≝ clk)}"
+non associative with precedence 80 for
+ @{ 'mkstatus $acc $pc $spc $fz $fc $m $clk }.
+
+interpretation "mk_status" 'mkstatus acc pc spc fz fc m clk =
+ (cic:/matita/assembly/vm/status.ind#xpointer(1/1/1) acc pc spc fz fc m clk).
+
definition update ≝
λf: addr → byte.λa.λv.λx.
match eqb x a with
[ true ⇒ v
| false ⇒ f x ].
+notation "hvbox(m break {a ↦ break v})" non associative with precedence 80 for
+ @{ 'update $m $a $v }.
+
+notation "hvbox(m break {a ↦ break v} \nbsp x)" non associative with precedence 80 for
+ @{ 'update4 $m $a $v $x }.
+
+interpretation "update" 'update m a v =
+ (cic:/matita/assembly/vm/update.con m a v).
+
+interpretation "update4" 'update4 m a v x =
+ (cic:/matita/assembly/vm/update.con m a v x).
+
lemma update_update_a_a:
∀s,a,v1,v2,b.
update (update s a v1) a v2 b = update s a v2 b.