From 1d90fa730df7d1cb2ee726d3b41b9915ae4e4e2e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Sep 2019 23:02:06 +0200 Subject: moved trapping_mode to a more appropriate place --- common/AST.v | 9 +++++++++ 1 file changed, 9 insertions(+) (limited to 'common/AST.v') diff --git a/common/AST.v b/common/AST.v index a91138c9..d98f954a 100644 --- a/common/AST.v +++ b/common/AST.v @@ -193,6 +193,15 @@ Definition chunk_of_type (ty: typ) := Lemma chunk_of_Tptr: chunk_of_type Tptr = Mptr. Proof. unfold Mptr, Tptr; destruct Archi.ptr64; auto. Qed. +(** Trapping mode: does undefined behavior result in a trap or an undefined value (e.g. for loads) *) +Inductive trapping_mode : Type := TRAP | NOTRAP. + +Definition trapping_mode_eq : forall x y : trapping_mode, + { x=y } + { x <> y}. +Proof. + decide equality. +Defined. + (** Initialization data for global variables. *) Inductive init_data: Type := -- cgit From c4cc75dc6abcb0eee6f3288e96fea4aec540fd68 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 5 Sep 2019 11:19:22 +0200 Subject: more proofs going through --- common/AST.v | 1 + 1 file changed, 1 insertion(+) (limited to 'common/AST.v') diff --git a/common/AST.v b/common/AST.v index d98f954a..bb8508b7 100644 --- a/common/AST.v +++ b/common/AST.v @@ -202,6 +202,7 @@ Proof. decide equality. Defined. + (** Initialization data for global variables. *) Inductive init_data: Type := -- cgit From 7b4e6a522bcf1f247ef9b3517af328b5da670a98 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 4 Oct 2019 17:57:25 +0200 Subject: Ibuiltin proof --- common/AST.v | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) (limited to 'common/AST.v') 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 -- cgit