Skip to content

Commit 2d55a49

Browse files
authored
Move signed to Z.signed in ZUtil.Definitions (#2039)
1 parent 0579c78 commit 2d55a49

File tree

5 files changed

+52
-45
lines changed

5 files changed

+52
-45
lines changed

src/Assembly/Symbolic.v

+5-30
Original file line numberDiff line numberDiff line change
@@ -24,6 +24,7 @@ Require Import Crypto.Util.ZUtil.Land.
2424
Require Import Crypto.Util.ZUtil.Land.Fold.
2525
Require Import Crypto.Util.ZUtil.Ones.
2626
Require Import Crypto.Util.ZUtil.Definitions.
27+
Require Import Crypto.Util.ZUtil.Signed.
2728
Require Import Crypto.Util.Equality.
2829
Require Import Crypto.Util.Bool.Reflect.
2930
Require Import Crypto.Util.Option.
@@ -596,7 +597,6 @@ Import ListNotations.
596597

597598
Section WithContext.
598599
Context (ctx : symbol -> option Z).
599-
Definition signed s n : Z := (Z.land (Z.shiftl 1 (Z.of_N s-1) + n) (Z.ones (Z.of_N s)) - Z.shiftl 1 (Z.of_N s-1))%Z.
600600
Definition op_to_Z_binop (o : op) : option _
601601
:= match o with
602602
| add _ => Some Z.add
@@ -627,12 +627,12 @@ Section WithContext.
627627
| subborrow s, cons a args' =>
628628
Some ((- Z.shiftr (a - List.fold_right Z.add 0 args') (Z.of_N s)) mod 2)
629629
| addoverflow s, args => Some (Z.b2z (negb (Z.eqb
630-
(signed s (keep s (List.fold_right Z.add 0 args)))
631-
(List.fold_right Z.add 0%Z (List.map (signed s) args)))))
630+
(Z.signed s (keep s (List.fold_right Z.add 0 args)))
631+
(List.fold_right Z.add 0%Z (List.map (Z.signed s) args)))))
632632
| neg s, [a] => Some (keep s (- a))
633633
| shl s, [a; b] => Some (keep s (Z.shiftl a b))
634634
| shr s, [a; b] => Some (keep s (Z.shiftr a b))
635-
| sar s, [a; b] => Some (keep s (Z.shiftr (signed s a) b))
635+
| sar s, [a; b] => Some (keep s (Z.shiftr (Z.signed s a) b))
636636
| rcr s, [v1; cf; cnt] => Some (
637637
let v1c := Z.lor v1 (Z.shiftl cf (Z.of_N s)) in
638638
let l := Z.lor v1c (Z.shiftl v1 (1+Z.of_N s)) in
@@ -2604,18 +2604,6 @@ Proof using Type.
26042604
rewrite Z.ones_equiv in * |- ; rewrite Z.shiftr_div_pow2, Z.div_small; cbn; lia.
26052605
Qed.
26062606

2607-
Lemma signed_small s v (Hv : (0 <= v <= Z.ones (Z.of_N s-1))%Z) : signed s v = v.
2608-
Proof using Type.
2609-
destruct (N.eq_dec s 0); subst; cbv [signed].
2610-
{ rewrite Z.land_0_r. cbn in *; Lia.lia. }
2611-
rewrite !Z.land_ones, !Z.shiftl_mul_pow2, ?Z.add_0_r, ?Z.mul_1_l by Lia.lia.
2612-
rewrite Z.ones_equiv in Hv.
2613-
rewrite Z.mod_small; try ring.
2614-
enough (2 ^ Z.of_N s = 2 ^ (Z.of_N s - 1) + 2 ^ (Z.of_N s - 1))%Z; try Lia.lia.
2615-
replace (Z.of_N s) with (1+(Z.of_N s-1))%Z at 1 by Lia.lia.
2616-
rewrite Z.pow_add_r; try Lia.lia.
2617-
Qed.
2618-
26192607
Definition addoverflow_small (d : dag) :=
26202608
fun e => match e with
26212609
ExprApp (addoverflow s, ([_]|[_;_]|[_;_;_]) as args) =>
@@ -2632,7 +2620,7 @@ Proof using Type.
26322620
BreakMatch.break_match_hyps; Option.inversion_option; t;
26332621
epose proof Z.ones_equiv (Z.of_N s -1);
26342622
destruct_head'_and; reflect_hyps; destruct_head'_and.
2635-
all : rewrite Z.land_ones, !Z.mod_small, !signed_small, !Z.eqb_refl; trivial.
2623+
all : rewrite Z.land_ones, !Z.mod_small, !Z.signed_small, !Z.eqb_refl; trivial.
26362624
all : try split; try Lia.lia.
26372625
all : replace (Z.of_N s) with (1+(Z.of_N s-1))%Z at 1 by Lia.lia;
26382626
rewrite Z.pow_add_r; try Lia.lia.
@@ -2694,19 +2682,6 @@ Proof.
26942682
apply fold_right_filter_identity_gen with (G:=id); cbv [id]; intuition (subst; eauto).
26952683
Qed.
26962684

2697-
Lemma signed_0 s : signed s 0 = 0%Z.
2698-
Proof using Type.
2699-
destruct (N.eq_dec s 0); subst; trivial.
2700-
cbv [signed].
2701-
rewrite !Z.land_ones, !Z.shiftl_mul_pow2, ?Z.add_0_r, ?Z.mul_1_l by Lia.lia.
2702-
rewrite Z.mod_small; try ring.
2703-
split; try (eapply Z.pow_lt_mono_r; Lia.lia).
2704-
eapply Z.pow_nonneg; Lia.lia.
2705-
Qed.
2706-
#[global]
2707-
Hint Rewrite signed_0 : zsimplify_const zsimplify zsimplify_fast.
2708-
Global Hint Resolve signed_0 : zarith.
2709-
27102685
Lemma interp_op_drop_identity o id : identity o = Some id ->
27112686
forall G xs, interp_op G o xs = interp_op G o (List.filter (fun v => negb (Z.eqb v id)) xs).
27122687
Proof using Type.

src/Assembly/WithBedrock/Semantics.v

+7-9
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ Require Import Crypto.Util.ListUtil.
1212
Require Import Crypto.Util.Tactics.DestructHead.
1313
Require Import Crypto.Util.Tactics.BreakMatch.
1414
Require Import Crypto.Util.Notations.
15+
Require Import Crypto.Util.ZUtil.Definitions.
1516
Require Import Crypto.Assembly.Syntax.
1617
Require Crypto.Util.Tuple.
1718
Import ListNotations.
@@ -148,9 +149,6 @@ Definition HavocFlagsFromResult s st v : machine_state :=
148149
Definition PreserveFlag (st : machine_state) (f : FLAG) st' :=
149150
update_flag_with st (fun fs => set_flag_internal fs f (get_flag st' f)).
150151

151-
Definition signed (s : N) (z : Z) : Z :=
152-
Z.land (Z.shiftl 1 (Z.of_N s-1) + z) (Z.ones (Z.of_N s)) - Z.shiftl 1 (Z.of_N s-1).
153-
154152
Definition rcrcnt s cnt : Z :=
155153
if N.eqb s 8 then Z.land cnt 0x1f mod 9 else
156154
if N.eqb s 16 then Z.land cnt 0x1f mod 17 else
@@ -207,7 +205,7 @@ Definition DenoteNormalInstruction (st : machine_state) (instr : NormalInstructi
207205
st <- SetOperand sa s st dst v;
208206
let st := HavocFlagsFromResult s st v in
209207
let st := SetFlag st CF (Z.odd (Z.shiftr (v1 + v2 + c) (Z.of_N s))) in
210-
Some (SetFlag st OF (negb (signed s v =? signed s v1 + signed s v2 + signed s c)%Z))
208+
Some (SetFlag st OF (negb (Z.signed s v =? Z.signed s v1 + Z.signed s v2 + Z.signed s c)%Z))
211209
| (adcx | adox) as opc, [dst; src] =>
212210
let flag := match opc with adcx => CF | _ => OF end in
213211
v1 <- DenoteOperand sa s st dst;
@@ -226,19 +224,19 @@ Definition DenoteNormalInstruction (st : machine_state) (instr : NormalInstructi
226224
st <- SetOperand sa s st dst v;
227225
let st := HavocFlagsFromResult s st v in
228226
let st := SetFlag st CF (Z.odd (Z.shiftr (v1 - (v2 + c)) (Z.of_N s))) in
229-
Some (SetFlag st OF (negb (signed s v =? signed s v1 - (signed s v2 + c))%Z))
227+
Some (SetFlag st OF (negb (Z.signed s v =? Z.signed s v1 - (Z.signed s v2 + c))%Z))
230228
| dec, [dst] =>
231229
v1 <- DenoteOperand sa s st dst;
232230
let v := Z.land (v1 - 1) (Z.ones (Z.of_N s)) in
233231
st <- SetOperand sa s st dst v;
234232
let st := PreserveFlag (HavocFlagsFromResult s st v) CF st in
235-
Some (SetFlag st OF (negb (signed s v =? signed s v1 - 1)%Z))
233+
Some (SetFlag st OF (negb (Z.signed s v =? Z.signed s v1 - 1)%Z))
236234
| inc, [dst] =>
237235
v1 <- DenoteOperand sa s st dst;
238236
let v := Z.land (v1 + 1) (Z.ones (Z.of_N s)) in
239237
st <- SetOperand sa s st dst v;
240238
let st := PreserveFlag (HavocFlagsFromResult s st v) CF st in
241-
Some (SetFlag st OF (negb (signed s v =? signed s v1 + 1)%Z))
239+
Some (SetFlag st OF (negb (Z.signed s v =? Z.signed s v1 + 1)%Z))
242240
| mulx, [hi; lo; src2] => (* Flags Affected: None *)
243241
let src1 : ARG := rdx in (* Note: assumes s=64 *)
244242
v1 <- DenoteOperand sa s st src1;
@@ -269,7 +267,7 @@ Definition DenoteNormalInstruction (st : machine_state) (instr : NormalInstructi
269267
v1 <- DenoteOperand sa s st dst;
270268
cnt' <- DenoteOperand sa s st cnt;
271269
let cnt := Z.land cnt' (Z.of_N s-1) in
272-
let v := Z.land (Z.shiftr (signed s v1) cnt) (Z.ones (Z.of_N s)) in
270+
let v := Z.land (Z.shiftr (Z.signed s v1) cnt) (Z.ones (Z.of_N s)) in
273271
st <- SetOperand sa s st dst v;
274272
Some (if cnt =? 0 then st else
275273
let st := HavocFlagsFromResult s st v in
@@ -324,7 +322,7 @@ Definition DenoteNormalInstruction (st : machine_state) (instr : NormalInstructi
324322
if cnt =? 0 then Some st else
325323
if Z.of_N s <? cnt then Some (HavocFlags st) else
326324
let st := HavocFlagsFromResult s st l in
327-
let signchange := xorb (signed s lv <? 0)%Z (signed s v <? 0)%Z in
325+
let signchange := xorb (Z.signed s lv <? 0)%Z (Z.signed s v <? 0)%Z in
328326
(* Note: IA-32 SDM does not make it clear what sign change is in question *)
329327
let st := if cnt =? 1 then SetFlag st OF signchange else st in
330328
let st := SetFlag st CF (Z.testbit l (cnt-1)) in

src/Assembly/WithBedrock/SymbolicProofs.v

+4-6
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ Require Import Crypto.Util.ErrorT.
88
Import Coq.Lists.List. (* [map] is [List.map] not [ErrorT.map] *)
99
Require Import Crypto.Util.ListUtil.IndexOf.
1010
Require Import Crypto.Util.Tactics.WarnIfGoalsRemain.
11+
Require Import Crypto.Util.ZUtil.Definitions.
1112
Require Crypto.Util.Option.
1213
Require Import Crypto.Assembly.Syntax.
1314
Require Import Crypto.Assembly.Symbolic.
@@ -1198,7 +1199,6 @@ Proof using Type.
11981199
end; eauto; try Lia.lia; try congruence.
11991200
eexists. split. eauto.
12001201
f_equal. f_equal.
1201-
change Symbolic.signed with Semantics.signed.
12021202
rewrite ?Z.add_0_r.
12031203
f_equal.
12041204
1:congruence.
@@ -1225,7 +1225,6 @@ Proof using Type.
12251225
end; eauto; try Lia.lia; try congruence.
12261226
eexists. split. eauto.
12271227
f_equal. f_equal.
1228-
change Symbolic.signed with Semantics.signed.
12291228
rewrite ?Z.add_0_r.
12301229
f_equal.
12311230
1:congruence.
@@ -1244,9 +1243,9 @@ Proof using Type.
12441243
| _ => destruct_one_match
12451244
| _ => progress intuition idtac
12461245
end; rewrite ?Z.add_0_r, ?Z.odd_opp; eauto; try Lia.lia; try congruence.
1247-
replace (Semantics.signed n 0) with 0; cycle 1.
1246+
replace (Z.signed n 0) with 0; cycle 1.
12481247
{ pose_operation_size_cases. clear -H0; intuition (subst; cbv; trivial). }
1249-
rewrite Z.add_0_r; cbv [Semantics.signed Symbolic.signed]; congruence. }
1248+
rewrite Z.add_0_r; cbv [Z.signed]; congruence. }
12501249

12511250
Unshelve. all : match goal with H : context[Syntax.adc] |- _ => idtac | _ => shelve end.
12521251
{ destruct s';
@@ -1259,8 +1258,7 @@ Proof using Type.
12591258
| _ => progress (cbv [R_flags Tuple.fieldwise Tuple.fieldwise'] in *; cbn -[Syntax.operation_size] in * ; subst)
12601259
| _ => destruct_one_match
12611260
| _ => progress intuition idtac
1262-
end; rewrite ?Z.add_assoc, ?Z.add_0_r, ?Z.odd_opp; eauto; try Lia.lia; try congruence.
1263-
change Symbolic.signed with Semantics.signed. congruence. }
1261+
end; rewrite ?Z.add_assoc, ?Z.add_0_r, ?Z.odd_opp; eauto; try Lia.lia; try congruence. }
12641262

12651263
Unshelve. all : match goal with H : context[Syntax.adcx] |- _ => idtac | _ => shelve end.
12661264
{ cbn [fold_right] in *; rewrite ?Z.bit0_odd, ?Z.add_0_r, ?Z.add_assoc in *; assumption. }

src/Util/ZUtil/Definitions.v

+3
Original file line numberDiff line numberDiff line change
@@ -166,6 +166,9 @@ Module Z.
166166

167167
Definition twos_complement_mul ma mb a b :=
168168
(sign_extend ma (ma + mb) a) * (sign_extend mb (ma + mb) b).
169+
170+
Definition signed (s : N) (n : Z) : Z :=
171+
Z.land (Z.shiftl 1 (Z.of_N s-1) + n) (Z.ones (Z.of_N s)) - Z.shiftl 1 (Z.of_N s-1).
169172
End Z.
170173

171174
Module Export Notations.

src/Util/ZUtil/Signed.v

+33
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,33 @@
1+
From Coq Require Import ZArith.
2+
Require Import Crypto.Util.Decidable.
3+
Require Import Crypto.Util.ZUtil.Notations.
4+
Require Import Crypto.Util.ZUtil.Definitions.
5+
Require Import Crypto.Util.LetIn.
6+
Local Open Scope Z_scope.
7+
8+
Module Z.
9+
Lemma signed_small s v (Hv : (0 <= v <= Z.ones (Z.of_N s-1))) : Z.signed s v = v.
10+
Proof.
11+
destruct (N.eq_dec s 0); subst; cbv [Z.signed].
12+
{ rewrite Z.land_0_r. cbn in *; Lia.lia. }
13+
rewrite !Z.land_ones, !Z.shiftl_mul_pow2, ?Z.add_0_r, ?Z.mul_1_l by Lia.lia.
14+
rewrite Z.ones_equiv in Hv.
15+
rewrite Z.mod_small; try ring.
16+
enough (2 ^ Z.of_N s = 2 ^ (Z.of_N s - 1) + 2 ^ (Z.of_N s - 1))%Z; try Lia.lia.
17+
replace (Z.of_N s) with (1+(Z.of_N s-1))%Z at 1 by Lia.lia.
18+
rewrite Z.pow_add_r; try Lia.lia.
19+
Qed.
20+
21+
Lemma signed_0 s : Z.signed s 0 = 0.
22+
Proof.
23+
destruct (N.eq_dec s 0); subst; trivial.
24+
cbv [Z.signed].
25+
rewrite !Z.land_ones, !Z.shiftl_mul_pow2, ?Z.add_0_r, ?Z.mul_1_l by Lia.lia.
26+
rewrite Z.mod_small; try ring.
27+
split; try (eapply Z.pow_lt_mono_r; Lia.lia).
28+
eapply Z.pow_nonneg; Lia.lia.
29+
Qed.
30+
#[global]
31+
Hint Rewrite signed_0 : zsimplify_const zsimplify zsimplify_fast.
32+
Global Hint Resolve signed_0 : zarith.
33+
End Z.

0 commit comments

Comments
 (0)