aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2019-07-04 17:49:11 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-07-17 09:17:28 +0200
commit633e60ed36c07c4b6cb4b1dc93b9eea312882ceb (patch)
tree22feeaf195a61a3ffecf280717ddbde60987a5c7
parent10aa130361a5a673a14a7b38ed9c077103f9155f (diff)
downloadcompcert-633e60ed36c07c4b6cb4b1dc93b9eea312882ceb.tar.gz
compcert-633e60ed36c07c4b6cb4b1dc93b9eea312882ceb.zip
Make __builtin_sel available from C source code
It is type-checked like a conditional expression then translated to a call to the known builtin function.
-rw-r--r--cfrontend/C2C.ml9
-rw-r--r--cfrontend/Cop.v21
-rw-r--r--cfrontend/Csem.v69
-rw-r--r--cfrontend/Csyntax.v12
-rw-r--r--cfrontend/Ctyping.v77
-rw-r--r--cparser/Elab.ml36
-rw-r--r--extraction/extraction.v3
7 files changed, 195 insertions, 32 deletions
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index 4d0e52e1..2e6bbe67 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -181,6 +181,11 @@ let builtins_generic = {
TInt(IULong, []);
TInt(IULong, [])],
false);
+ (* Selection *)
+ "__builtin_sel",
+ (TVoid [],
+ [TInt(C.IBool, [])],
+ true);
(* Annotations *)
"__builtin_annot",
(TVoid [],
@@ -941,6 +946,10 @@ let rec convertExpr env e =
Econs(va_list_ptr dst, Econs(va_list_ptr src, Enil)),
Tvoid)
+ | C.ECall({edesc = C.EVar {name = "__builtin_sel"}}, [arg1; arg2; arg3]) ->
+ ewrap (Ctyping.eselection (convertExpr env arg1)
+ (convertExpr env arg2) (convertExpr env arg3))
+
| C.ECall({edesc = C.EVar {name = "printf"}}, args)
when !Clflags.option_interp ->
let targs = convertTypArgs env [] args
diff --git a/cfrontend/Cop.v b/cfrontend/Cop.v
index c395a2cb..782fb32a 100644
--- a/cfrontend/Cop.v
+++ b/cfrontend/Cop.v
@@ -1580,6 +1580,27 @@ Proof.
intros. apply cast_val_casted. eapply cast_val_is_casted; eauto.
Qed.
+(** Moreover, casted values belong to the machine type corresponding to the
+ C type. *)
+
+Lemma val_casted_has_type:
+ forall v ty, val_casted v ty -> ty <> Tvoid -> Val.has_type v (typ_of_type ty).
+Proof.
+ intros. inv H; simpl typ_of_type.
+- exact I.
+- exact I.
+- exact I.
+- exact I.
+- apply Val.Vptr_has_type.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- apply Val.Vptr_has_type.
+- apply Val.Vptr_has_type.
+- congruence.
+Qed.
+
(** Relation with the arithmetic conversions of ISO C99, section 6.3.1 *)
Module ArithConv.
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index 0c3e00de..a76a14ba 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -15,19 +15,9 @@
(** Dynamic semantics for the Compcert C language *)
-Require Import Coqlib.
-Require Import Errors.
-Require Import Maps.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import AST.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Ctypes.
-Require Import Cop.
-Require Import Csyntax.
+Require Import Coqlib Errors Maps.
+Require Import Integers Floats Values AST Memory Builtins Events Globalenvs.
+Require Import Ctypes Cop Csyntax.
Require Import Smallstep.
(** * Operational semantics *)
@@ -437,6 +427,59 @@ Definition not_stuck (e: expr) (m: mem) : Prop :=
forall k C e' ,
context k RV C -> e = C e' -> imm_safe k e' m.
+(** ** Derived forms. *)
+
+(** The following are admissible reduction rules for some derived forms
+ of the CompCert C language. They help showing that the derived forms
+ make sense. *)
+
+Lemma red_selection:
+ forall v1 ty1 v2 ty2 v3 ty3 ty m b v2' v3',
+ ty <> Tvoid ->
+ bool_val v1 ty1 m = Some b ->
+ sem_cast v2 ty2 ty m = Some v2' ->
+ sem_cast v3 ty3 ty m = Some v3' ->
+ rred (Eselection (Eval v1 ty1) (Eval v2 ty2) (Eval v3 ty3) ty) m
+ E0 (Eval (if b then v2' else v3') ty) m.
+Proof.
+ intros. unfold Eselection.
+ set (t := typ_of_type ty).
+ set (sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default).
+ assert (LK: lookup_builtin_function "__builtin_sel"%string sg = Some (BI_standard (BI_select t))).
+ { unfold sg, t; destruct ty as [ | ? ? ? | ? | [] ? | ? ? | ? ? ? | ? ? ? | ? ? | ? ? ];
+ simpl; unfold Tptr; destruct Archi.ptr64; reflexivity. }
+ set (v' := if b then v2' else v3').
+ assert (C: val_casted v' ty).
+ { unfold v'; destruct b; eapply cast_val_is_casted; eauto. }
+ assert (EQ: Val.normalize v' t = v').
+ { apply Val.normalize_idem. apply val_casted_has_type; auto. }
+ econstructor.
+- constructor. rewrite cast_bool_bool_val, H0. eauto.
+ constructor. eauto.
+ constructor. eauto.
+ constructor.
+- red. red. rewrite LK. constructor. simpl. rewrite <- EQ.
+ destruct b; auto.
+Qed.
+
+Lemma ctx_selection_1:
+ forall k C r2 r3 ty, context k RV C -> context k RV (fun x => Eselection (C x) r2 r3 ty).
+Proof.
+ intros. apply ctx_builtin. constructor; auto.
+Qed.
+
+Lemma ctx_selection_2:
+ forall k r1 C r3 ty, context k RV C -> context k RV (fun x => Eselection r1 (C x) r3 ty).
+Proof.
+ intros. apply ctx_builtin. constructor; constructor; auto.
+Qed.
+
+Lemma ctx_selection_3:
+ forall k r1 r2 C ty, context k RV C -> context k RV (fun x => Eselection r1 r2 (C x) ty).
+Proof.
+ intros. apply ctx_builtin. constructor; constructor; constructor; auto.
+Qed.
+
End EXPR.
(** ** Transition semantics. *)
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index 00565309..c34a5e13 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -100,6 +100,18 @@ Definition Epreincr (id: incr_or_decr) (l: expr) (ty: type) :=
Eassignop (match id with Incr => Oadd | Decr => Osub end)
l (Eval (Vint Int.one) type_int32s) (typeconv ty) ty.
+(** Selection is a conditional expression that always evaluates both arms
+ and can be implemented by "conditional move" instructions.
+ It is expressed as an invocation of a builtin function. *)
+
+Definition Eselection (r1 r2 r3: expr) (ty: type) :=
+ let t := typ_of_type ty in
+ let sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default in
+ Ebuiltin (EF_builtin "__builtin_sel"%string sg)
+ (Tcons type_bool (Tcons ty (Tcons ty Tnil)))
+ (Econs r1 (Econs r2 (Econs r3 Enil)))
+ ty.
+
(** Extract the type part of a type-annotated expression. *)
Definition typeof (a: expr) : type :=
diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v
index ba1d34fb..b92a9bac 100644
--- a/cfrontend/Ctyping.v
+++ b/cfrontend/Ctyping.v
@@ -18,7 +18,7 @@
Require Import String.
Require Import Coqlib Maps Integers Floats Errors.
Require Import AST Linking.
-Require Import Values Memory Globalenvs Events.
+Require Import Values Memory Globalenvs Builtins Events.
Require Import Ctypes Cop Csyntax Csem.
Local Open Scope error_monad_scope.
@@ -392,13 +392,17 @@ Inductive wt_rvalue : expr -> Prop :=
classify_fun (typeof r1) = fun_case_f tyargs tyres cconv ->
wt_arguments rargs tyargs ->
wt_rvalue (Ecall r1 rargs tyres)
- | wt_Ebuiltin: forall ef tyargs rargs,
+ | wt_Ebuiltin: forall ef tyargs rargs ty,
wt_exprlist rargs ->
wt_arguments rargs tyargs ->
- (* This is specialized to builtins returning void, the only
- case generated by C2C. *)
- sig_res (ef_sig ef) = None ->
- wt_rvalue (Ebuiltin ef tyargs rargs Tvoid)
+ (* This typing rule is specialized to the builtin invocations generated
+ by C2C, which are either __builtin_sel or builtins returning void. *)
+ (ty = Tvoid /\ sig_res (ef_sig ef) = None)
+ \/ (tyargs = Tcons type_bool (Tcons ty (Tcons ty Tnil))
+ /\ let t := typ_of_type ty in
+ let sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default in
+ ef = EF_builtin "__builtin_sel"%string sg) ->
+ wt_rvalue (Ebuiltin ef tyargs rargs ty)
| wt_Eparen: forall r tycast ty,
wt_rvalue r ->
wt_cast (typeof r) tycast -> subtype tycast ty ->
@@ -745,6 +749,12 @@ Definition ebuiltin (ef: external_function) (tyargs: typelist) (args: exprlist)
then OK (Ebuiltin ef tyargs args tyres)
else Error (msg "builtin: wrong type decoration").
+Definition eselection (r1 r2 r3: expr) : res expr :=
+ do x1 <- check_rval r1; do x2 <- check_rval r2; do x3 <- check_rval r3;
+ do y1 <- check_bool (typeof r1);
+ do ty <- type_conditional (typeof r2) (typeof r3);
+ OK (Eselection r1 r2 r3 ty).
+
Definition sdo (a: expr) : res statement :=
do x <- check_rval a; OK (Sdo a).
@@ -981,6 +991,15 @@ Proof.
destruct i; auto.
Qed.
+Lemma wt_bool_cast:
+ forall ty, wt_bool ty -> wt_cast ty type_bool.
+Proof.
+ unfold wt_bool, wt_cast; unfold classify_bool; intros.
+ destruct ty; simpl in *; try congruence;
+ try (destruct Archi.ptr64; congruence).
+ destruct f; congruence.
+Qed.
+
Lemma wt_cast_int:
forall i1 s1 a1 i2 s2 a2, wt_cast (Tint i1 s1 a1) (Tint i2 s2 a2).
Proof.
@@ -1225,6 +1244,17 @@ Proof.
econstructor; eauto. eapply check_arguments_sound; eauto.
Qed.
+Lemma eselection_sound:
+ forall r1 r2 r3 a, eselection r1 r2 r3 = OK a ->
+ wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a.
+Proof.
+ intros. monadInv H. apply type_conditional_cast in EQ3. destruct EQ3.
+ eapply wt_Ebuiltin.
+ repeat (constructor; eauto with ty).
+ repeat (constructor; eauto with ty). apply wt_bool_cast; eauto with ty.
+ right; auto.
+Qed.
+
Lemma sdo_sound:
forall a s, sdo a = OK s -> wt_expr ce e a -> wt_stmt ce e rt s.
Proof.
@@ -1632,15 +1662,6 @@ Proof.
destruct f; discriminate.
Qed.
-Lemma wt_bool_cast:
- forall ty, wt_bool ty -> wt_cast ty type_bool.
-Proof.
- unfold wt_bool, wt_cast; unfold classify_bool; intros.
- destruct ty; simpl in *; try congruence;
- try (destruct Archi.ptr64; congruence).
- destruct f; congruence.
-Qed.
-
Lemma wt_cast_self:
forall t1 t2, wt_cast t1 t2 -> wt_cast t2 t2.
Proof.
@@ -1725,7 +1746,27 @@ Proof.
constructor; auto.
- (* comma *) auto.
- (* paren *) inv H3. constructor. apply H5. eapply pres_sem_cast; eauto.
-- (* builtin *) subst. auto with ty.
+- (* builtin *) subst. destruct H7 as [(A & B) | (A & B)].
++ subst ty. auto with ty.
++ simpl in B. set (T := typ_of_type ty) in *.
+ set (sg := mksignature (AST.Tint :: T :: T :: nil) (Some T) cc_default) in *.
+ assert (LK: lookup_builtin_function "__builtin_sel"%string sg = Some (BI_standard (BI_select T))).
+ { unfold sg, T; destruct ty as [ | ? ? ? | ? | [] ? | ? ? | ? ? ? | ? ? ? | ? ? | ? ? ];
+ simpl; unfold Tptr; destruct Archi.ptr64; reflexivity. }
+ subst ef. red in H0. red in H0. rewrite LK in H0. inv H0.
+ inv H. inv H8. inv H9. inv H10. simpl in H1.
+ assert (A: val_casted v1 type_bool) by (eapply cast_val_is_casted; eauto).
+ inv A.
+ set (v' := if Int.eq n Int.zero then v4 else v2) in *.
+ constructor.
+ destruct (type_eq ty Tvoid).
+ subst. constructor.
+ inv H1.
+ assert (C: val_casted v' ty).
+ { unfold v'; destruct (Int.eq n Int.zero); eapply cast_val_is_casted; eauto. }
+ assert (EQ: Val.normalize v' T = v').
+ { apply Val.normalize_idem. apply val_casted_has_type; auto. }
+ rewrite EQ. apply wt_val_casted; auto.
Qed.
Lemma wt_lred:
@@ -1767,8 +1808,8 @@ with wt_subexprlist:
wt_exprlist cenv tenv (C a) ->
wt_expr_kind cenv tenv from a.
Proof.
- destruct 1; intros WT; auto; inv WT; eauto.
- destruct 1; intros WT; inv WT; eauto.
+- destruct 1; intros WT; auto; inv WT; eauto.
+- destruct 1; intros WT; inv WT; eauto.
Qed.
Lemma typeof_context:
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index eea60127..3797164d 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -1802,6 +1802,42 @@ let elab_expr ctx loc env a =
(print_typ env) ty (print_typ env) ty' (print_typ env) ty' (print_typ env) ty;
{ edesc = ECall(ident, [b2; b3]); etyp = ty },env
+ | CALL((VARIABLE "__builtin_sel" as a0), al) ->
+ begin match al with
+ | [a1; a2; a3] ->
+ let b0,env = elab env a0 in
+ let b1,env = elab env a1 in
+ let b2,env = elab env a2 in
+ let b3,env = elab env a3 in
+ if not (is_scalar_type env b1.etyp) then
+ error "first argument of '__builtin_sel' is not a scalar type (invalid %a)"
+ (print_typ env) b1.etyp;
+ let tyres =
+ match pointer_decay env b2.etyp, pointer_decay env b3.etyp with
+ | (TInt _ | TFloat _ | TEnum _), (TInt _ | TFloat _ | TEnum _) ->
+ binary_conversion env b2.etyp b3.etyp
+ | (TPtr(ty1, a1) as pty1), (TPtr(ty2, a2) as pty2) ->
+ if is_void_type env ty1 || is_void_type env ty2 then
+ TPtr(TVoid (add_attributes a1 a2), [])
+ else begin
+ match combine_types AttrIgnoreAll env pty1 pty2 with
+ | None ->
+ warning Pointer_type_mismatch "the second and third arguments of '__builtin_sel' have incompatible pointer types (%a and %a)"
+ (print_typ env) pty1 (print_typ env) pty2;
+ (* tolerance *)
+ TPtr(TVoid (add_attributes a1 a2), [])
+ | Some ty -> ty
+ end
+ | _, _ ->
+ fatal_error "wrong types (%a and %a) for the second and third arguments of '__builtin_sel'"
+ (print_typ env) b2.etyp (print_typ env) b3.etyp
+
+ in
+ { edesc = ECall(b0, [b1; b2; b3]); etyp = tyres }, env
+ | _ ->
+ fatal_error "'__builtin_sel' expect 3 arguments"
+ end
+
| CALL(a1, al) ->
let b1,env =
(* Catch the old-style usage of calling a function without
diff --git a/extraction/extraction.v b/extraction/extraction.v
index 25319475..521c0cdd 100644
--- a/extraction/extraction.v
+++ b/extraction/extraction.v
@@ -162,9 +162,10 @@ Separate Extraction
Cexec.do_initial_state Cexec.do_step Cexec.at_final_state
Ctypes.merge_attributes Ctypes.remove_attributes Ctypes.build_composite_env
Initializers.transl_init Initializers.constval
- Csyntax.Eindex Csyntax.Epreincr
+ Csyntax.Eindex Csyntax.Epreincr Csyntax.Eselection
Ctyping.typecheck_program
Ctyping.epostincr Ctyping.epostdecr Ctyping.epreincr Ctyping.epredecr
+ Ctyping.eselection
Ctypes.make_program
Clight.type_of_function
Conventions1.callee_save_type Conventions1.is_float_reg