Skip to content

Commit 2f8617c

Browse files
committed
- Proof of cap_encode_valid
- Proof of bv_to_bits_app_last, and other auxiliary lemmas - Simplified proof of cap_encode_decode
1 parent da69ac4 commit 2f8617c

File tree

1 file changed

+67
-50
lines changed

1 file changed

+67
-50
lines changed

theories/Morello/Capabilities.v

Lines changed: 67 additions & 50 deletions
Original file line numberDiff line numberDiff line change
@@ -40,7 +40,7 @@ Local Notation "x >=? y" := (geb x y) (at level 70, no associativity) : bv_scope
4040
Local Notation "(<@{ A } )" := (@lt A) (only parsing) : stdpp_scope.
4141
Local Notation LtDecision A := (RelDecision (<@{A})).
4242

43-
(** Utility converters **)
43+
(** Utility converters and lemmas **)
4444

4545
Definition bv_to_Z_unsigned {n} (v : bv n) : Z := v.(bv_unsigned).
4646
Definition bv_to_N {n} (v : bv n) : N := Z.to_N v.(bv_unsigned).
@@ -67,6 +67,22 @@ Definition invert_bits {n} (m : mword n) : (mword n) :=
6767
let x : Z := int_of_mword false x in
6868
mword_of_int x.
6969

70+
Lemma bv_eqb_eq {n} : forall (a b: bv n), (a =? b) = true <-> a = b.
71+
Proof.
72+
split.
73+
- intro H. unfold Capabilities.eqb in H.
74+
apply Z.eqb_eq in H. apply bv_eq in H. exact H.
75+
- intro H. rewrite H. unfold Capabilities.eqb. lia.
76+
Defined.
77+
78+
Lemma bv1_neqb_eq0 : forall (a : bv 1), (a =? 1%bv) = false <-> a = 0%bv.
79+
Proof.
80+
intros. split.
81+
+ intro H. unfold eqb in H. apply Z.eqb_neq in H. apply bv_eq.
82+
assert (P: (0 ≤ bv_unsigned a < 2)%Z); [ apply (bv_unsigned_in_range 1 a) |].
83+
change (bv_unsigned 0) with 0%Z. change (bv_unsigned 1) with 1%Z in H. lia.
84+
+ intro H. rewrite H. done.
85+
Qed.
7086

7187
Module Permissions <: PERMISSIONS.
7288
Definition len:N := 18. (* CAP_PERMS_NUM_BITS = 16 bits of actual perms + 2 bits for Executive and Global *)
@@ -1135,12 +1151,7 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
11351151
Qed.
11361152

11371153
Lemma eqb_eq : forall (a b:t), (a =? b) = true <-> a = b.
1138-
Proof.
1139-
split.
1140-
- intro H. unfold Capabilities.eqb in H.
1141-
apply Z.eqb_eq in H. apply bv_eq in H. exact H.
1142-
- intro H. rewrite H. apply eqb_refl.
1143-
Defined.
1154+
Proof. apply bv_eqb_eq. Defined.
11441155

11451156
(* Beginning of Translation definitions and lemmas from Capabilities to bitvectors. *)
11461157

@@ -1477,7 +1488,11 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
14771488
Lemma cap_encode_valid:
14781489
forall cap cb b,
14791490
cap_is_valid cap = true -> encode cap = Some (cb, b) -> b = true.
1480-
Admitted.
1491+
Proof.
1492+
intros cap bytes tag H_valid H_encode.
1493+
unfold encode in H_encode. repeat case_match; try done.
1494+
inversion H_encode. exact H_valid.
1495+
Qed.
14811496

14821497
Lemma cap_is_valid_bv_extract c :
14831498
cap_is_valid c ↔ bv_extract 128 1 c = (BV 1 1).
@@ -1511,39 +1526,11 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
15111526
symmetry. apply Z.eqb_neq. by apply bv_neq.
15121527
Qed.
15131528

1514-
Lemma bits_of_vec_hd h l :
1515-
bits_of (vec_of_bits (h::l)) = bits_of ((concat_vec (vec_of_bits [h]) (vec_of_bits l))).
1529+
Lemma mword_to_bits_to_mword {n} (c : mword (n)) :
1530+
vec_of_bits (bits_of c) = autocast c.
15161531
Proof.
15171532
Admitted.
15181533

1519-
Lemma hd_bits_of_vec h l :
1520-
h ≠ BU →
1521-
bits_of (concat_vec (vec_of_bits [h]) (vec_of_bits l)) = h :: bits_of (vec_of_bits l).
1522-
Proof.
1523-
Admitted.
1524-
1525-
Lemma bits_to_vec_to_bits l :
1526-
bits_of (vec_of_bits l) = l.
1527-
Proof.
1528-
assert (H: ∀ n l, length l = n → bits_of (vec_of_bits l) = l).
1529-
{ intro n. induction n as [ | m IH ].
1530-
+ intros l' H. apply nil_length_inv in H. subst. by vm_compute.
1531-
+ intros l' H. destruct l' as [| h tl]; try discriminate.
1532-
rewrite cons_length in H. inversion H as [H1].
1533-
specialize IH with (l:=tl). apply IH in H1 as H2.
1534-
1535-
rewrite bits_of_vec_hd.
1536-
rewrite hd_bits_of_vec; [| by rewrite H2 ].
1537-
admit.
1538-
}
1539-
by apply (H (length l)).
1540-
Admitted.
1541-
1542-
Lemma cap_to_bits_to_cap (c : mword 129) :
1543-
vec_of_bits (bits_of c) = c.
1544-
Proof.
1545-
Admitted.
1546-
15471534
Lemma bytes_to_bits_to_bytes l l' l'' :
15481535
byte_chunks l = Some l' →
15491536
try_map memory_byte_to_ascii (rev l') = Some l'' →
@@ -1553,7 +1540,30 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
15531540

15541541
Lemma bv_to_bits_app_last (c : bv 129) :
15551542
bv_to_bits c = (bv_to_bits (bv_extract 0 128 c) ++ [bv_extract 128 1 c =? 1])%list.
1556-
Admitted.
1543+
Proof.
1544+
have -> : bv_to_bits c = bv_to_bits (@bv_concat 129 1 128 (@bv_extract 129 128 1 c) (@bv_extract 129 0 128 c)).
1545+
{ apply f_equal. bv_simplify. bitblast. }
1546+
set (l1 := bv_to_bits (bv_concat 129 (bv_extract 128 1 c) (bv_extract 0 128 c))).
1547+
set (l2 := (bv_to_bits (bv_extract 0 128 c) ++ [bv_extract 128 1 c =? 1])%list).
1548+
apply list_eq. intro i.
1549+
destruct (128 <? i)%nat eqn:Hi1.
1550+
{ apply Nat.ltb_lt in Hi1. rewrite !lookup_ge_None_2; try done. }
1551+
destruct (i =? 128)%nat eqn:Hi2.
1552+
- apply Nat.eqb_eq in Hi2. rewrite Hi2. unfold l1, l2.
1553+
rewrite lookup_app_r bv_to_bits_length; try done; simpl.
1554+
apply bv_to_bits_lookup_Some.
1555+
split; try lia.
1556+
destruct (_ =? _) eqn:P; [ apply bv_eqb_eq in P | apply bv1_neqb_eq0 in P ];
1557+
rewrite P; bv_simplify; bitblast.
1558+
- assert (Z.of_nat i < 128)%Z.
1559+
{ apply Nat.eqb_neq in Hi2. apply Nat.ltb_ge in Hi1. lia. }
1560+
unfold l1, l2. rewrite lookup_app_l; [ rewrite bv_to_bits_length; try lia |].
1561+
destruct (bv_to_bits (bv_extract 0 128 c) !! i) eqn:P.
1562+
+ apply bv_to_bits_lookup_Some. split; try lia.
1563+
apply bv_to_bits_lookup_Some in P. destruct P as [_ P]. rewrite P.
1564+
bv_simplify. bitblast.
1565+
+ rewrite list_lookup_lookup_total_lt in P; [ rewrite bv_to_bits_length; try lia | discriminate ].
1566+
Qed.
15571567

15581568
Lemma mword_to_bools_cons_last (c: bv 129) :
15591569
@mword_to_bools 129 c = (((bv_extract 128 1 c =? 1)) :: (@mword_to_bools 128 (bv_extract 0 128 c)))%list.
@@ -1572,13 +1582,21 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
15721582
inversion E; clear E; subst.
15731583
unfold mem_bytes_of_bits in E1. unfold bytes_of_bits in E1. unfold option_map in E1. case_match eqn:E3; try done.
15741584
inversion E1; clear E1; subst.
1575-
apply (bytes_to_bits_to_bytes _ l0 bytes) in E3; try assumption.
1576-
rewrite bits_to_vec_to_bits in E3.
1577-
1578-
unfold t, len in *.
1585+
apply (bytes_to_bits_to_bytes _ l0 bytes) in E3; try assumption.
1586+
unfold t, len in *.
1587+
change (Z.of_N 129) with 129%Z in *. Set Printing Implicit.
1588+
replace (list.tail (@bits_of 129 c)) with (@bits_of 128 (bv_extract 0 128 c)) in E3. 2:{ unfold bits_of. rewrite mword_to_bools_cons_last. simpl. reflexivity. }
1589+
rewrite mword_to_bits_to_mword in E3.
1590+
15791591
unfold decode. rewrite E_len; simpl. rewrite E3.
15801592
unfold uint, MachineWord.word_to_N, len; rewrite Z2N.id; [ apply bv_unsigned_in_range |];
1581-
simpl. set (vec := vec_of_bits (bitU_of_bool (cap_is_valid c) :: list.tail (@bits_of 129 c))).
1593+
simpl. Set Printing Implicit. set (vec := vec_of_bits
1594+
(bitU_of_bool (cap_is_valid c)
1595+
:: @bits_of (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c)))
1596+
(@autocast mword 128
1597+
(@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c)))
1598+
(@dummy_mword (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c))))
1599+
(@bv_extract 129 0 128 c)))).
15821600

15831601
have -> : bv_unsigned vec = bv_unsigned (@bv_to_bv _ 129 vec).
15841602
{ unfold bv_to_bv, bv_to_Z_unsigned. by rewrite Z_to_bv_bv_unsigned. }
@@ -1587,20 +1605,19 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
15871605
rewrite cap_is_valid_bv_extract_bool.
15881606
destruct (bv_extract 128 1 c =? 1) eqn:H_c_valid; simpl;
15891607
[ replace B1 with (bitU_of_bool (bv_extract 128 1 c =? 1)); [| by rewrite H_c_valid ]
1590-
| replace B0 with (bitU_of_bool (bv_extract 128 1 c =? 1)); [| by rewrite H_c_valid ]].
1591-
all: unfold bits_of;
1592-
have -> : @mword_to_bools 129 c = (bv_extract 128 1 c =? 1) :: @mword_to_bools 128 (bv_extract 0 128 c); [ by rewrite mword_to_bools_cons_last |];
1593-
rewrite map_cons; simpl;
1608+
| replace B0 with (bitU_of_bool (bv_extract 128 1 c =? 1)); [| by rewrite H_c_valid ]].
1609+
all: unfold bits_of;
1610+
change (autocast (bv_extract 0 128 c)) with ((bv_extract 0 128 c));
15941611
have -> : vec_of_bits (bitU_of_bool (bv_extract 128 1 c =? 1) :: map bitU_of_bool (@mword_to_bools 128 (bv_extract 0 128 c))) = (vec_of_bits (map bitU_of_bool ((bv_extract 128 1 c =? 1) :: @mword_to_bools 128 (bv_extract 0 128 c))));
15951612
[ unfold vec_of_bits; apply f_equal; by rewrite map_cons
15961613
| have -> : bv_unsigned (vec_of_bits (map bitU_of_bool ((@bv_extract 129 128 1 c =? 1) :: @mword_to_bools 128 (@bv_extract 129 0 128 c)))) = bv_unsigned (vec_of_bits (map bitU_of_bool (@mword_to_bools 129 c)));
15971614
[ change (N.of_nat (Pos.to_nat (Pos.of_succ_nat _ ))) with 129%N;
15981615
apply f_equal; unfold vec_of_bits; apply f_equal; by rewrite <- mword_to_bools_cons_last
1599-
| fold (@bits_of 129 c); rewrite cap_to_bits_to_cap; by rewrite Z_to_bv_bv_unsigned ]
1616+
| fold (@bits_of 129 c); rewrite mword_to_bits_to_mword; by rewrite Z_to_bv_bv_unsigned ]
16001617
].
16011618
}
16021619
Qed.
1603-
1620+
16041621
Lemma cap_encode_decode_bounds:
16051622
∀ cap bytes t,
16061623
encode cap = Some (bytes, t) →

0 commit comments

Comments
 (0)