Skip to content

Commit efaa7ba

Browse files
committed
additional alignment checks in repr
1 parent 6d711f1 commit efaa7ba

File tree

2 files changed

+19
-15
lines changed

2 files changed

+19
-15
lines changed

coq/CheriMemory/CheriMorelloMemory.v

Lines changed: 5 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -750,19 +750,14 @@ Module Type CheriMemoryImpl
750750
let align := IMP.get.(alignof_pointer) in
751751
(AddressValue.to_Z addr) mod (Z.of_nat align) =? 0.
752752

753-
(** Update [capmeta] dictionary for capability [c] stored at [addr].
754-
If address is capability-aligned, then the tag and ghost state
755-
is stored. Otherwise capmeta is left unchanged. *)
753+
(** Update [capmeta] dictionary for capability [c] stored at [addr]. *)
756754
Definition update_capmeta
757755
(c: C.t)
758756
(addr: AddressValue.t)
759757
(capmeta : AMap.M.t (bool*CapGhostState))
760758
: AMap.M.t (bool*CapGhostState)
761759
:=
762-
if is_pointer_algined addr
763-
then AMap.M.add addr (C.cap_is_valid c, C.get_ghost_state c) capmeta
764-
else capmeta.
765-
760+
AMap.M.add addr (C.cap_is_valid c, C.get_ghost_state c) capmeta.
766761

767762
Fixpoint repr
768763
(fuel: nat)
@@ -797,6 +792,7 @@ Module Type CheriMemoryImpl
797792
| CoqIntegerType.Unsigned CoqIntegerType.Intptr_t
798793
=>
799794
'(cb, ct) <- option2serr "int encoding error" (C.encode true c_value) ;;
795+
sassert (is_pointer_algined addr) "unaligned pointer to cap" ;;
800796
let capmeta := update_capmeta c_value addr capmeta in
801797
sassert (AddressValue.to_Z addr + (Z.of_nat (length cb)) <=? AddressValue.ADDR_LIMIT) "object does not fit in address space" ;;
802798
ret (funptrmap, capmeta, List.map (Some) cb)
@@ -816,11 +812,13 @@ Module Type CheriMemoryImpl
816812
fp) =>
817813
let '(funptrmap, c_value) := resolve_function_pointer funptrmap fp in
818814
'(cb, ct) <- option2serr "valid function pointer encoding error" (C.encode true c_value) ;;
815+
sassert (is_pointer_algined addr) "unaligned pointer to cap" ;;
819816
let capmeta := update_capmeta c_value addr capmeta in
820817
sassert (AddressValue.to_Z addr + (Z.of_nat (length cb)) <=? AddressValue.ADDR_LIMIT) "object does not fit in address space" ;;
821818
ret (funptrmap, capmeta, List.map (Some) cb)
822819
| (PVfunction (FP_invalid c_value) | PVconcrete c_value) =>
823820
'(cb, ct) <- option2serr "pointer encoding error" (C.encode true c_value) ;;
821+
sassert (is_pointer_algined addr) "unaligned pointer to cap" ;;
824822
let capmeta := update_capmeta c_value addr capmeta in
825823
sassert (AddressValue.to_Z addr + (Z.of_nat (length cb)) <=? AddressValue.ADDR_LIMIT) "object does not fit in address space" ;;
826824
ret (funptrmap, capmeta, List.map (Some) cb)

coq/Proofs/Revocation.v

Lines changed: 14 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -5826,12 +5826,16 @@ Module CheriMemoryImplWithProofs
58265826
forall (c : Capability_GS.t) (cb : list ascii) (b:bool),
58275827
Capability_GS.encode true c = Some (cb, b) ->
58285828
forall start : AddressValue.t,
5829+
is_pointer_algined start = true ->
58295830
AddressValue.to_Z start + Z.of_nat (Datatypes.length cb) <= AddressValue.ADDR_LIMIT ->
58305831
forall bs : list (option ascii),
58315832
bs = map Some cb ->
58325833
mem_invariant (mem_state_with_funptrmap_bytemap_capmeta (funptrmap s) (AMap.map_add_list_at (bytemap s) bs start) (update_capmeta c start (capmeta s)) s).
58335834
Proof.
5834-
intros s Bdead Bnooverlap Bfit Balign Bnextallocid Blastaddr MIcap c cb ct start RSZ bs Heqbs.
5835+
intros s Bdead Bnooverlap Bfit Balign Bnextallocid Blastaddr MIcap c cb ct E start SA RSZ bs Heqbs.
5836+
5837+
5838+
(* TODO: this proof must be substantionally similar to [mem_state_with_bytes_preserves] *)
58355839
Admitted.
58365840

58375841
Lemma repr_preserves
@@ -5895,13 +5899,13 @@ Module CheriMemoryImplWithProofs
58955899
repeat break_match_hyp; state_inv_steps.
58965900
*
58975901
generalize dependent (Capability_GS.cap_get_value c); clear c.
5898-
intros start RSZ IHfuel.
5902+
intros start R4 RSZ IHfuel.
58995903
bool_to_prop_hyp.
59005904
remember (map Some l) as bs.
59015905
eapply mem_state_with_cap_preserves;eauto.
59025906
*
59035907
generalize dependent (Capability_GS.cap_get_value c); clear c.
5904-
intros start RSZ IHfuel.
5908+
intros start R4 RSZ IHfuel.
59055909
bool_to_prop_hyp.
59065910
remember (map Some l) as bs.
59075911
eapply mem_state_with_cap_preserves;eauto.
@@ -5934,23 +5938,25 @@ Module CheriMemoryImplWithProofs
59345938
state_inv_step_quick.
59355939
break_let.
59365940
state_inv_step_quick.
5937-
apply ret_inr in R5.
5938-
inversion R5; clear R5.
5941+
state_inv_step_quick.
5942+
apply ret_inr in R6.
5943+
inversion R6; clear R6.
59395944
bool_to_prop_hyp.
59405945
eapply mem_state_with_cap_preserves;eauto.
59415946
*
59425947
state_inv_step_quick.
59435948
break_let.
59445949
state_inv_step_quick.
59455950
state_inv_step_quick.
5946-
apply ret_inr in R5.
5947-
inversion R5; clear R5.
5951+
state_inv_step_quick.
5952+
apply ret_inr in R6.
5953+
inversion R6; clear R6.
59485954
bool_to_prop_hyp.
59495955
eapply mem_state_with_cap_preserves;eauto.
59505956
+
59515957
state_inv_steps.
59525958
generalize dependent (Capability_GS.cap_get_value c); clear c.
5953-
intros start RSZ.
5959+
intros start R4 RSZ.
59545960
bool_to_prop_hyp.
59555961
remember (map Some l) as bs.
59565962
eapply mem_state_with_cap_preserves;eauto.

0 commit comments

Comments
 (0)