Skip to content

Commit b705ac5

Browse files
Add support for loop assigns in loop contracts (#4174)
This PR adds support for loop assigns in loop contracts. Resolves #3871 By submitting this pull request, I confirm that my contribution is made under the terms of the Apache 2.0 and MIT licenses.
1 parent ee23e07 commit b705ac5

22 files changed

+443
-21
lines changed

cprover_bindings/src/goto_program/stmt.rs

Lines changed: 27 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -87,6 +87,7 @@ pub enum StmtBody {
8787
// The loop invariants annotated to the goto, which can be
8888
// applied as loop contracts in CBMC if it is a backward goto.
8989
loop_invariants: Option<Expr>,
90+
loop_modifies: Option<Vec<Expr>>,
9091
},
9192
/// `if (i) { t } else { e }`
9293
Ifthenelse {
@@ -288,7 +289,7 @@ impl Stmt {
288289
pub fn goto<T: Into<InternedString>>(dest: T, loc: Location) -> Self {
289290
let dest = dest.into();
290291
assert!(!dest.is_empty());
291-
stmt!(Goto { dest, loop_invariants: None }, loc)
292+
stmt!(Goto { dest, loop_invariants: None, loop_modifies: None }, loc)
292293
}
293294

294295
/// `if (i) { t } else { e }` or `if (i) { t }`
@@ -333,13 +334,36 @@ impl Stmt {
333334

334335
/// `goto dest;` with loop invariant
335336
pub fn with_loop_contracts(self, inv: Expr) -> Self {
336-
if let Goto { dest, loop_invariants } = self.body() {
337+
if let Goto { dest, loop_invariants, loop_modifies } = self.body() {
337338
assert!(loop_invariants.is_none());
338-
stmt!(Goto { dest: *dest, loop_invariants: Some(inv) }, *self.location())
339+
stmt!(
340+
Goto {
341+
dest: *dest,
342+
loop_invariants: Some(inv),
343+
loop_modifies: loop_modifies.clone()
344+
},
345+
*self.location()
346+
)
339347
} else {
340348
unreachable!("Loop contracts should be annotated only to goto stmt")
341349
}
342350
}
351+
352+
pub fn with_loop_modifies(self, asg: Vec<Expr>) -> Self {
353+
if let Goto { dest, loop_invariants, loop_modifies } = self.body() {
354+
assert!(loop_modifies.is_none());
355+
stmt!(
356+
Goto {
357+
dest: *dest,
358+
loop_invariants: loop_invariants.clone(),
359+
loop_modifies: Some(asg)
360+
},
361+
*self.location()
362+
)
363+
} else {
364+
unreachable!("Loop assigns should be annotated only to goto stmt")
365+
}
366+
}
343367
}
344368

345369
/// Predicates

cprover_bindings/src/irep/to_irep.rs

Lines changed: 13 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -548,17 +548,19 @@ impl ToIrep for StmtBody {
548548
arguments_irep(arguments.iter(), mm),
549549
],
550550
),
551-
StmtBody::Goto { dest, loop_invariants } => {
552-
let stmt_goto = code_irep(IrepId::Goto, vec![])
553-
.with_named_sub(IrepId::Destination, Irep::just_string_id(dest.to_string()));
554-
if let Some(inv) = loop_invariants {
555-
stmt_goto.with_named_sub(
556-
IrepId::CSpecLoopInvariant,
557-
inv.clone().and(Expr::bool_true()).to_irep(mm),
558-
)
559-
} else {
560-
stmt_goto
561-
}
551+
StmtBody::Goto { dest, loop_invariants, loop_modifies } => {
552+
let inv = loop_invariants
553+
.clone()
554+
.map(|inv| inv.clone().and(Expr::bool_true()).to_irep(mm));
555+
let assigns = loop_modifies.clone().map(|assigns| {
556+
Irep::just_sub(vec![Irep::just_sub(
557+
assigns.iter().map(|assign| assign.to_irep(mm)).collect(),
558+
)])
559+
});
560+
code_irep(IrepId::Goto, vec![])
561+
.with_named_sub(IrepId::Destination, Irep::just_string_id(dest.to_string()))
562+
.with_named_sub_option(IrepId::CSpecLoopInvariant, inv)
563+
.with_named_sub_option(IrepId::CSpecAssigns, assigns.clone())
562564
}
563565
StmtBody::Ifthenelse { i, t, e } => code_irep(
564566
IrepId::Ifthenelse,

docs/src/reference/experimental/loop-contracts.md

Lines changed: 64 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -61,7 +61,7 @@ Check 10: simple_loop_with_loop_contracts.loop_invariant_base.1
6161
- Description: "Check invariant before entry for loop simple_loop_with_loop_contracts.0"
6262
- Location: simple_while_loop.rs:15:5 in function simple_loop_with_loop_contracts
6363
64-
Check 11: simple_loop_with_loop_contracts.loop_assigns.1
64+
Check 11: simple_loop_with_loop_contracts.loop_modifies.1
6565
- Status: SUCCESS
6666
- Description: "Check assigns clause inclusion for loop simple_loop_with_loop_contracts.0"
6767
- Location: simple_while_loop.rs:15:5 in function simple_loop_with_loop_contracts
@@ -164,6 +164,69 @@ fn contract_proof() {
164164

165165
When loop contracts and function contracts are both enabled (by flags `-Z loop-contracts -Z function-contracts`),
166166
Kani automatically contracts (instead of unwinds) all loops in the functions that we want to prove contracts for.
167+
168+
## Loop modifies clauses:
169+
We allow users to manually specified the `loop_modifies` clauses for memory allocated addresses which can be modified inside the loop body.
170+
The concept is very similar to the `__CPROVER_assigns` clause of CBMC (https://diffblue.github.io/cbmc/contracts-assigns.html).
171+
However, in Kani, the CBMC target is replaced by three Rust types which can be used in the `loop_modifies` clauses:
172+
1. `RawPtr`: We don't allow variable names as targets. Users must use pointers to them instead, which also allows checking modification using borrowed references and aliases.
173+
```Rust
174+
#[kani::proof]
175+
fn main() {
176+
let mut i = 0;
177+
#[kani::loop_invariant(i <= 20)]
178+
#[kani::loop_modifies(&i as *const _)]
179+
while i < 20 {
180+
i = i + 1;
181+
}
182+
}
183+
```
184+
2. `Reference`: Similar to RawPtr, but we also can use it to replace `__CPROVER_object_whole(ptr-expr)`,
185+
Example
186+
```Rust
187+
#[kani::proof]
188+
fn main() {
189+
let mut i = 0;
190+
let mut a: [u8; 20] = kani::any();
191+
#[kani::loop_invariant(i <= 20)]
192+
#[kani::loop_modifies(&i, &a)]
193+
while i < 20 {
194+
a[i] = 1;
195+
i = i + 1;
196+
}
197+
}
198+
```
199+
3. `FatPtr (Slice)`: We use this to replace `__CPROVER_object_from(ptr-expr)`, and `__CPROVER_object_upto(ptr-expr, uint-expr)`.
200+
```Rust
201+
#[kani::proof]
202+
fn main() {
203+
let mut i = 3;
204+
let mut a: [u8; 100] = kani::any();
205+
#[kani::loop_invariant(i >=3 && i <= 20)]
206+
#[kani::loop_modifies(&i , &a[3..20])]
207+
while i < 20 {
208+
a[i] = 1;
209+
i = i + 1;
210+
}
211+
}
212+
```
213+
or
214+
215+
```Rust
216+
use std::ptr::slice_from_raw_parts;
217+
#[kani::proof]
218+
fn main() {
219+
let mut i = 0;
220+
let mut a: [u8; 100] = kani::any();
221+
#[kani::loop_invariant(i <= 20)]
222+
#[kani::loop_modifies(&i , slice_from_raw_parts(a.as_ptr(), 20))]
223+
while i < 20 {
224+
a[i] = 1;
225+
i = i + 1;
226+
}
227+
}
228+
```
229+
167230
## Limitations
168231

169232
Loop contracts comes with the following limitations.

kani-compiler/src/codegen_cprover_gotoc/codegen/statement.rs

Lines changed: 62 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ use super::{PropertyClass, bb_label};
66
use crate::codegen_cprover_gotoc::codegen::function::rustc_smir::region_from_coverage_opaque;
77
use crate::codegen_cprover_gotoc::{GotocCtx, VtableCtx};
88
use crate::unwrap_or_return_codegen_unimplemented_stmt;
9+
use cbmc::goto_program::ExprValue;
910
use cbmc::goto_program::{Expr, Location, Stmt, Type};
1011
use rustc_abi::Size;
1112
use rustc_abi::{FieldsShape, Primitive, TagEncoding, Variants};
@@ -17,12 +18,66 @@ use stable_mir::abi::{ArgAbi, FnAbi, PassMode};
1718
use stable_mir::mir::mono::{Instance, InstanceKind};
1819
use stable_mir::mir::{
1920
AssertMessage, BasicBlockIdx, CopyNonOverlapping, NonDivergingIntrinsic, Operand, Place,
20-
RETURN_LOCAL, Statement, StatementKind, SwitchTargets, Terminator, TerminatorKind,
21+
RETURN_LOCAL, Rvalue, Statement, StatementKind, SwitchTargets, Terminator, TerminatorKind,
2122
};
2223
use stable_mir::ty::{Abi, RigidTy, Span, Ty, TyKind, VariantIdx};
2324
use tracing::{debug, debug_span, trace};
2425

2526
impl GotocCtx<'_> {
27+
pub fn ty_to_assign_target(&self, ty: Ty, expr: &Expr) -> Expr {
28+
match ty.kind() {
29+
TyKind::RigidTy(RigidTy::Ref(_, unref_ty, _))
30+
| TyKind::RigidTy(RigidTy::RawPtr(unref_ty, _)) => match unref_ty.kind() {
31+
TyKind::RigidTy(RigidTy::Slice(slice_ty)) => {
32+
let size = slice_ty.layout().unwrap().shape().size.bytes();
33+
Expr::symbol_expression(
34+
"__CPROVER_object_upto",
35+
Type::code(
36+
vec![
37+
Type::empty().to_pointer().as_parameter(None, Some("ptr".into())),
38+
Type::size_t().as_parameter(None, Some("size".into())),
39+
],
40+
Type::empty(),
41+
),
42+
)
43+
.call(vec![
44+
expr.clone()
45+
.member("data", &self.symbol_table)
46+
.cast_to(Type::empty().to_pointer()),
47+
expr.clone()
48+
.member("len", &self.symbol_table)
49+
.mul(Expr::size_constant(size.try_into().unwrap(), &self.symbol_table)),
50+
])
51+
}
52+
_ => expr.clone().dereference(),
53+
},
54+
_ => expr.clone().dereference(),
55+
}
56+
}
57+
58+
pub fn rvalue_to_assign_targets(&mut self, rvalue: &Rvalue, location: Location) -> Vec<Expr> {
59+
let assigns = self.codegen_rvalue_stable(rvalue, location);
60+
let assigns_value = assigns.value().clone();
61+
let assign_exprs = if let ExprValue::Struct { values } = assigns_value {
62+
values.clone()
63+
} else {
64+
vec![assigns.clone()]
65+
};
66+
match rvalue {
67+
Rvalue::Aggregate(_agg_kind, operands) => {
68+
let mut ptr_exprs = Vec::new();
69+
for (operand, expr) in operands.iter().zip(assign_exprs.iter()) {
70+
let operand_ty = self.operand_ty_stable(operand);
71+
debug!("Ty {:?}", operand_ty);
72+
let ptr_expr = self.ty_to_assign_target(operand_ty, expr);
73+
ptr_exprs.push(ptr_expr)
74+
}
75+
ptr_exprs
76+
}
77+
_ => vec![assigns.dereference()],
78+
}
79+
}
80+
2681
/// Generate Goto-C for MIR [Statement]s.
2782
/// This does not cover all possible "statements" because MIR distinguishes between ordinary
2883
/// statements and [Terminator]s, which can exclusively appear at the end of a basic block.
@@ -36,6 +91,12 @@ impl GotocCtx<'_> {
3691
StatementKind::Assign(lhs, rhs) => {
3792
let lty = self.place_ty_stable(lhs);
3893
let rty = self.rvalue_ty_stable(rhs);
94+
let localname = self.codegen_var_name(&lhs.local);
95+
if localname.contains("kani_loop_modifies") {
96+
let assigns = self.rvalue_to_assign_targets(rhs, location);
97+
self.current_loop_modifies = assigns.clone();
98+
return Stmt::skip(location);
99+
}
39100
// we ignore assignment for all zero size types
40101
if self.is_zst_stable(lty) {
41102
Stmt::skip(location)

kani-compiler/src/codegen_cprover_gotoc/context/goto_ctx.rs

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -78,6 +78,8 @@ pub struct GotocCtx<'tcx> {
7878
pub transformer: BodyTransformation,
7979
/// If there exist some usage of loop contracts int context.
8080
pub has_loop_contracts: bool,
81+
/// Track loop assign clause
82+
pub current_loop_modifies: Vec<Expr>,
8183
}
8284

8385
/// Constructor
@@ -108,6 +110,7 @@ impl<'tcx> GotocCtx<'tcx> {
108110
concurrent_constructs: FxHashMap::default(),
109111
transformer,
110112
has_loop_contracts: false,
113+
current_loop_modifies: Vec::new(),
111114
}
112115
}
113116
}

kani-compiler/src/codegen_cprover_gotoc/overrides/hooks.rs

Lines changed: 9 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -733,16 +733,22 @@ impl GotocHook for LoopInvariantRegister {
733733
// free(0)
734734
// goto target --- with loop contracts annotated.
735735

736+
let mut stmt = Stmt::goto(bb_label(target.unwrap()), loc)
737+
.with_loop_contracts(func_exp.call(fargs).cast_to(Type::CInteger(CIntType::Bool)));
738+
let assigns = gcx.current_loop_modifies.clone();
739+
if !assigns.is_empty() {
740+
stmt = stmt.with_loop_modifies(assigns.clone());
741+
gcx.current_loop_modifies.clear();
742+
}
743+
736744
// Add `free(0)` to make sure the body of `free` won't be dropped to
737745
// satisfy the requirement of DFCC.
738746
Stmt::block(
739747
vec![
740748
BuiltinFn::Free
741749
.call(vec![Expr::pointer_constant(0, Type::void_pointer())], loc)
742750
.as_stmt(loc),
743-
Stmt::goto(bb_label(target.unwrap()), loc).with_loop_contracts(
744-
func_exp.call(fargs).cast_to(Type::CInteger(CIntType::Bool)),
745-
),
751+
stmt,
746752
],
747753
loc,
748754
)

library/kani_macros/src/lib.rs

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -440,6 +440,10 @@ pub fn loop_invariant(attr: TokenStream, item: TokenStream) -> TokenStream {
440440
attr_impl::loop_invariant(attr, item)
441441
}
442442

443+
#[proc_macro_attribute]
444+
pub fn loop_modifies(attr: TokenStream, item: TokenStream) -> TokenStream {
445+
attr_impl::loop_modifies(attr, item)
446+
}
443447
/// This module implements Kani attributes in a way that only Kani's compiler can understand.
444448
/// This code should only be activated when pre-building Kani's sysroot.
445449
#[cfg(kani_sysroot)]
@@ -450,7 +454,7 @@ mod sysroot {
450454
mod loop_contracts;
451455

452456
pub use contracts::{ensures, modifies, proof_for_contract, requires, stub_verified};
453-
pub use loop_contracts::loop_invariant;
457+
pub use loop_contracts::{loop_invariant, loop_modifies};
454458

455459
use super::*;
456460

@@ -629,4 +633,5 @@ mod regular {
629633
no_op!(proof_for_contract);
630634
no_op!(stub_verified);
631635
no_op!(loop_invariant);
636+
no_op!(loop_modifies);
632637
}

library/kani_macros/src/sysroot/loop_contracts/mod.rs

Lines changed: 24 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,9 +7,13 @@
77
use proc_macro::TokenStream;
88
use proc_macro_error2::abort_call_site;
99
use quote::{format_ident, quote};
10+
use syn::punctuated::Punctuated;
1011
use syn::spanned::Spanned;
1112
use syn::token::AndAnd;
12-
use syn::{BinOp, Block, Expr, ExprBinary, Ident, Stmt, parse_quote, visit_mut::VisitMut};
13+
use syn::{
14+
BinOp, Block, Expr, ExprBinary, Ident, Stmt, Token, parse_macro_input, parse_quote,
15+
visit_mut::VisitMut,
16+
};
1317

1418
/*
1519
Transform the loop to support on_entry(expr) : the value of expr before entering the loop
@@ -389,3 +393,22 @@ fn generate_unique_id_from_span(stmt: &Stmt) -> String {
389393
// Create a tuple of location information (file path, start line, start column, end line, end column)
390394
format!("_{:?}_{:?}_{:?}_{:?}", start.line(), start.column(), end.line(), end.column())
391395
}
396+
397+
pub fn loop_modifies(attr: TokenStream, item: TokenStream) -> TokenStream {
398+
let assigns = parse_macro_input!(attr with Punctuated::<Expr, Token![,]>::parse_terminated)
399+
.into_iter()
400+
.collect::<Vec<Expr>>();
401+
let loop_assign_name: String = "kani_loop_modifies".to_owned();
402+
let loop_assign_ident = format_ident!("{}", loop_assign_name);
403+
let loop_assign_stmt: Stmt = parse_quote! {
404+
let #loop_assign_ident = (#(#assigns),*);
405+
};
406+
let loop_stmt: Stmt = syn::parse(item.clone()).unwrap();
407+
let ret: TokenStream = quote!(
408+
{
409+
#loop_assign_stmt
410+
#loop_stmt
411+
})
412+
.into();
413+
ret
414+
}
Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
1+
main.assigns.1\
2+
- Status: SUCCESS\
3+
- Description: "Check that a[var_15] is assignable"
4+
5+
main.assigns.2\
6+
- Status: SUCCESS\
7+
- Description: "Check that i is assignable"
8+
9+
main.loop_invariant_step.1\
10+
- Status: SUCCESS\
11+
- Description: "Check invariant after step for loop main.0"
12+
13+
VERIFICATION:- SUCCESSFUL
Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,23 @@
1+
// Copyright Kani Contributors
2+
// SPDX-License-Identifier: Apache-2.0 OR MIT
3+
4+
// kani-flags: -Z loop-contracts
5+
6+
//! Check the use of loop_modifies for Rust's Fat ptr
7+
8+
#![feature(proc_macro_hygiene)]
9+
#![feature(stmt_expr_attributes)]
10+
11+
use std::ptr::slice_from_raw_parts;
12+
13+
#[kani::proof]
14+
fn main() {
15+
let mut i = 0;
16+
let mut a: [u8; 100] = kani::any();
17+
#[kani::loop_invariant(i <= 20)]
18+
#[kani::loop_modifies(&i , slice_from_raw_parts(a.as_ptr(), 20))]
19+
while i < 20 {
20+
a[i] = 1;
21+
i = i + 1;
22+
}
23+
}

0 commit comments

Comments
 (0)