Skip to content

Commit b570230

Browse files
committed
Several proofs of round-trip conversions using coq-sail functions
1 parent ba2c092 commit b570230

File tree

1 file changed

+110
-18
lines changed

1 file changed

+110
-18
lines changed

theories/Morello/Capabilities.v

Lines changed: 110 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -873,7 +873,7 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
873873

874874
Definition decode (bytes : list ascii) (tag : bool) : option t :=
875875
if bool_decide (List.length bytes = 16%nat) then
876-
let bytes := List.rev bytes in (* TODO: bypassing rev could make this more efficient *)
876+
let bytes := List.rev bytes in
877877
let bits : (list bool) := tag::(bool_bits_of_bytes bytes) in
878878
let bitsu := List.map bitU_of_bool bits in
879879
let w := vec_of_bits bitsu in
@@ -1296,13 +1296,18 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
12961296
rewrite Hdiv. rewrite Nat.div_mul; lia.
12971297
Qed.
12981298

1299+
Lemma mword_to_bools_len {n} (w : mword n) :
1300+
length (mword_to_bools w) = Z.to_nat n.
1301+
Proof.
1302+
unfold mword_to_bools.
1303+
by rewrite MachineWord.word_to_bools_length.
1304+
Qed.
1305+
12991306
Lemma bits_of_len {n} (w : mword n) :
13001307
length (bits_of w) = Z.to_nat n.
13011308
Proof.
13021309
unfold bits_of. rewrite map_length.
1303-
unfold mword_to_bools.
1304-
rewrite MachineWord.word_to_bools_length.
1305-
reflexivity.
1310+
by rewrite mword_to_bools_len.
13061311
Qed.
13071312

13081313
Lemma cap_encode_length:
@@ -1382,11 +1387,103 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
13821387
{ intro H. apply cap_is_valid_bv_extract in H. rewrite E in H. contradiction. }
13831388
symmetry. apply Z.eqb_neq. by apply bv_neq.
13841389
Qed.
1390+
1391+
Lemma bool_to_bit_to_bool b :
1392+
bool_of_bit (bitU_of_bool b) = b.
1393+
Proof.
1394+
unfold bitU_of_bool, bool_of_bit.
1395+
repeat case_match; try done.
1396+
Qed.
13851397

1386-
Lemma mword_to_bits_to_mword {n} (c : mword (n)) :
1398+
Lemma bv_to_bits_to_Z {n} (c : bv n) :
1399+
(n>0)%N →
1400+
Z.of_N (N_of_digits (bv_to_bits c)) = bv_unsigned c.
1401+
Proof.
1402+
intro H_n_min.
1403+
assert (H: length (bv_to_bits c) = N.to_nat n); [ by rewrite bv_to_bits_length|].
1404+
assert (G: (N_of_digits (bv_to_bits c) < 2 ^ (N.of_nat (N.to_nat n)))%N).
1405+
{ rewrite <- (bv_to_bits_length (c)). apply MachineWord.N_of_digits_limit. }
1406+
bitblast as m.
1407+
- set (i:=Z.to_N m).
1408+
have ->: m = Z.of_N i; [ lia |].
1409+
rewrite Z.testbit_of_N.
1410+
apply eq_bool_prop_intro.
1411+
destruct (i <? n)%N eqn:Hbound.
1412+
+ apply N.ltb_lt in Hbound.
1413+
set (i' := N.to_nat i).
1414+
assert (Q: i = N.of_nat i'); [ lia |].
1415+
assert (R: (i' < N.to_nat n)%nat); [ lia |].
1416+
split.
1417+
* intro P. rewrite Q in P.
1418+
apply Is_true_eq_true in P.
1419+
specialize (conj P R) as H'.
1420+
rewrite <- H in H'.
1421+
apply MachineWord.MachineWord.nth_error_N_of_digits in H'.
1422+
rewrite <- MachineWord.nth_error_lookup in H'.
1423+
apply bv_to_bits_lookup_Some in H'.
1424+
destruct H' as [H1 H2]. unfold i' in H2. rewrite N_nat_Z in H2.
1425+
by rewrite <- H2.
1426+
* intro P. rewrite Q in P. rewrite nat_N_Z in P.
1427+
apply Is_true_eq_true in P; symmetry in P.
1428+
specialize (conj R P) as H'.
1429+
apply bv_to_bits_lookup_Some in H'.
1430+
rewrite MachineWord.nth_error_lookup in H'.
1431+
apply MachineWord.MachineWord.nth_error_N_of_digits in H'.
1432+
destruct H' as [H1 H2]. unfold i' in H1. rewrite N2Nat.id in H1.
1433+
by rewrite H1.
1434+
+ apply N.ltb_ge in Hbound. rewrite bv_unsigned_spec_high; [ lia |].
1435+
rewrite N.bits_above_log2; try done.
1436+
assert ((N_of_digits (bv_to_bits c) < 2 ^ (N.of_nat (length (bv_to_bits c))))%N); [ apply MachineWord.N_of_digits_limit | lia ].
1437+
- assert (m >= Z.of_N n)%Z; [ lia |].
1438+
apply Z.bits_above_log2; try lia.
1439+
assert (H': N_of_digits (bv_to_bits c) = 0%N ∨ (N_of_digits (bv_to_bits c) > 0)%N); [ lia |].
1440+
destruct H' as [H' | H'].
1441+
+ rewrite H'. simpl. lia.
1442+
+ apply Z.log2_lt_pow2; try lia.
1443+
assert (H'': (N_of_digits (bv_to_bits c) < 2 ^ Z.to_N m)%N).
1444+
{ assert (2 ^ n <= 2 ^ Z.to_N m)%N;
1445+
[ apply N.pow_le_mono_r; try lia | lia ]. }
1446+
apply N2Z.inj_lt in H''.
1447+
replace (Z.of_N (2 ^ Z.to_N m)) with (2 ^ m)%Z in H''.
1448+
{ lia. } { rewrite N2Z.inj_pow. simpl. rewrite Z2N.id; lia. }
1449+
Qed.
1450+
1451+
Lemma mword_of_bits_to_mword_of_bools {n} (c : mword n) :
1452+
(n >= 0)%Z →
1453+
vec_of_bits (bits_of c) = of_bools (mword_to_bools c).
1454+
Proof.
1455+
intro Hbound.
1456+
unfold bits_of, vec_of_bits.
1457+
rewrite -> map_map.
1458+
set (f := (λ x : bool, bool_of_bit (bitU_of_bool x))).
1459+
set (g := λ x : bool, x).
1460+
rewrite (map_ext f g); [ apply bool_to_bit_to_bool |].
1461+
by rewrite map_id.
1462+
Qed.
1463+
1464+
Lemma mword128_to_bits_to_mword (c : mword 128) :
13871465
vec_of_bits (bits_of c) = autocast c.
13881466
Proof.
1389-
Admitted.
1467+
rewrite mword_of_bits_to_mword_of_bools; try lia.
1468+
unfold mword_to_bools, of_bools. simpl.
1469+
unfold bools_to_mword, MachineWord.word_to_bools, MachineWord.bools_to_word.
1470+
rewrite rev_involutive /MachineWord.N_to_word /to_word_nat.
1471+
rewrite cast_nat_refl /to_word; simpl.
1472+
rewrite bv_to_bits_to_Z; try lia. rewrite Z_to_bv_bv_unsigned.
1473+
by rewrite autocast_refl.
1474+
Qed.
1475+
1476+
Lemma mword129_to_bits_to_mword (c : mword 129) :
1477+
vec_of_bits (bits_of c) = autocast c.
1478+
Proof.
1479+
rewrite mword_of_bits_to_mword_of_bools; try lia.
1480+
unfold mword_to_bools, of_bools. simpl.
1481+
unfold bools_to_mword, MachineWord.word_to_bools, MachineWord.bools_to_word.
1482+
rewrite rev_involutive /MachineWord.N_to_word /to_word_nat.
1483+
rewrite cast_nat_refl /to_word; simpl.
1484+
rewrite bv_to_bits_to_Z; try lia. rewrite Z_to_bv_bv_unsigned.
1485+
by rewrite autocast_refl.
1486+
Qed.
13901487

13911488
Lemma bytes_to_bits_to_bytes l l' l'' :
13921489
byte_chunks l = Some l' →
@@ -1441,19 +1538,15 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
14411538
inversion E1; clear E1; subst.
14421539
apply (bytes_to_bits_to_bytes _ l0 bytes) in E3; try assumption.
14431540
unfold t, len in *.
1444-
change (Z.of_N 129) with 129%Z in *. Set Printing Implicit.
1445-
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. }
1446-
rewrite mword_to_bits_to_mword in E3.
1541+
change (Z.of_N 129) with 129%Z in *.
1542+
replace (list.tail (@bits_of 129 c)) with (@bits_of 128 (bv_extract 0 128 c)) in E3;
1543+
[| unfold bits_of; rewrite mword_to_bools_cons_last; reflexivity ].
1544+
rewrite mword128_to_bits_to_mword in E3; try lia.
14471545

14481546
unfold decode. rewrite E_len; simpl. rewrite E3.
14491547
unfold uint, MachineWord.word_to_N, len; rewrite Z2N.id; [ apply bv_unsigned_in_range |];
1450-
simpl. Set Printing Implicit. set (vec := vec_of_bits
1451-
(bitU_of_bool (cap_is_valid c)
1452-
:: @bits_of (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c)))
1453-
(@autocast mword 128
1454-
(@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c)))
1455-
(@dummy_mword (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c))))
1456-
(@bv_extract 129 0 128 c)))).
1548+
simpl.
1549+
set (vec := vec_of_bits (bitU_of_bool (cap_is_valid c) :: @bits_of (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c))) (@autocast mword 128 (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c))) (@dummy_mword (@length_list bitU (@bits_of 128 (@bv_extract 129 0 128 c)))) (@bv_extract 129 0 128 c)))).
14571550

14581551
have -> : bv_unsigned vec = bv_unsigned (@bv_to_bv _ 129 vec).
14591552
{ unfold bv_to_bv, bv_to_Z_unsigned. by rewrite Z_to_bv_bv_unsigned. }
@@ -1470,8 +1563,7 @@ Module Capability <: CAPABILITY (AddressValue) (Flags) (ObjType) (SealType) (Bou
14701563
| 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)));
14711564
[ change (N.of_nat (Pos.to_nat (Pos.of_succ_nat _ ))) with 129%N;
14721565
apply f_equal; unfold vec_of_bits; apply f_equal; by rewrite <- mword_to_bools_cons_last
1473-
| fold (@bits_of 129 c); rewrite mword_to_bits_to_mword; by rewrite Z_to_bv_bv_unsigned ]
1474-
].
1566+
| rewrite mword129_to_bits_to_mword; try lia; rewrite Z_to_bv_bv_unsigned; by rewrite autocast_refl ]].
14751567
}
14761568
Qed.
14771569

0 commit comments

Comments
 (0)