From 7a6bb90048db7a254e959b1e3c308bac5fe6c418 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 11 Oct 2015 17:43:59 +0200 Subject: Use Coq strings instead of idents to name external and builtin functions. The AST.ident type represents source-level identifiers as unique positive numbers. However, the mapping identifiers <-> AST.ident differs between runs of CompCert on different source files. This is problematic when we need to produce or recognize external functions and builtin functions with fixed names, for example: * in $ARCH/Machregs.v to define the register conventions for builtin functions; * in the VST program logic from Princeton to treat thread primitives specially. So far, we used AST.ident_of_string to recover the ident associated with a string. However, this function is defined in OCaml and doesn't execute within Coq. This is a problem both for VST and for future executability of CompCert within Coq. This commit replaces "ident" by "string" in the arguments of EF_external, EF_builtin, EF_inline_asm, EF_annot, and EF_annot_val. This provides stable names for externals and builtins, as needed. For inline asm and annotations, it's a matter of taste, but using strings feels more natural. EF_debug keeps using idents, since some kinds of EF_debug annotations talk about program variables. --- ia32/Machregs.v | 23 ++++++++--------------- 1 file changed, 8 insertions(+), 15 deletions(-) (limited to 'ia32/Machregs.v') diff --git a/ia32/Machregs.v b/ia32/Machregs.v index ace193b7..f3801900 100644 --- a/ia32/Machregs.v +++ b/ia32/Machregs.v @@ -129,17 +129,14 @@ Fixpoint destroyed_by_clobber (cl: list string): list mreg := end end. -Definition builtin_write16_reversed := ident_of_string "__builtin_write16_reversed". -Definition builtin_write32_reversed := ident_of_string "__builtin_write32_reversed". - Definition destroyed_by_builtin (ef: external_function): list mreg := match ef with | EF_memcpy sz al => if zle sz 32 then CX :: X7 :: nil else CX :: SI :: DI :: nil | EF_vstore (Mint8unsigned|Mint8signed) => AX :: CX :: nil - | EF_builtin id sg => - if ident_eq id builtin_write16_reversed - || ident_eq id builtin_write32_reversed + | EF_builtin name sg => + if string_dec name "__builtin_write16_reversed" + || string_dec name "__builtin_write32_reversed" then CX :: DX :: nil else nil | EF_inline_asm txt sg clob => destroyed_by_clobber clob | _ => nil @@ -173,21 +170,17 @@ Definition mregs_for_operation (op: operation): list (option mreg) * option mreg | _ => (nil, None) end. -Definition builtin_negl := ident_of_string "__builtin_negl". -Definition builtin_addl := ident_of_string "__builtin_addl". -Definition builtin_subl := ident_of_string "__builtin_subl". -Definition builtin_mull := ident_of_string "__builtin_mull". - Definition mregs_for_builtin (ef: external_function): list (option mreg) * list (option mreg) := match ef with | EF_memcpy sz al => if zle sz 32 then (Some AX :: Some DX :: nil, nil) else (Some DI :: Some SI :: nil, nil) - | EF_builtin id sg => - if ident_eq id builtin_negl then + | EF_builtin name sg => + if string_dec name "__builtin_negl" then (Some DX :: Some AX :: nil, Some DX :: Some AX :: nil) - else if ident_eq id builtin_addl || ident_eq id builtin_subl then + else if string_dec name "__builtin_addl" + || string_dec name "__builtin_subl" then (Some DX :: Some AX :: Some CX :: Some BX :: nil, Some DX :: Some AX :: nil) - else if ident_eq id builtin_mull then + else if string_dec name "__builtin_mull" then (Some AX :: Some DX :: nil, Some DX :: Some AX :: nil) else (nil, nil) -- cgit