Skip to content

Instantly share code, notes, and snippets.

@ildyria
Last active November 2, 2017 15:53
Show Gist options
  • Save ildyria/9f3d00499f637f8c5bf6fdd6404257f8 to your computer and use it in GitHub Desktop.
Save ildyria/9f3d00499f637f8c5bf6fdd6404257f8 to your computer and use it in GitHub Desktop.
initialization of arrays is not convenient.
Require Import VST.floyd.proofauto.
(*
#define FOR(i,n) for (i = 0;i < n;++i)
#define sv static void
typedef long long _Alignas(8) i64;
typedef i64 gf[16];
sv set(gf r, const gf a)
{
int i;
FOR(i,16) r[i]=a[i];
}
sv call_set(gf a)
{
int i; // dummy stuff
gf r;
i = 0;
set(r,a);
}
int main(){
gf a;
call_set(a);
}
*)
Local Open Scope Z_scope.
Definition ___builtin_annot : ident := 3%positive.
Definition ___builtin_annot_intval : ident := 4%positive.
Definition ___builtin_bswap : ident := 27%positive.
Definition ___builtin_bswap16 : ident := 29%positive.
Definition ___builtin_bswap32 : ident := 28%positive.
Definition ___builtin_clz : ident := 30%positive.
Definition ___builtin_clzl : ident := 31%positive.
Definition ___builtin_clzll : ident := 32%positive.
Definition ___builtin_ctz : ident := 33%positive.
Definition ___builtin_ctzl : ident := 34%positive.
Definition ___builtin_ctzll : ident := 35%positive.
Definition ___builtin_debug : ident := 48%positive.
Definition ___builtin_fabs : ident := 1%positive.
Definition ___builtin_fmadd : ident := 39%positive.
Definition ___builtin_fmax : ident := 37%positive.
Definition ___builtin_fmin : ident := 38%positive.
Definition ___builtin_fmsub : ident := 40%positive.
Definition ___builtin_fnmadd : ident := 41%positive.
Definition ___builtin_fnmsub : ident := 42%positive.
Definition ___builtin_fsqrt : ident := 36%positive.
Definition ___builtin_membar : ident := 5%positive.
Definition ___builtin_memcpy_aligned : ident := 2%positive.
Definition ___builtin_nop : ident := 47%positive.
Definition ___builtin_read16_reversed : ident := 43%positive.
Definition ___builtin_read32_reversed : ident := 44%positive.
Definition ___builtin_va_arg : ident := 7%positive.
Definition ___builtin_va_copy : ident := 8%positive.
Definition ___builtin_va_end : ident := 9%positive.
Definition ___builtin_va_start : ident := 6%positive.
Definition ___builtin_write16_reversed : ident := 45%positive.
Definition ___builtin_write32_reversed : ident := 46%positive.
Definition ___compcert_va_composite : ident := 13%positive.
Definition ___compcert_va_float64 : ident := 12%positive.
Definition ___compcert_va_int32 : ident := 10%positive.
Definition ___compcert_va_int64 : ident := 11%positive.
Definition ___i64_dtos : ident := 14%positive.
Definition ___i64_dtou : ident := 15%positive.
Definition ___i64_sar : ident := 26%positive.
Definition ___i64_sdiv : ident := 20%positive.
Definition ___i64_shl : ident := 24%positive.
Definition ___i64_shr : ident := 25%positive.
Definition ___i64_smod : ident := 22%positive.
Definition ___i64_stod : ident := 16%positive.
Definition ___i64_stof : ident := 18%positive.
Definition ___i64_udiv : ident := 21%positive.
Definition ___i64_umod : ident := 23%positive.
Definition ___i64_utod : ident := 17%positive.
Definition ___i64_utof : ident := 19%positive.
Definition _a : ident := 50%positive.
Definition _call_set : ident := 53%positive.
Definition _i : ident := 51%positive.
Definition _main : ident := 54%positive.
Definition _r : ident := 49%positive.
Definition _set : ident := 52%positive.
Definition _t'1 : ident := 55%positive.
Definition f_set := {|
fn_return := tvoid;
fn_callconv := cc_default;
fn_params := ((_r, (tptr (talignas 3%N tlong))) ::
(_a, (tptr (talignas 3%N tlong))) :: nil);
fn_vars := nil;
fn_temps := ((_i, tint) :: (_t'1, (talignas 3%N tlong)) :: nil);
fn_body :=
(Ssequence
(Sset _i (Econst_int (Int.repr 0) tint))
(Sloop
(Ssequence
(Sifthenelse (Ebinop Olt (Etempvar _i tint)
(Econst_int (Int.repr 16) tint) tint)
Sskip
Sbreak)
(Ssequence
(Sset _t'1
(Ederef
(Ebinop Oadd (Etempvar _a (tptr (talignas 3%N tlong)))
(Etempvar _i tint) (tptr (talignas 3%N tlong)))
(talignas 3%N tlong)))
(Sassign
(Ederef
(Ebinop Oadd (Etempvar _r (tptr (talignas 3%N tlong)))
(Etempvar _i tint) (tptr (talignas 3%N tlong)))
(talignas 3%N tlong)) (Etempvar _t'1 (talignas 3%N tlong)))))
(Sset _i
(Ebinop Oadd (Etempvar _i tint) (Econst_int (Int.repr 1) tint) tint))))
|}.
Definition f_call_set := {|
fn_return := tvoid;
fn_callconv := cc_default;
fn_params := ((_a, (tptr (talignas 3%N tlong))) :: nil);
fn_vars := ((_r, (tarray (talignas 3%N tlong) 16)) :: nil);
fn_temps := ((_i, tint) :: nil);
fn_body :=
(Ssequence
(Sset _i (Econst_int (Int.repr 0) tint))
(Scall None
(Evar _set (Tfunction
(Tcons (tptr (talignas 3%N tlong))
(Tcons (tptr (talignas 3%N tlong)) Tnil)) tvoid
cc_default))
((Evar _r (tarray (talignas 3%N tlong) 16)) ::
(Etempvar _a (tptr (talignas 3%N tlong))) :: nil)))
|}.
Definition f_main := {|
fn_return := tint;
fn_callconv := cc_default;
fn_params := nil;
fn_vars := ((_a, (tarray (talignas 3%N tlong) 16)) :: nil);
fn_temps := nil;
fn_body :=
(Ssequence
(Scall None
(Evar _call_set (Tfunction (Tcons (tptr (talignas 3%N tlong)) Tnil) tvoid
cc_default))
((Evar _a (tarray (talignas 3%N tlong) 16)) :: nil))
(Sreturn (Some (Econst_int (Int.repr 0) tint))))
|}.
Definition composites : list composite_definition :=
nil.
Definition prog : Clight.program := {|
prog_defs :=
((___builtin_fabs,
Gfun(External (EF_builtin "__builtin_fabs"
(mksignature (AST.Tfloat :: nil) (Some AST.Tfloat)
cc_default)) (Tcons tdouble Tnil) tdouble cc_default)) ::
(___builtin_memcpy_aligned,
Gfun(External (EF_builtin "__builtin_memcpy_aligned"
(mksignature
(AST.Tint :: AST.Tint :: AST.Tint :: AST.Tint :: nil)
None cc_default))
(Tcons (tptr tvoid)
(Tcons (tptr tvoid) (Tcons tuint (Tcons tuint Tnil)))) tvoid
cc_default)) ::
(___builtin_annot,
Gfun(External (EF_builtin "__builtin_annot"
(mksignature (AST.Tint :: nil) None
{|cc_vararg:=true; cc_unproto:=false; cc_structret:=false|}))
(Tcons (tptr tschar) Tnil) tvoid
{|cc_vararg:=true; cc_unproto:=false; cc_structret:=false|})) ::
(___builtin_annot_intval,
Gfun(External (EF_builtin "__builtin_annot_intval"
(mksignature (AST.Tint :: AST.Tint :: nil) (Some AST.Tint)
cc_default)) (Tcons (tptr tschar) (Tcons tint Tnil))
tint cc_default)) ::
(___builtin_membar,
Gfun(External (EF_builtin "__builtin_membar"
(mksignature nil None cc_default)) Tnil tvoid cc_default)) ::
(___builtin_va_start,
Gfun(External (EF_builtin "__builtin_va_start"
(mksignature (AST.Tint :: nil) None cc_default))
(Tcons (tptr tvoid) Tnil) tvoid cc_default)) ::
(___builtin_va_arg,
Gfun(External (EF_builtin "__builtin_va_arg"
(mksignature (AST.Tint :: AST.Tint :: nil) None
cc_default)) (Tcons (tptr tvoid) (Tcons tuint Tnil))
tvoid cc_default)) ::
(___builtin_va_copy,
Gfun(External (EF_builtin "__builtin_va_copy"
(mksignature (AST.Tint :: AST.Tint :: nil) None
cc_default))
(Tcons (tptr tvoid) (Tcons (tptr tvoid) Tnil)) tvoid cc_default)) ::
(___builtin_va_end,
Gfun(External (EF_builtin "__builtin_va_end"
(mksignature (AST.Tint :: nil) None cc_default))
(Tcons (tptr tvoid) Tnil) tvoid cc_default)) ::
(___compcert_va_int32,
Gfun(External (EF_external "__compcert_va_int32"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons (tptr tvoid) Tnil) tuint cc_default)) ::
(___compcert_va_int64,
Gfun(External (EF_external "__compcert_va_int64"
(mksignature (AST.Tint :: nil) (Some AST.Tlong)
cc_default)) (Tcons (tptr tvoid) Tnil) tulong
cc_default)) ::
(___compcert_va_float64,
Gfun(External (EF_external "__compcert_va_float64"
(mksignature (AST.Tint :: nil) (Some AST.Tfloat)
cc_default)) (Tcons (tptr tvoid) Tnil) tdouble
cc_default)) ::
(___compcert_va_composite,
Gfun(External (EF_external "__compcert_va_composite"
(mksignature (AST.Tint :: AST.Tint :: nil) (Some AST.Tint)
cc_default)) (Tcons (tptr tvoid) (Tcons tuint Tnil))
(tptr tvoid) cc_default)) ::
(___i64_dtos,
Gfun(External (EF_runtime "__i64_dtos"
(mksignature (AST.Tfloat :: nil) (Some AST.Tlong)
cc_default)) (Tcons tdouble Tnil) tlong cc_default)) ::
(___i64_dtou,
Gfun(External (EF_runtime "__i64_dtou"
(mksignature (AST.Tfloat :: nil) (Some AST.Tlong)
cc_default)) (Tcons tdouble Tnil) tulong cc_default)) ::
(___i64_stod,
Gfun(External (EF_runtime "__i64_stod"
(mksignature (AST.Tlong :: nil) (Some AST.Tfloat)
cc_default)) (Tcons tlong Tnil) tdouble cc_default)) ::
(___i64_utod,
Gfun(External (EF_runtime "__i64_utod"
(mksignature (AST.Tlong :: nil) (Some AST.Tfloat)
cc_default)) (Tcons tulong Tnil) tdouble cc_default)) ::
(___i64_stof,
Gfun(External (EF_runtime "__i64_stof"
(mksignature (AST.Tlong :: nil) (Some AST.Tsingle)
cc_default)) (Tcons tlong Tnil) tfloat cc_default)) ::
(___i64_utof,
Gfun(External (EF_runtime "__i64_utof"
(mksignature (AST.Tlong :: nil) (Some AST.Tsingle)
cc_default)) (Tcons tulong Tnil) tfloat cc_default)) ::
(___i64_sdiv,
Gfun(External (EF_runtime "__i64_sdiv"
(mksignature (AST.Tlong :: AST.Tlong :: nil)
(Some AST.Tlong) cc_default))
(Tcons tlong (Tcons tlong Tnil)) tlong cc_default)) ::
(___i64_udiv,
Gfun(External (EF_runtime "__i64_udiv"
(mksignature (AST.Tlong :: AST.Tlong :: nil)
(Some AST.Tlong) cc_default))
(Tcons tulong (Tcons tulong Tnil)) tulong cc_default)) ::
(___i64_smod,
Gfun(External (EF_runtime "__i64_smod"
(mksignature (AST.Tlong :: AST.Tlong :: nil)
(Some AST.Tlong) cc_default))
(Tcons tlong (Tcons tlong Tnil)) tlong cc_default)) ::
(___i64_umod,
Gfun(External (EF_runtime "__i64_umod"
(mksignature (AST.Tlong :: AST.Tlong :: nil)
(Some AST.Tlong) cc_default))
(Tcons tulong (Tcons tulong Tnil)) tulong cc_default)) ::
(___i64_shl,
Gfun(External (EF_runtime "__i64_shl"
(mksignature (AST.Tlong :: AST.Tint :: nil)
(Some AST.Tlong) cc_default))
(Tcons tlong (Tcons tint Tnil)) tlong cc_default)) ::
(___i64_shr,
Gfun(External (EF_runtime "__i64_shr"
(mksignature (AST.Tlong :: AST.Tint :: nil)
(Some AST.Tlong) cc_default))
(Tcons tulong (Tcons tint Tnil)) tulong cc_default)) ::
(___i64_sar,
Gfun(External (EF_runtime "__i64_sar"
(mksignature (AST.Tlong :: AST.Tint :: nil)
(Some AST.Tlong) cc_default))
(Tcons tlong (Tcons tint Tnil)) tlong cc_default)) ::
(___builtin_bswap,
Gfun(External (EF_builtin "__builtin_bswap"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tuint cc_default)) ::
(___builtin_bswap32,
Gfun(External (EF_builtin "__builtin_bswap32"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tuint cc_default)) ::
(___builtin_bswap16,
Gfun(External (EF_builtin "__builtin_bswap16"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tushort Tnil) tushort cc_default)) ::
(___builtin_clz,
Gfun(External (EF_builtin "__builtin_clz"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tint cc_default)) ::
(___builtin_clzl,
Gfun(External (EF_builtin "__builtin_clzl"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tint cc_default)) ::
(___builtin_clzll,
Gfun(External (EF_builtin "__builtin_clzll"
(mksignature (AST.Tlong :: nil) (Some AST.Tint)
cc_default)) (Tcons tulong Tnil) tint cc_default)) ::
(___builtin_ctz,
Gfun(External (EF_builtin "__builtin_ctz"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tint cc_default)) ::
(___builtin_ctzl,
Gfun(External (EF_builtin "__builtin_ctzl"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons tuint Tnil) tint cc_default)) ::
(___builtin_ctzll,
Gfun(External (EF_builtin "__builtin_ctzll"
(mksignature (AST.Tlong :: nil) (Some AST.Tint)
cc_default)) (Tcons tulong Tnil) tint cc_default)) ::
(___builtin_fsqrt,
Gfun(External (EF_builtin "__builtin_fsqrt"
(mksignature (AST.Tfloat :: nil) (Some AST.Tfloat)
cc_default)) (Tcons tdouble Tnil) tdouble cc_default)) ::
(___builtin_fmax,
Gfun(External (EF_builtin "__builtin_fmax"
(mksignature (AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble Tnil)) tdouble cc_default)) ::
(___builtin_fmin,
Gfun(External (EF_builtin "__builtin_fmin"
(mksignature (AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble Tnil)) tdouble cc_default)) ::
(___builtin_fmadd,
Gfun(External (EF_builtin "__builtin_fmadd"
(mksignature
(AST.Tfloat :: AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble (Tcons tdouble Tnil))) tdouble
cc_default)) ::
(___builtin_fmsub,
Gfun(External (EF_builtin "__builtin_fmsub"
(mksignature
(AST.Tfloat :: AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble (Tcons tdouble Tnil))) tdouble
cc_default)) ::
(___builtin_fnmadd,
Gfun(External (EF_builtin "__builtin_fnmadd"
(mksignature
(AST.Tfloat :: AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble (Tcons tdouble Tnil))) tdouble
cc_default)) ::
(___builtin_fnmsub,
Gfun(External (EF_builtin "__builtin_fnmsub"
(mksignature
(AST.Tfloat :: AST.Tfloat :: AST.Tfloat :: nil)
(Some AST.Tfloat) cc_default))
(Tcons tdouble (Tcons tdouble (Tcons tdouble Tnil))) tdouble
cc_default)) ::
(___builtin_read16_reversed,
Gfun(External (EF_builtin "__builtin_read16_reversed"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons (tptr tushort) Tnil) tushort cc_default)) ::
(___builtin_read32_reversed,
Gfun(External (EF_builtin "__builtin_read32_reversed"
(mksignature (AST.Tint :: nil) (Some AST.Tint) cc_default))
(Tcons (tptr tuint) Tnil) tuint cc_default)) ::
(___builtin_write16_reversed,
Gfun(External (EF_builtin "__builtin_write16_reversed"
(mksignature (AST.Tint :: AST.Tint :: nil) None
cc_default)) (Tcons (tptr tushort) (Tcons tushort Tnil))
tvoid cc_default)) ::
(___builtin_write32_reversed,
Gfun(External (EF_builtin "__builtin_write32_reversed"
(mksignature (AST.Tint :: AST.Tint :: nil) None
cc_default)) (Tcons (tptr tuint) (Tcons tuint Tnil))
tvoid cc_default)) ::
(___builtin_nop,
Gfun(External (EF_builtin "__builtin_nop"
(mksignature nil None cc_default)) Tnil tvoid cc_default)) ::
(___builtin_debug,
Gfun(External (EF_external "__builtin_debug"
(mksignature (AST.Tint :: nil) None
{|cc_vararg:=true; cc_unproto:=false; cc_structret:=false|}))
(Tcons tint Tnil) tvoid
{|cc_vararg:=true; cc_unproto:=false; cc_structret:=false|})) ::
(_set, Gfun(Internal f_set)) :: (_call_set, Gfun(Internal f_call_set)) ::
(_main, Gfun(Internal f_main)) :: nil);
prog_public :=
(_main :: ___builtin_debug :: ___builtin_nop ::
___builtin_write32_reversed :: ___builtin_write16_reversed ::
___builtin_read32_reversed :: ___builtin_read16_reversed ::
___builtin_fnmsub :: ___builtin_fnmadd :: ___builtin_fmsub ::
___builtin_fmadd :: ___builtin_fmin :: ___builtin_fmax ::
___builtin_fsqrt :: ___builtin_ctzll :: ___builtin_ctzl :: ___builtin_ctz ::
___builtin_clzll :: ___builtin_clzl :: ___builtin_clz ::
___builtin_bswap16 :: ___builtin_bswap32 :: ___builtin_bswap ::
___i64_sar :: ___i64_shr :: ___i64_shl :: ___i64_umod :: ___i64_smod ::
___i64_udiv :: ___i64_sdiv :: ___i64_utof :: ___i64_stof :: ___i64_utod ::
___i64_stod :: ___i64_dtou :: ___i64_dtos :: ___compcert_va_composite ::
___compcert_va_float64 :: ___compcert_va_int64 :: ___compcert_va_int32 ::
___builtin_va_end :: ___builtin_va_copy :: ___builtin_va_arg ::
___builtin_va_start :: ___builtin_membar :: ___builtin_annot_intval ::
___builtin_annot :: ___builtin_memcpy_aligned :: ___builtin_fabs :: nil);
prog_main := _main;
prog_types := composites;
prog_comp_env := make_composite_env composites;
prog_comp_env_eq := refl_equal _
|}.
Instance CompSpecs : compspecs. make_compspecs prog. Defined.
Definition Vprog : varspecs. mk_varspecs prog. Defined.
Open Scope Z.
Notation tlg := (Tlong Signed {| attr_volatile := false; attr_alignas := Some 3%N |}).
Definition set_spec :=
DECLARE _set
WITH r: val, a: val, sh : share, tsh:share, contents_r : list val, contents_a : list Z
PRE [ _r OF (tptr tlg), _a OF (tptr tlg)]
PROP (readable_share sh; writable_share tsh;
Zlength contents_a = 16;
Zlength contents_r = 16)
LOCAL (temp _r r; temp _a a)
SEP (data_at sh (tarray tlg 16) (map Vlong (map Int64.repr contents_a)) a;
data_at tsh (tarray tlg 16) contents_r r)
POST [ tvoid ]
PROP (readable_share sh; writable_share tsh;
Zlength contents_a = 16;
Zlength contents_r = 16
)
LOCAL()
SEP (data_at sh (tarray tlg 16) (map Vlong (map Int64.repr contents_a)) a;
data_at tsh (tarray tlg 16) (map Vlong (map Int64.repr contents_a)) r).
Definition Gprog : funspecs :=
ltac:(with_library prog [set_spec]).
Lemma body_set: semax_body Vprog Gprog f_set set_spec.
Proof.
start_function.
(* just to "prove" that it actually matches the signature. *)
(* the proof is in my repo and is not interesting here. *)
Admitted.
Definition call_set_spec :=
DECLARE _call_set
WITH a: val, sh : share, contents_a : list Z
PRE [ _a OF (tptr tlg)]
PROP (writable_share sh;
Zlength contents_a = 16)
LOCAL (temp _a a)
SEP (data_at sh (tarray tlg 16) (map Vlong (map Int64.repr contents_a)) a)
POST [ tvoid ]
PROP ()
LOCAL()
SEP (data_at sh (tarray tlg 16) (map Vlong (map Int64.repr contents_a)) a).
Definition Hprog : funspecs :=
ltac:(with_library prog [set_spec; call_set_spec]).
Lemma body_call_set: semax_body Vprog Hprog f_call_set call_set_spec.
Proof.
start_function.
forward.
(* now we need to do the call...
we need to know the content of r.
and for that... we have to unfold unfold unfold... *)
unfold data_at_.
unfold field_at_.
unfold default_val.
unfold reptype_gen.
simpl.
(* these 5 previous steps should be automatics !!! *)
assert(Hundef16 : exists undef16, undef16 = (list_repeat (Z.to_nat 16) Vundef)) by (eexists ; reflexivity).
destruct Hundef16 as [undef16 Hundef16].
replace [Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef; Vundef;
Vundef; Vundef; Vundef] with undef16.
replace (field_at Tsh (tarray tlg 16) [] undef16 v_r) with (data_at Tsh (tarray tlg 16) undef16 v_r)
by reflexivity.
forward_call (v_r, a, sh, Tsh, undef16 , contents_a).
(* nothing interesting afterwards... *)
repeat split ; auto.
rewrite Hundef16.
rewrite Zlength_list_repeat ; omega.
forward.
cancel.
Qed.
@lennartberinger
Copy link

There's a useful lemma in progs.conclib: data_at__eq.
(yes, we should move it to floyd and maybe add it to the automation).

But the spec doesn't seem to hold....

Require Import VST.progs.conclib.

Lemma body_call_set: semax_body Vprog Hprog f_call_set call_set_spec.
Proof.
start_function.
forward.
forward_call (lvar0, a, sh, Tsh, list_repeat (Z.to_nat 16) Vundef, contents_a).
{ rewrite data_at__eq. cancel. }
forward. cancel. Admitted.

Sign up for free to join this conversation on GitHub. Already have an account? Sign in to comment