X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fassembly%2Ftest.ma;h=644cebaa9d57d3edb1182e63a4fc164faba9a4e2;hb=32d8d8d419e0b910435da275361bb55d49bc43a9;hp=1bbc1e7103171613343915d3a003080b6eda0a2f;hpb=006db2fb5a2ccd6e5043b6809a620ac96603648c;p=helm.git diff --git a/helm/software/matita/library/assembly/test.ma b/helm/software/matita/library/assembly/test.ma index 1bbc1e710..644cebaa9 100644 --- a/helm/software/matita/library/assembly/test.ma +++ b/helm/software/matita/library/assembly/test.ma @@ -65,12 +65,10 @@ lemma test_O_O: let i ≝ 14 in let s ≝ execute (mult_status 〈x0, x0〉 〈x0, x0〉) i in pc s = 20 ∧ mem s 32 = byte_of_nat 0. - normalize; split; reflexivity. qed. -(* lemma test_0_2: let x ≝ 〈x0, x0〉 in let y ≝ 〈x0, x2〉 in @@ -81,7 +79,6 @@ lemma test_0_2: split; reflexivity. qed. -*) lemma test_x_1: ∀x. @@ -98,7 +95,6 @@ lemma test_x_1: ]. qed. -(* lemma test_x_2: ∀x. let y ≝ 〈x0, x2〉 in @@ -114,7 +110,6 @@ lemma test_x_2: reflexivity ]. qed. -*) lemma loop_invariant': ∀x,y:byte.∀j:nat. j ≤ y → @@ -137,45 +132,31 @@ lemma loop_invariant': rewrite < minus_n_O; normalize 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); + simplify in ⊢ (? ? ? %); change in ⊢ (? ? ? (? (? (? ? ? %) ? ?) ? ? ?)) with (mult_memory x y 30); rewrite > byte_of_nat_nat_of_byte; change in ⊢ (? ? ? (? (? ? ? %) ? ? ?)) with (mult_memory x y 31); apply inj_update; intro; rewrite > (eq_update_s_a_sa (update (mult_memory x y) 30 (mult_memory x y 30)) - 31 a) in ⊢ (? ? ? %); + 31 a); rewrite > eq_update_s_a_sa; reflexivity ] | cut (5 + 23 * S n = 5 + 23 * n + 23); - [ letin K ≝ (breakpoint (mult_status x y) (5 + 23 * n) 23); clearbody K; - letin H' ≝ (H ?); clearbody H'; clear H; - [ autobatch - | letin xxx ≝ (eq_f ? ? (λz. execute (mult_status x y) z) ? ? Hcut); clearbody xxx; - clear Hcut; - rewrite > xxx; - clear xxx; - apply (transitive_eq ? ? ? ? K); - clear K; - rewrite > H'; - clear H'; - cut (∃z.y-n=S z ∧ z < 255); + [ rewrite > Hcut; clear Hcut; + rewrite > breakpoint; + rewrite > H; clear H; + [2: apply le_S_S_to_le; + apply le_S; + apply H1 + | cut (∃z.y-n=S z ∧ z < 255); [ elim Hcut; clear Hcut; elim H; clear H; rewrite > H2; (* instruction LDAd *) - letin K ≝ - (breakpoint - (mk_status (byte_of_nat (x*n)) 4 O - (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) - 3 20); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + change in ⊢ (? ? (? ? %) ?) with (3+20); + rewrite > breakpoint in ⊢ (? ? % ?); whd in ⊢ (? ? (? % ?) ?); normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?); change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) @@ -183,46 +164,22 @@ lemma loop_invariant': change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with (byte_of_nat (S a)); (* instruction BEQ *) - letin K ≝ - (breakpoint - (mk_status (byte_of_nat (S a)) 6 O - (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) - 3 17); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + change in ⊢ (? ? (? ? %) ?) with (3+17); + rewrite > breakpoint in ⊢ (? ? % ?); whd in ⊢ (? ? (? % ?) ?); letin K ≝ (eq_eqbyte_x0_x0_byte_of_nat_S_false ? H3); clearbody K; rewrite > K; clear K; simplify in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?); (* instruction LDAd *) - letin K ≝ - (breakpoint - (mk_status (byte_of_nat (S a)) 8 O - (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) - 3 14); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + change in ⊢ (? ? (? ? %) ?) with (3+14); + rewrite > breakpoint in ⊢ (? ? % ?); whd in ⊢ (? ? (? % ?) ?); change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with (byte_of_nat (x*n)); normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?); 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 〈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) - 5 9); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + change in ⊢ (? ? (? ? %) ?) with (5+9); + rewrite > breakpoint in ⊢ (? ? % ?); whd in ⊢ (? ? (? % ?) ?); change in ⊢ (? ? (? (? ? ? ? (? ? %) ? ? ?) ?) ?) with (bpred (byte_of_nat (S a))); rewrite > (eq_bpred_S_a_a ? H3); @@ -234,18 +191,8 @@ lemma loop_invariant': reflexivity | ]; rewrite < Hcut; clear Hcut; clear H3; clear H2; clear a; (* instruction ADDd *) - letin K ≝ - (breakpoint - (mk_status (byte_of_nat (x*n)) 12 - 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)))) - 32 (byte_of_nat (x*n))) 31 - (byte_of_nat (y-S n))) O) - 3 6); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + change in ⊢ (? ? (? ? %) ?) with (3+6); + rewrite > breakpoint in ⊢ (? ? % ?); whd in ⊢ (? ? (? % ?) ?); change in ⊢ (? ? (? (? % ? ? ? ? ? ?) ?) ?) with (plusbytenc (byte_of_nat (x*n)) x); @@ -256,18 +203,7 @@ lemma loop_invariant': with (plusbytec (byte_of_nat (x*n)) x); rewrite > plusbytenc_S; (* instruction STAd *) - letin K ≝ - (breakpoint - (mk_status (byte_of_nat (x*S n)) 14 O - (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)))) - 32 (byte_of_nat (x*n))) 31 - (byte_of_nat (y-S n))) O) - 3 3); clearbody K; - normalize in K:(? ? (? ? %) ?); - apply transitive_eq; [2: apply K | skip | ]; clear K; + rewrite > (breakpoint ? 3 3); whd in ⊢ (? ? (? % ?) ?); normalize in ⊢ (? ? (? (? ? % ? ? ? ? ?) ?) ?); (* instruction BRA *) @@ -276,32 +212,20 @@ lemma loop_invariant': rewrite < pred_Sn; apply status_eq; [1,2,3,4,7: normalize; reflexivity - | change with (plusbytec (byte_of_nat (x*n)) x = - plusbytec (byte_of_nat (x*n)) x); + | change with (plusbytec #(x*n) x = plusbytec #(x*n) x); reflexivity |6: intro; simplify in ⊢ (? ? ? %); + normalize in ⊢ (? ? (? (? ? ? ? ? ? (? ? (? %) ?) ?) ?) ?); change in ⊢ (? ? % ?) with - (update - (update - (update - (update (update (mult_memory x y) 30 x) 31 - (byte_of_nat (S (nat_of_byte y-S n)))) 32 (byte_of_nat (nat_of_byte x*n))) 31 - (byte_of_nat (nat_of_byte y-S n))) - (nat_of_byte - (update - (update - (update (update (mult_memory x y) 30 x) 31 - (byte_of_nat (S (nat_of_byte y-S n)))) 32 (byte_of_nat (nat_of_byte x*n))) 31 - (byte_of_nat (nat_of_byte y-S n)) 15)) - (byte_of_nat (nat_of_byte x*S n)) a); - normalize in ⊢ (? ? (? ? % ? ?) ?); + ((mult_memory x y){30↦x}{31↦#(S (y-S n))}{32↦#(x*n)}{31↦#(y-S n)} + {〈x2,x0〉↦ #(x*S n)} a); apply inj_update; intro; apply inj_update; intro; - rewrite > not_eq_a_b_to_eq_update_a_b in ⊢ (? ? % ?); [2: apply H | ]; - rewrite > not_eq_a_b_to_eq_update_a_b in ⊢ (? ? % ?); + rewrite > not_eq_a_b_to_eq_update_a_b; [2: apply H | ]; + rewrite > not_eq_a_b_to_eq_update_a_b; [ reflexivity | assumption ] @@ -313,7 +237,11 @@ lemma loop_invariant': 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 + [ apply (lt_to_le_to_lt O (S n) (nat_of_byte y) ? ?); + autobatch + | autobatch + | autobatch + ] ] ] ] @@ -349,16 +277,13 @@ theorem test_x_y: 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 + simplify in ⊢ (? ? ? %); + change in ⊢ (? ? % ?) 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); - apply inj_update; intro; - apply inj_update; intro; - change in ⊢ (? ? (? ? ? % ?) ?) with (mult_memory x y 30); - apply eq_update_s_a_sa + repeat (apply inj_update; intro); + apply (eq_update_s_a_sa ? 30) ] ] ]. -qed. +qed. \ No newline at end of file