aboutsummaryrefslogtreecommitdiff
path: root/src
diff options
context:
space:
mode:
Diffstat (limited to 'src')
-rw-r--r--src/Compilers/Z/ArithmeticSimplifier.v3
-rw-r--r--src/Compilers/Z/Bounds/Interpretation.v33
-rw-r--r--src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v24
-rw-r--r--src/Compilers/Z/CommonSubexpressionElimination.v20
-rw-r--r--src/Compilers/Z/Syntax.v4
-rw-r--r--src/Compilers/Z/Syntax/Equality.v6
-rw-r--r--src/Compilers/Z/Syntax/Util.v2
-rw-r--r--src/Util/ZUtil/Definitions.v10
-rw-r--r--src/Util/ZUtil/Morphisms.v2
9 files changed, 92 insertions, 12 deletions
diff --git a/src/Compilers/Z/ArithmeticSimplifier.v b/src/Compilers/Z/ArithmeticSimplifier.v
index f0d3e19ab..7a2a73434 100644
--- a/src/Compilers/Z/ArithmeticSimplifier.v
+++ b/src/Compilers/Z/ArithmeticSimplifier.v
@@ -32,6 +32,7 @@ Section language.
| OpConst _ z => fun _ _ => const_of _ z
| Opp TZ TZ => fun _ args => neg_expr _ args
| AddWithGetCarry _ _ _ _ _ _ => fun _ _ => I
+ | SubWithGetBorrow _ _ _ _ _ _ => fun _ _ => I
| _ => fun e _ => gen_expr _ e
end (Op opc args) args)
| TT => Some tt
@@ -187,6 +188,8 @@ Section language.
| Zselect _ _ _ _ as opc
| AddWithCarry _ _ _ _ as opc
| AddWithGetCarry _ _ _ _ _ _ as opc
+ | SubWithBorrow _ _ _ _ as opc
+ | SubWithGetBorrow _ _ _ _ _ _ as opc
=> Op opc
end.
End with_var.
diff --git a/src/Compilers/Z/Bounds/Interpretation.v b/src/Compilers/Z/Bounds/Interpretation.v
index f4cbb3bbd..dc29fe654 100644
--- a/src/Compilers/Z/Bounds/Interpretation.v
+++ b/src/Compilers/Z/Bounds/Interpretation.v
@@ -105,6 +105,10 @@ Module Import Bounds.
:= t_map3' Z.add_with_carry.
Definition add_with_carry : t -> t -> t -> t
:= t_map3 Z.add_with_carry.
+ Definition sub_with_borrow' : t -> t -> t -> t
+ := t_map3' Z.sub_with_borrow.
+ Definition sub_with_borrow : t -> t -> t -> t
+ := t_map3 Z.sub_with_borrow.
Definition modulo_pow2_constant : Z -> t -> t
:= fun e x
=> let d := 2^e in
@@ -116,6 +120,11 @@ Module Import Bounds.
=> let d := 2^e in
let (l, u) := (lower x, upper x) in
truncation_bounds {| lower := l / d ; upper := u / d |}.
+ Definition opp_div_pow2_constant : Z -> t -> t
+ := fun e x
+ => let d := 2^e in
+ let (l, u) := (lower x, upper x) in
+ truncation_bounds {| lower := -(u / d) ; upper := -(l / d) |}.
Definition neg' (int_width : Z) : t -> t
:= fun v
=> let (lb, ub) := v in
@@ -143,12 +152,22 @@ Module Import Bounds.
:= truncation_bounds (cmovle' r1 r2).
End with_bitwidth.
Section with_bitwidth2.
- Context (bit_width1 bit_width2 : option Z).
- Definition add_with_get_carry (carry_boundary_bit_width : Z) : t -> t -> t -> t * t
+ Context (bit_width1 bit_width2 : option Z)
+ (carry_boundary_bit_width : Z).
+ Definition get_carry : t -> t * t
+ := fun v =>
+ (modulo_pow2_constant bit_width1 carry_boundary_bit_width v,
+ div_pow2_constant bit_width2 carry_boundary_bit_width v).
+ Definition get_borrow : t -> t * t
+ := fun v =>
+ (modulo_pow2_constant bit_width1 carry_boundary_bit_width v,
+ opp_div_pow2_constant bit_width2 carry_boundary_bit_width v).
+ Definition add_with_get_carry : t -> t -> t -> t * t
:= fun c x y
- => let xpy := add_with_carry' c x y in
- (modulo_pow2_constant bit_width1 carry_boundary_bit_width xpy,
- div_pow2_constant bit_width2 carry_boundary_bit_width xpy).
+ => get_carry (add_with_carry' c x y).
+ Definition sub_with_get_borrow : t -> t -> t -> t * t
+ := fun c x y
+ => get_borrow (sub_with_borrow' c x y).
End with_bitwidth2.
Module Export Notations.
@@ -189,6 +208,10 @@ Module Import Bounds.
| AddWithGetCarry carry_boundary_bit_width _ _ _ T1 T2
=> fun cxy => let '(c, x, y) := eta3 cxy in
add_with_get_carry (bit_width_of_base_type T1) (bit_width_of_base_type T2) carry_boundary_bit_width c x y
+ | SubWithBorrow _ _ _ T => fun cxy => let '(c, x, y) := eta3 cxy in sub_with_borrow (bit_width_of_base_type T) c x y
+ | SubWithGetBorrow carry_boundary_bit_width _ _ _ T1 T2
+ => fun cxy => let '(c, x, y) := eta3 cxy in
+ sub_with_get_borrow (bit_width_of_base_type T1) (bit_width_of_base_type T2) carry_boundary_bit_width c x y
end%bounds.
Definition of_Z (z : Z) : t := ZToZRange z.
diff --git a/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v b/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v
index c74c319d5..68cc13d98 100644
--- a/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v
+++ b/src/Compilers/Z/Bounds/InterpretationLemmas/IsBoundedBy.v
@@ -268,7 +268,7 @@ Proof.
repeat (apply Z.max_case_strong || apply Zabs_ind); omega.
Qed.
-Local Existing Instances Z.log2_up_le_Proper Z.add_le_Proper.
+Local Existing Instances Z.log2_up_le_Proper Z.add_le_Proper Z.sub_with_borrow_le_Proper.
Lemma land_upper_lor_land_bounds a b
: Z.abs (Z.land a b) <= Bounds.upper_lor_and_bounds (Z.abs a) (Z.abs b).
Proof.
@@ -324,6 +324,10 @@ Local Arguments Z.pow : simpl never.
Local Arguments Z.add !_ !_.
Local Existing Instances Z.add_le_Proper Z.sub_le_flip_le_Proper Z.log2_up_le_Proper Z.pow_Zpos_le_Proper Z.sub_le_eq_Proper Z.add_with_carry_le_Proper.
Local Hint Extern 1 => progress cbv beta iota : typeclass_instances.
+Local Ltac ibbio_do_cbv :=
+ cbv [Bounds.interp_op Zinterp_op Z.add_with_get_carry SmartFlatTypeMapUnInterp Bounds.add_with_get_carry Bounds.sub_with_get_borrow Bounds.get_carry Bounds.get_borrow Z.get_carry cast_const]; cbn [fst snd].
+Local Ltac ibbio_prefin_by_apply :=
+ [ > | intros; apply is_bounded_by_truncation_bounds | simpl; reflexivity ].
Lemma is_bounded_by_interp_op t tR (opc : op t tR)
(bs : interp_flat_type Bounds.interp_base_type _)
(v : interp_flat_type interp_base_type _)
@@ -332,12 +336,17 @@ Lemma is_bounded_by_interp_op t tR (opc : op t tR)
Proof.
destruct opc;
[ apply is_bounded_by_truncation_bounds..
- | split;
- cbv [Bounds.interp_op Zinterp_op Z.add_with_get_carry SmartFlatTypeMapUnInterp Bounds.add_with_get_carry Z.get_carry cast_const]; cbn [fst snd];
+ | split; ibbio_do_cbv;
[ eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v mod _)) (v:=ZToInterp _);
- [ | intros; apply is_bounded_by_truncation_bounds | simpl; reflexivity ]
+ ibbio_prefin_by_apply
| eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v / _)) (v:=ZToInterp _);
- [ | intros; apply is_bounded_by_truncation_bounds | simpl; reflexivity ] ] ];
+ ibbio_prefin_by_apply ]
+ | apply is_bounded_by_truncation_bounds
+ | split; ibbio_do_cbv;
+ [ eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (v mod _)) (v:=ZToInterp _);
+ ibbio_prefin_by_apply
+ | eapply is_bounded_by_compose with (T1:=TZ) (f_v := fun v => ZToInterp (-(v / _))) (v:=ZToInterp _);
+ ibbio_prefin_by_apply ] ];
repeat first [ progress simpl in *
| progress cbv [interp_op lift_op cast_const Bounds.interp_base_type Bounds.is_bounded_by' ZRange.is_bounded_by'] in *
| progress destruct_head'_prod
@@ -388,4 +397,9 @@ Proof.
{ apply Z.mod_bound_min_max; auto. }
{ apply (@monotone_eight_corners true true true _ _ _); split; auto. }
{ auto with zarith. }
+ { apply (@monotone_eight_corners false true false _ _ _); split; auto. }
+ { apply (@monotone_eight_corners false true false _ _ _); split; auto. }
+ { apply Z.mod_bound_min_max; auto. }
+ { apply (@monotone_eight_corners false true false _ _ _); split; auto. }
+ { auto with zarith. }
Qed.
diff --git a/src/Compilers/Z/CommonSubexpressionElimination.v b/src/Compilers/Z/CommonSubexpressionElimination.v
index 81f6553ba..d545d9b47 100644
--- a/src/Compilers/Z/CommonSubexpressionElimination.v
+++ b/src/Compilers/Z/CommonSubexpressionElimination.v
@@ -23,12 +23,15 @@ Inductive symbolic_op :=
| SZselect
| SAddWithCarry
| SAddWithGetCarry (bitwidth : Z)
+| SSubWithBorrow
+| SSubWithGetBorrow (bitwidth : Z)
.
Definition symbolic_op_leb (x y : symbolic_op) : bool
:= match x, y with
| SOpConst z1, SOpConst z2 => Z.leb z1 z2
| SAddWithGetCarry bw1, SAddWithGetCarry bw2 => Z.leb bw1 bw2
+ | SSubWithGetBorrow bw1, SSubWithGetBorrow bw2 => Z.leb bw1 bw2
| SOpConst _, _ => true
| _, SOpConst _ => false
| SAdd, _ => true
@@ -51,8 +54,12 @@ Definition symbolic_op_leb (x y : symbolic_op) : bool
| _, SZselect => false
| SAddWithCarry, _ => true
| _, SAddWithCarry => false
- (*| SAddWithGetCarry _, _ => true
- | _, SAddWithGetCarry _ => false*)
+ | SAddWithGetCarry _, _ => true
+ | _, SAddWithGetCarry _ => false
+ | SSubWithBorrow, _ => true
+ | _, SSubWithBorrow => false
+ (*| SSubWithGetBorrow _, _ => true
+ | _, SSubWithGetBorrow _ => false*)
end.
Local Notation symbolic_expr := (@symbolic_expr base_type symbolic_op).
@@ -73,6 +80,8 @@ Definition symbolize_op s d (opc : op s d) : symbolic_op
| Zselect T1 T2 T3 Tout => SZselect
| AddWithCarry T1 T2 T3 Tout => SAddWithCarry
| AddWithGetCarry bitwidth T1 T2 T3 Tout1 Tout2 => SAddWithGetCarry bitwidth
+ | SubWithBorrow T1 T2 T3 Tout => SSubWithBorrow
+ | SubWithGetBorrow bitwidth T1 T2 T3 Tout1 Tout2 => SSubWithGetBorrow bitwidth
end.
Definition denote_symbolic_op s d (opc : symbolic_op) : option (op s d)
@@ -90,6 +99,9 @@ Definition denote_symbolic_op s d (opc : symbolic_op) : option (op s d)
| SAddWithCarry, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Tbase _ => Some (AddWithCarry _ _ _ _)
| SAddWithGetCarry bitwidth, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Prod (Tbase _) (Tbase _)
=> Some (AddWithGetCarry bitwidth _ _ _ _ _)
+ | SSubWithBorrow, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Tbase _ => Some (SubWithBorrow _ _ _ _)
+ | SSubWithGetBorrow bitwidth, Prod (Prod (Tbase _) (Tbase _)) (Tbase _), Prod (Tbase _) (Tbase _)
+ => Some (SubWithGetBorrow bitwidth _ _ _ _ _)
| SAdd, _, _
| SSub, _, _
| SMul, _, _
@@ -102,6 +114,8 @@ Definition denote_symbolic_op s d (opc : symbolic_op) : option (op s d)
| SZselect, _, _
| SAddWithCarry, _, _
| SAddWithGetCarry _, _, _
+ | SSubWithBorrow, _, _
+ | SSubWithGetBorrow _, _, _
=> None
end.
@@ -172,6 +186,8 @@ Definition normalize_symbolic_expr_mod_c (opc : symbolic_op) (args : symbolic_ex
| SZselect
| SAddWithCarry
| SAddWithGetCarry _
+ | SSubWithBorrow
+ | SSubWithGetBorrow _
=> args
end.
diff --git a/src/Compilers/Z/Syntax.v b/src/Compilers/Z/Syntax.v
index 7d0c84421..5bfa1168c 100644
--- a/src/Compilers/Z/Syntax.v
+++ b/src/Compilers/Z/Syntax.v
@@ -30,6 +30,8 @@ Inductive op : flat_type base_type -> flat_type base_type -> Type :=
| Zselect T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
| AddWithCarry T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
| AddWithGetCarry (bitwidth : Z) T1 T2 T3 Tout1 Tout2 : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout1 * Tbase Tout2)
+| SubWithBorrow T1 T2 T3 Tout : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout)
+| SubWithGetBorrow (bitwidth : Z) T1 T2 T3 Tout1 Tout2 : op (Tbase T1 * Tbase T2 * Tbase T3) (Tbase Tout1 * Tbase Tout2)
.
Definition interp_base_type (v : base_type) : Type :=
@@ -85,6 +87,8 @@ Definition Zinterp_op src dst (f : op src dst)
| Zselect _ _ _ _ => fun ctf => let '(c, t, f) := eta3 ctf in Z.zselect c t f
| AddWithCarry _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.add_with_carry c x y
| AddWithGetCarry bitwidth _ _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.add_with_get_carry bitwidth c x y
+ | SubWithBorrow _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.sub_with_borrow c x y
+ | SubWithGetBorrow bitwidth _ _ _ _ _ => fun cxy => let '(c, x, y) := eta3 cxy in Z.sub_with_get_borrow bitwidth c x y
end%Z.
Definition interp_op src dst (f : op src dst) : interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst
diff --git a/src/Compilers/Z/Syntax/Equality.v b/src/Compilers/Z/Syntax/Equality.v
index b2075c49c..a9003d354 100644
--- a/src/Compilers/Z/Syntax/Equality.v
+++ b/src/Compilers/Z/Syntax/Equality.v
@@ -47,6 +47,10 @@ Definition op_beq_hetero {t1 tR t1' tR'} (f : op t1 tR) (g : op t1' tR') : bool
=> base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout Tout'
| AddWithGetCarry bitwidth T1 T2 T3 Tout1 Tout2, AddWithGetCarry bitwidth' T1' T2' T3' Tout1' Tout2'
=> Z.eqb bitwidth bitwidth' && base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout1 Tout1' && base_type_beq Tout2 Tout2'
+ | SubWithBorrow T1 T2 T3 Tout, SubWithBorrow T1' T2' T3' Tout'
+ => base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout Tout'
+ | SubWithGetBorrow bitwidth T1 T2 T3 Tout1 Tout2, SubWithGetBorrow bitwidth' T1' T2' T3' Tout1' Tout2'
+ => Z.eqb bitwidth bitwidth' && base_type_beq T1 T1' && base_type_beq T2 T2' && base_type_beq T3 T3' && base_type_beq Tout1 Tout1' && base_type_beq Tout2 Tout2'
| OpConst _ _, _
| Add _ _ _, _
| Sub _ _ _, _
@@ -59,6 +63,8 @@ Definition op_beq_hetero {t1 tR t1' tR'} (f : op t1 tR) (g : op t1' tR') : bool
| Zselect _ _ _ _, _
| AddWithCarry _ _ _ _, _
| AddWithGetCarry _ _ _ _ _ _, _
+ | SubWithBorrow _ _ _ _, _
+ | SubWithGetBorrow _ _ _ _ _ _, _
=> false
end%bool.
diff --git a/src/Compilers/Z/Syntax/Util.v b/src/Compilers/Z/Syntax/Util.v
index daf3d65be..2338efc48 100644
--- a/src/Compilers/Z/Syntax/Util.v
+++ b/src/Compilers/Z/Syntax/Util.v
@@ -71,6 +71,8 @@ Definition genericize_op {var' src dst} (opc : op src dst) {f}
| Zselect _ _ _ _ => fun _ _ => Zselect _ _ _ _
| AddWithCarry _ _ _ _ => fun _ _ => AddWithCarry _ _ _ _
| AddWithGetCarry bitwidth _ _ _ _ _ => fun _ _ => AddWithGetCarry bitwidth _ _ _ _ _
+ | SubWithBorrow _ _ _ _ => fun _ _ => SubWithBorrow _ _ _ _
+ | SubWithGetBorrow bitwidth _ _ _ _ _ => fun _ _ => SubWithGetBorrow bitwidth _ _ _ _ _
end.
Lemma cast_const_id {t} v
diff --git a/src/Util/ZUtil/Definitions.v b/src/Util/ZUtil/Definitions.v
index 4a5d3a0ab..6fc969dfd 100644
--- a/src/Util/ZUtil/Definitions.v
+++ b/src/Util/ZUtil/Definitions.v
@@ -18,6 +18,16 @@ Module Z.
Definition add_get_carry (bitwidth : Z) (x y : Z) : Z * Z
:= add_with_get_carry bitwidth 0 x y.
+ Definition get_borrow (bitwidth : Z) (v : Z) : Z * Z
+ := let '(v, c) := get_carry bitwidth v in
+ (v, -c).
+ Definition sub_with_borrow (c : Z) (x y : Z) : Z
+ := add_with_carry (-c) x (-y).
+ Definition sub_with_get_borrow (bitwidth : Z) (c : Z) (x y : Z) : Z * Z
+ := get_borrow bitwidth (sub_with_borrow c x y).
+ Definition sub_get_borrow (bitwidth : Z) (x y : Z) : Z * Z
+ := sub_with_get_borrow bitwidth 0 x y.
+
(* splits at [bound], not [2^bitwidth]; wrapper to make add_getcarry
work if input is not known to be a power of 2 *)
Definition add_get_carry_full (bound : Z) (x y : Z) : Z * Z
diff --git a/src/Util/ZUtil/Morphisms.v b/src/Util/ZUtil/Morphisms.v
index e0ab179a8..285fa5261 100644
--- a/src/Util/ZUtil/Morphisms.v
+++ b/src/Util/ZUtil/Morphisms.v
@@ -46,4 +46,6 @@ Module Z.
Proof. intros ???; apply Z.pow_le_mono_r; try reflexivity; try assumption. Qed.
Lemma add_with_carry_le_Proper : Proper (Z.le ==> Z.le ==> Z.le ==> Z.le) Z.add_with_carry.
Proof. unfold Z.add_with_carry; repeat (omega || intro). Qed.
+ Lemma sub_with_borrow_le_Proper : Proper (Basics.flip Z.le ==> Z.le ==> Basics.flip Z.le ==> Z.le) Z.sub_with_borrow.
+ Proof. unfold Z.sub_with_borrow, Z.add_with_carry, Basics.flip; repeat (omega || intro). Qed.
End Z.