aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--backend/Duplicate.v24
-rw-r--r--backend/Duplicateproof.v7
-rw-r--r--common/AST.v19
-rw-r--r--lib/Floats.v63
-rw-r--r--lib/Integers.v38
5 files changed, 138 insertions, 13 deletions
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
index a5e7d92a..c313e3fa 100644
--- a/backend/Duplicate.v
+++ b/backend/Duplicate.v
@@ -53,6 +53,20 @@ Proof.
intros H H'. intros. decide equality.
Qed.
+(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application"
+ * error when doing so *)
+Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}.
+Proof.
+ intros.
+ apply (builtin_arg_eq Pos.eq_dec).
+Defined.
+Global Opaque builtin_arg_eq_pos.
+
+Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}.
+Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed.
+Global Opaque builtin_res_eq_pos.
+
+
Definition verify_match_inst revmap inst tinst :=
match inst with
| Inop n => match tinst with Inop n' => do u <- verify_is_copy revmap n n'; OK tt | _ => Error(msg "verify_match_inst Inop") end
@@ -103,6 +117,16 @@ Definition verify_match_inst revmap inst tinst :=
else Error (msg "Different ri in Icall")
else Error (msg "Different signatures in Icall")
| _ => Error (msg "verify_match_inst Icall") end
+ | Ibuiltin ef lbar brr n => match tinst with
+ | Ibuiltin ef' lbar' brr' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (external_function_eq ef ef') then
+ if (list_eq_dec builtin_arg_eq_pos lbar lbar') then
+ if (builtin_res_eq_pos brr brr') then OK tt
+ else Error (msg "Different brr in Ibuiltin")
+ else Error (msg "Different lbar in Ibuiltin")
+ else Error (msg "Different ef in Ibuiltin")
+ | _ => Error (msg "verify_match_inst Ibuiltin") end
| _ => Error(msg "not implemented")
end.
diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v
index e9799f08..7369c3ea 100644
--- a/backend/Duplicateproof.v
+++ b/backend/Duplicateproof.v
@@ -114,6 +114,13 @@ Proof.
destruct (Pos.eq_dec _ _); try discriminate.
eapply verify_is_copy_correct_one. destruct x. eassumption. subst.
constructor.
+(* Ibuiltin *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct (external_function_eq _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (builtin_res_eq_pos _ _); try discriminate.
+ eapply verify_is_copy_correct_one. destruct x. eassumption. subst.
+ constructor.
Qed.
Lemma verify_mapping_mn_correct:
diff --git a/common/AST.v b/common/AST.v
index a91138c9..7ffe355d 100644
--- a/common/AST.v
+++ b/common/AST.v
@@ -17,7 +17,7 @@
the abstract syntax trees of many of the intermediate languages. *)
Require Import String.
-Require Import Coqlib Maps Errors Integers Floats.
+Require Import Coqlib Maps Errors Integers Floats BinPos.
Require Archi.
Set Implicit Arguments.
@@ -630,11 +630,28 @@ Inductive builtin_arg (A: Type) : Type :=
| BA_splitlong (hi lo: builtin_arg A)
| BA_addptr (a1 a2: builtin_arg A).
+Definition builtin_arg_eq {A: Type}:
+ (forall x y : A, {x = y} + {x <> y}) ->
+ forall (ba1 ba2: (builtin_arg A)), {ba1=ba2} + {ba1<>ba2}.
+Proof.
+ intro. generalize Integers.int_eq int64_eq float_eq float32_eq chunk_eq ptrofs_eq ident_eq.
+ decide equality.
+Defined.
+Global Opaque builtin_arg_eq.
+
Inductive builtin_res (A: Type) : Type :=
| BR (x: A)
| BR_none
| BR_splitlong (hi lo: builtin_res A).
+Definition builtin_res_eq {A: Type}:
+ (forall x y : A, {x = y} + {x <> y}) ->
+ forall (a b: builtin_res A), {a=b} + {a<>b}.
+Proof.
+ intro. decide equality.
+Defined.
+Global Opaque builtin_res_eq.
+
Fixpoint globals_of_builtin_arg (A: Type) (a: builtin_arg A) : list ident :=
match a with
| BA_loadglobal chunk id ofs => id :: nil
diff --git a/lib/Floats.v b/lib/Floats.v
index 13350dd0..272efa52 100644
--- a/lib/Floats.v
+++ b/lib/Floats.v
@@ -16,7 +16,7 @@
(** Formalization of floating-point numbers, using the Flocq library. *)
-Require Import Coqlib Zbits Integers.
+Require Import Coqlib Zbits Integers Axioms.
(*From Flocq*)
Require Import Binary Bits Core.
Require Import IEEE754_extra.
@@ -27,8 +27,69 @@ Close Scope R_scope.
Open Scope Z_scope.
Definition float := binary64. (**r the type of IEE754 double-precision FP numbers *)
+
+Definition float_eq: forall (i1 i2: float), {i1=i2} + {i1<>i2}.
+Proof.
+ intros. destruct i1.
+(* B754_zero *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_infinity *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_nan *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + generalize (Pos.eq_dec pl pl0). intro. inv H.
+ ++ left. apply eqb_prop in BEQ. subst.
+ assert (e = e0) by (apply proof_irr). congruence.
+ ++ right. intro. inv H. contradiction.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_finite *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ; [apply eqb_prop in BEQ | apply eqb_false_iff in BEQ].
+ generalize (Pos.eq_dec m m0). intro. inv H.
+ generalize (Z.eq_dec e e1). intro. inv H.
+ 1: { left. assert (e0 = e2) by (apply proof_irr). congruence. }
+ all: right; intro; inv H; contradiction.
+Qed.
+
Definition float32 := binary32. (**r the type of IEE754 single-precision FP numbers *)
+Definition float32_eq: forall (i1 i2: float32), {i1=i2} + {i1<>i2}.
+Proof.
+ intros. destruct i1.
+(* B754_zero *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_infinity *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_nan *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + generalize (Pos.eq_dec pl pl0). intro. inv H.
+ ++ left. apply eqb_prop in BEQ. subst.
+ assert (e = e0) by (apply proof_irr). congruence.
+ ++ right. intro. inv H. contradiction.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_finite *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ; [apply eqb_prop in BEQ | apply eqb_false_iff in BEQ].
+ generalize (Pos.eq_dec m m0). intro. inv H.
+ generalize (Z.eq_dec e e1). intro. inv H.
+ 1: { left. assert (e0 = e2) by (apply proof_irr). congruence. }
+ all: right; intro; inv H; contradiction.
+Qed.
+
(** Boolean-valued comparisons *)
Definition cmp_of_comparison (c: comparison) (x: option Datatypes.comparison) : bool :=
diff --git a/lib/Integers.v b/lib/Integers.v
index fca44b1e..9c6fcf1d 100644
--- a/lib/Integers.v
+++ b/lib/Integers.v
@@ -101,17 +101,6 @@ Hint Resolve modulus_pos: ints.
Record int: Type := mkint { intval: Z; intrange: -1 < intval < modulus }.
-Definition int_eq: forall (i1 i2: int), {i1=i2} + {i1<>i2}.
-Proof.
- generalize Z.eq_dec. intros.
- destruct i1. destruct i2. generalize (H intval0 intval1). intro.
- inversion H0.
- - subst. left. assert (intrange0 = intrange1) by (apply proof_irr). congruence.
- - right. intro. inversion H2. contradiction.
-Qed.
-
-(* TODO - continue for the rest *)
-
(** Fast normalization modulo [2^wordsize] *)
Definition Z_mod_modulus (x: Z) : Z :=
@@ -4167,8 +4156,26 @@ End Int64.
Strategy 0 [Wordsize_64.wordsize].
+Definition int_eq: forall (i1 i2: int), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Notation int64 := Int64.int.
+Definition int64_eq: forall (i1 i2: int64), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Global Opaque Int.repr Int64.repr Byte.repr.
(** * Specialization to offsets in pointer values *)
@@ -4445,6 +4452,15 @@ Strategy 0 [Wordsize_Ptrofs.wordsize].
Notation ptrofs := Ptrofs.int.
+Definition ptrofs_eq: forall (i1 i2: ptrofs), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Global Opaque Ptrofs.repr.
Hint Resolve Int.modulus_pos Int.eqm_refl Int.eqm_refl2 Int.eqm_sym Int.eqm_trans