charon_lib/transform/
ctx.rs

1use crate::ast::*;
2use crate::errors::{ErrorCtx, Level};
3use crate::formatter::{FmtCtx, IntoFormatter};
4use crate::llbc_ast;
5use crate::options::TranslateOptions;
6use crate::pretty::FmtWithCtx;
7use crate::ullbc_ast;
8use std::cell::RefCell;
9use std::{fmt, mem};
10
11/// Simpler context used for rustc-independent code transformation. This only depends on rustc for
12/// its error reporting machinery.
13pub struct TransformCtx {
14    /// The options that control transformation.
15    pub options: TranslateOptions,
16    /// The translated data.
17    pub translated: TranslatedCrate,
18    /// Context for tracking and reporting errors.
19    pub errors: RefCell<ErrorCtx>,
20}
21
22/// A pass that modifies ullbc bodies.
23pub trait UllbcPass: Sync {
24    /// Whether the pass should run.
25    fn should_run(&self, _options: &TranslateOptions) -> bool {
26        true
27    }
28
29    /// Transform a body.
30    fn transform_body(&self, _ctx: &mut TransformCtx, _body: &mut ullbc_ast::ExprBody) {}
31
32    /// Transform a function declaration. This forwards to `transform_body` by default.
33    fn transform_function(&self, ctx: &mut TransformCtx, decl: &mut FunDecl) {
34        if let Some(body) = decl.body.as_unstructured_mut() {
35            self.transform_body(ctx, body)
36        }
37    }
38
39    /// Transform the given context. This forwards to the other methods by default.
40    fn transform_ctx(&self, ctx: &mut TransformCtx) {
41        ctx.for_each_fun_decl(|ctx, decl| {
42            self.log_before_body(ctx, &decl.item_meta.name, &decl.body);
43            self.transform_function(ctx, decl);
44        });
45    }
46
47    /// The name of the pass, used for debug logging. The default implementation uses the type
48    /// name.
49    fn name(&self) -> &str {
50        std::any::type_name::<Self>()
51    }
52
53    /// Log that the pass is about to be run on this body.
54    fn log_before_body(&self, ctx: &TransformCtx, name: &Name, body: &Body) {
55        let fmt_ctx = &ctx.into_fmt();
56        let body_str = body.to_string_with_ctx(fmt_ctx);
57        trace!(
58            "# About to run pass [{}] on `{}`:\n{}",
59            self.name(),
60            name.with_ctx(fmt_ctx),
61            body_str,
62        );
63    }
64}
65
66/// A pass that modifies llbc bodies.
67pub trait LlbcPass: Sync {
68    /// Whether the pass should run.
69    fn should_run(&self, _options: &TranslateOptions) -> bool {
70        true
71    }
72
73    /// Transform a body.
74    fn transform_body(&self, _ctx: &mut TransformCtx, _body: &mut llbc_ast::ExprBody) {}
75
76    /// Transform a function declaration. This forwards to `transform_body` by default.
77    fn transform_function(&self, ctx: &mut TransformCtx, decl: &mut FunDecl) {
78        if let Some(body) = decl.body.as_structured_mut() {
79            self.transform_body(ctx, body)
80        }
81    }
82
83    /// Transform the given context. This forwards to the other methods by default.
84    fn transform_ctx(&self, ctx: &mut TransformCtx) {
85        ctx.for_each_fun_decl(|ctx, decl| {
86            self.log_before_body(ctx, &decl.item_meta.name, &decl.body);
87            self.transform_function(ctx, decl);
88        });
89    }
90
91    /// The name of the pass, used for debug logging. The default implementation uses the type
92    /// name.
93    fn name(&self) -> &str {
94        std::any::type_name::<Self>()
95    }
96
97    /// Log that the pass is about to be run on this body.
98    fn log_before_body(&self, ctx: &TransformCtx, name: &Name, body: &Body) {
99        let fmt_ctx = &ctx.into_fmt();
100        let body_str = body.to_string_with_ctx(fmt_ctx);
101        trace!(
102            "# About to run pass [{}] on `{}`:\n{}",
103            self.name(),
104            name.with_ctx(fmt_ctx),
105            body_str,
106        );
107    }
108}
109
110/// A pass that transforms the crate data.
111pub trait TransformPass: Sync {
112    /// Whether the pass should run.
113    fn should_run(&self, _options: &TranslateOptions) -> bool {
114        true
115    }
116
117    fn transform_ctx(&self, ctx: &mut TransformCtx);
118
119    /// The name of the pass, used for debug logging. The default implementation uses the type
120    /// name.
121    fn name(&self) -> &str {
122        std::any::type_name::<Self>()
123    }
124}
125
126impl<'ctx> TransformCtx {
127    pub(crate) fn has_errors(&self) -> bool {
128        self.errors.borrow().has_errors()
129    }
130
131    /// Span an error and register the error.
132    pub(crate) fn span_err(&self, span: Span, msg: &str, level: Level) -> Error {
133        self.errors
134            .borrow_mut()
135            .span_err(&self.translated, span, msg, level)
136    }
137
138    pub(crate) fn opacity_for_name(&self, name: &Name) -> ItemOpacity {
139        self.options.opacity_for_name(&self.translated, name)
140    }
141
142    pub(crate) fn with_def_id<F, T>(
143        &mut self,
144        def_id: impl Into<ItemId>,
145        def_id_is_local: bool,
146        f: F,
147    ) -> T
148    where
149        F: FnOnce(&mut Self) -> T,
150    {
151        let mut errors = self.errors.borrow_mut();
152        let current_def_id = mem::replace(&mut errors.def_id, Some(def_id.into()));
153        let current_def_id_is_local = mem::replace(&mut errors.def_id_is_local, def_id_is_local);
154        drop(errors); // important: release the refcell "lock"
155        let ret = f(self);
156        let mut errors = self.errors.borrow_mut();
157        errors.def_id = current_def_id;
158        errors.def_id_is_local = current_def_id_is_local;
159        ret
160    }
161
162    /// Mutably iterate over the bodies.
163    /// Warning: we replace each body with `Err(Opaque)` while inspecting it so we can keep access
164    /// to the rest of the crate.
165    pub(crate) fn for_each_body(&mut self, mut f: impl FnMut(&mut Self, &mut Body)) {
166        let fn_ids = self.translated.fun_decls.all_indices();
167        for id in fn_ids {
168            if let Some(decl) = self.translated.fun_decls.get_mut(id) {
169                if decl.body.has_contents() {
170                    let mut body = mem::replace(&mut decl.body, Body::Opaque);
171                    let fun_decl_id = decl.def_id;
172                    let is_local = decl.item_meta.is_local;
173                    self.with_def_id(fun_decl_id, is_local, |ctx| f(ctx, &mut body));
174                    self.translated.fun_decls[id].body = body;
175                }
176            }
177        }
178    }
179
180    /// Mutably iterate over the function declarations.
181    /// Warning: each inspected fundecl becomes inaccessible from `ctx` during the course of this function.
182    pub(crate) fn for_each_fun_decl(&mut self, mut f: impl FnMut(&mut Self, &mut FunDecl)) {
183        let fn_ids = self.translated.fun_decls.all_indices();
184        for id in fn_ids {
185            if let Some(mut decl) = self.translated.fun_decls.remove(id) {
186                let fun_decl_id = decl.def_id;
187                let is_local = decl.item_meta.is_local;
188                self.with_def_id(fun_decl_id, is_local, |ctx| f(ctx, &mut decl));
189                self.translated.fun_decls.set_slot(id, decl);
190            }
191        }
192    }
193
194    /// Iterate mutably over all items, keeping access to `self`. To make this work, we move out
195    /// each item before iterating over it. Items added during traversal will not be iterated over.
196    pub fn for_each_item_mut(&mut self, mut f: impl for<'a> FnMut(&'a mut Self, ItemRefMut<'a>)) {
197        for id in self.translated.all_ids() {
198            if let Some(mut decl) = self.translated.remove_item(id) {
199                f(self, decl.as_mut());
200                self.translated.set_item_slot(id, decl);
201            }
202        }
203    }
204}
205
206impl<'a> IntoFormatter for &'a TransformCtx {
207    type C = FmtCtx<'a>;
208
209    fn into_fmt(self) -> Self::C {
210        self.translated.into_fmt()
211    }
212}
213
214impl fmt::Display for TransformCtx {
215    fn fmt(&self, f: &mut fmt::Formatter) -> fmt::Result {
216        self.translated.fmt(f)
217    }
218}
219
220/// A helper trait that captures common operations in body transformation.
221pub trait BodyTransformCtx: Sized {
222    fn get_crate(&self) -> &TranslatedCrate;
223    fn get_options(&self) -> &TranslateOptions;
224    fn get_params(&self) -> &GenericParams;
225    fn get_locals_mut(&mut self) -> &mut Locals;
226
227    fn insert_storage_live_stmt(&mut self, local: LocalId);
228    fn insert_storage_dead_stmt(&mut self, local: LocalId);
229    fn insert_assn_stmt(&mut self, place: Place, rvalue: Rvalue);
230
231    fn into_fmt(&self) -> FmtCtx<'_> {
232        self.get_crate().into_fmt()
233    }
234
235    /// Create a local & return the place pointing to it
236    fn fresh_var(&mut self, name: Option<String>, ty: Ty) -> Place {
237        let var = self.get_locals_mut().new_var(name, ty);
238        self.insert_storage_live_stmt(var.local_id().unwrap());
239        var
240    }
241
242    /// Assign an rvalue to a place, unless the rvalue is a move in which case we just use the
243    /// moved place.
244    fn rval_to_place(&mut self, rvalue: Rvalue, ty: Ty) -> Place {
245        if let Rvalue::Use(Operand::Move(place)) = rvalue {
246            place
247        } else {
248            let var = self.fresh_var(None, ty);
249            self.insert_assn_stmt(var.clone(), rvalue);
250            var
251        }
252    }
253
254    /// When `from_end` is true, we need to compute `len(p) - last_arg` instead of just using `last_arg`.
255    /// Otherwise, we simply return `last_arg`.
256    /// New local variables are created as needed.
257    ///
258    /// The `last_arg` is either the `offset` for `Index` or the `to` for `Subslice` for the projections.
259    fn compute_subslice_end_idx(
260        &mut self,
261        len_place: &Place,
262        last_arg: Operand,
263        from_end: bool,
264    ) -> Operand {
265        if from_end {
266            // `storage_live(len_var)`
267            // `len_var = len(p)`
268            let len_var = self.fresh_var(None, Ty::mk_usize());
269            self.insert_assn_stmt(
270                len_var.clone(),
271                Rvalue::Len(
272                    len_place.clone(),
273                    len_place.ty().clone(),
274                    len_place
275                        .ty()
276                        .as_adt()
277                        .unwrap()
278                        .generics
279                        .const_generics
280                        .get(0.into())
281                        .cloned(),
282                ),
283            );
284
285            // `storage_live(index_var)`
286            // `index_var = len_var - last_arg`
287            // `storage_dead(len_var)`
288            let index_var = self.fresh_var(None, Ty::mk_usize());
289            self.insert_assn_stmt(
290                index_var.clone(),
291                Rvalue::BinaryOp(
292                    BinOp::Sub(OverflowMode::UB),
293                    Operand::Copy(len_var.clone()),
294                    last_arg,
295                ),
296            );
297            self.insert_storage_dead_stmt(len_var.local_id().unwrap());
298            Operand::Copy(index_var)
299        } else {
300            last_arg
301        }
302    }
303
304    fn is_sized_type_var(&mut self, ty: &Ty) -> bool {
305        match ty.kind() {
306            TyKind::TypeVar(..) => {
307                if self.get_options().hide_marker_traits {
308                    // If we're hiding `Sized`, let's consider everything to be sized.
309                    return true;
310                }
311                let params = self.get_params();
312                for clause in &params.trait_clauses {
313                    let tref = clause.trait_.clone().erase();
314                    // Check if it is `Sized<T>`
315                    if tref.generics.types[0] == *ty
316                        && self
317                            .get_crate()
318                            .trait_decls
319                            .get(tref.id)
320                            .and_then(|decl| decl.item_meta.lang_item.clone())
321                            == Some("sized".into())
322                    {
323                        return true;
324                    }
325                }
326                false
327            }
328            _ => false,
329        }
330    }
331
332    /// Emit statements that compute the metadata of the given place. Returns an operand containing the
333    /// metadata value.
334    ///
335    /// E.g., for:
336    /// ```ignore
337    /// let x = &(*ptr).field;
338    /// ```
339    /// if `(*ptr).field` is a DST like `[i32]`, this will get the metadata from the appropriate
340    /// pointer:
341    /// ```ignore
342    /// let len = ptr.metadata;
343    /// ```
344    /// and return `Operand::Move(len)`.
345    ///
346    fn compute_place_metadata(&mut self, place: &Place) -> Operand {
347        /// No metadata. We use the `unit_metadata` global to avoid having to define unit locals
348        /// everywhere.
349        fn no_metadata<T: BodyTransformCtx>(ctx: &T) -> Operand {
350            let unit_meta = ctx.get_crate().unit_metadata.clone().unwrap();
351            Operand::Copy(Place::new_global(unit_meta, Ty::mk_unit()))
352        }
353
354        /// Compute the metadata for a place. Return `None` if the place has no metadata.
355        fn compute_place_metadata_inner<T: BodyTransformCtx>(
356            ctx: &mut T,
357            place: &Place,
358            metadata_ty: &Ty,
359        ) -> Option<Operand> {
360            let (subplace, proj) = place.as_projection()?;
361            match proj {
362                // The outermost deref we encountered gives us the metadata of the place.
363                ProjectionElem::Deref => {
364                    let metadata_place = subplace
365                        .clone()
366                        .project(ProjectionElem::PtrMetadata, metadata_ty.clone());
367                    Some(Operand::Copy(metadata_place))
368                }
369                ProjectionElem::Field { .. } => {
370                    compute_place_metadata_inner(ctx, subplace, metadata_ty)
371                }
372                // Indexing for array & slice will only result in sized types, hence no metadata
373                ProjectionElem::Index { .. } => None,
374                // Ptr metadata is always sized.
375                ProjectionElem::PtrMetadata { .. } => None,
376                // Subslice must have metadata length, compute the metadata here as `to` - `from`
377                ProjectionElem::Subslice { from, to, from_end } => {
378                    let to_idx = ctx.compute_subslice_end_idx(subplace, *to.clone(), *from_end);
379                    let diff_place = ctx.fresh_var(None, Ty::mk_usize());
380                    ctx.insert_assn_stmt(
381                        diff_place.clone(),
382                        // Overflow is UB and should have been prevented by a bound check beforehand.
383                        Rvalue::BinaryOp(BinOp::Sub(OverflowMode::UB), to_idx, *from.clone()),
384                    );
385                    Some(Operand::Copy(diff_place))
386                }
387            }
388        }
389        trace!(
390            "getting ptr metadata for place: {}",
391            place.with_ctx(&self.into_fmt())
392        );
393        let metadata_ty = place.ty().get_ptr_metadata(&self.get_crate()).into_type();
394        if metadata_ty.is_unit()
395            || matches!(metadata_ty.kind(), TyKind::PtrMetadata(ty) if self.is_sized_type_var(ty))
396        {
397            // If the type var is known to be `Sized`, then no metadata is needed
398            return no_metadata(self);
399        }
400        trace!(
401            "computed metadata type: {}",
402            metadata_ty.with_ctx(&self.into_fmt())
403        );
404        compute_place_metadata_inner(self, place, &metadata_ty).unwrap_or_else(|| no_metadata(self))
405    }
406
407    /// Create a `&` borrow of the place.
408    fn borrow(&mut self, place: Place, kind: BorrowKind) -> Rvalue {
409        let ptr_metadata = self.compute_place_metadata(&place);
410        Rvalue::Ref {
411            place,
412            kind,
413            ptr_metadata,
414        }
415    }
416    /// Create a `&raw` borrow of the place.
417    fn raw_borrow(&mut self, place: Place, kind: RefKind) -> Rvalue {
418        let ptr_metadata = self.compute_place_metadata(&place);
419        Rvalue::RawPtr {
420            place,
421            kind,
422            ptr_metadata,
423        }
424    }
425
426    /// Store a `&` borrow of the place into a new place.
427    fn borrow_to_new_var(&mut self, place: Place, kind: BorrowKind, name: Option<String>) -> Place {
428        let ref_ty = TyKind::Ref(Region::Erased, place.ty().clone(), kind.into()).into_ty();
429        let target_place = self.fresh_var(name, ref_ty);
430        let rvalue = self.borrow(place, kind);
431        self.insert_assn_stmt(target_place.clone(), rvalue);
432        target_place
433    }
434    /// Store a `&raw` borrow of the place into a new place.
435    fn raw_borrow_to_new_var(
436        &mut self,
437        place: Place,
438        kind: RefKind,
439        name: Option<String>,
440    ) -> Place {
441        let ref_ty = TyKind::RawPtr(place.ty().clone(), kind).into_ty();
442        let target_place = self.fresh_var(name, ref_ty);
443        let rvalue = self.raw_borrow(place, kind);
444        self.insert_assn_stmt(target_place.clone(), rvalue);
445        target_place
446    }
447}
448
449pub struct UllbcStatementTransformCtx<'a> {
450    pub ctx: &'a mut TransformCtx,
451    pub params: &'a GenericParams,
452    pub locals: &'a mut Locals,
453    /// Span of the statement being explored
454    pub span: Span,
455    /// Statements to prepend to the statement currently being explored.
456    pub statements: Vec<ullbc_ast::Statement>,
457}
458
459impl BodyTransformCtx for UllbcStatementTransformCtx<'_> {
460    fn get_crate(&self) -> &TranslatedCrate {
461        &self.ctx.translated
462    }
463    fn get_options(&self) -> &TranslateOptions {
464        &self.ctx.options
465    }
466    fn get_params(&self) -> &GenericParams {
467        self.params
468    }
469    fn get_locals_mut(&mut self) -> &mut Locals {
470        self.locals
471    }
472
473    fn insert_storage_live_stmt(&mut self, local: LocalId) {
474        self.statements.push(ullbc_ast::Statement::new(
475            self.span,
476            ullbc_ast::StatementKind::StorageLive(local),
477        ));
478    }
479
480    fn insert_assn_stmt(&mut self, place: Place, rvalue: Rvalue) {
481        self.statements.push(ullbc_ast::Statement::new(
482            self.span,
483            ullbc_ast::StatementKind::Assign(place, rvalue),
484        ));
485    }
486
487    fn insert_storage_dead_stmt(&mut self, local: LocalId) {
488        self.statements.push(ullbc_ast::Statement::new(
489            self.span,
490            ullbc_ast::StatementKind::StorageDead(local),
491        ));
492    }
493}
494
495pub struct LlbcStatementTransformCtx<'a> {
496    pub ctx: &'a mut TransformCtx,
497    pub params: &'a GenericParams,
498    pub locals: &'a mut Locals,
499    /// Span of the statement being explored
500    pub span: Span,
501    /// Statements to prepend to the statement currently being explored.
502    pub statements: Vec<llbc_ast::Statement>,
503}
504
505impl BodyTransformCtx for LlbcStatementTransformCtx<'_> {
506    fn get_crate(&self) -> &TranslatedCrate {
507        &self.ctx.translated
508    }
509    fn get_options(&self) -> &TranslateOptions {
510        &self.ctx.options
511    }
512    fn get_params(&self) -> &GenericParams {
513        self.params
514    }
515    fn get_locals_mut(&mut self) -> &mut Locals {
516        self.locals
517    }
518
519    fn insert_storage_live_stmt(&mut self, local: LocalId) {
520        self.statements.push(llbc_ast::Statement::new(
521            self.span,
522            llbc_ast::StatementKind::StorageLive(local),
523        ));
524    }
525
526    fn insert_assn_stmt(&mut self, place: Place, rvalue: Rvalue) {
527        self.statements.push(llbc_ast::Statement::new(
528            self.span,
529            llbc_ast::StatementKind::Assign(place, rvalue),
530        ));
531    }
532
533    fn insert_storage_dead_stmt(&mut self, local: LocalId) {
534        self.statements.push(llbc_ast::Statement::new(
535            self.span,
536            llbc_ast::StatementKind::StorageDead(local),
537        ));
538    }
539}
540
541impl FunDecl {
542    pub fn transform_ullbc_statements(
543        &mut self,
544        ctx: &mut TransformCtx,
545        mut f: impl FnMut(&mut UllbcStatementTransformCtx, &mut ullbc_ast::Statement),
546    ) {
547        if let Some(body) = self.body.as_unstructured_mut() {
548            let mut ctx = UllbcStatementTransformCtx {
549                ctx,
550                params: &self.generics,
551                locals: &mut body.locals,
552                span: self.item_meta.span,
553                statements: Vec::new(),
554            };
555            body.body.iter_mut().for_each(|block| {
556                ctx.statements = Vec::with_capacity(block.statements.len());
557                for mut st in mem::take(&mut block.statements) {
558                    ctx.span = st.span;
559                    f(&mut ctx, &mut st);
560                    ctx.statements.push(st);
561                }
562                block.statements = mem::take(&mut ctx.statements);
563            });
564        }
565    }
566
567    pub fn transform_ullbc_terminators(
568        &mut self,
569        ctx: &mut TransformCtx,
570        mut f: impl FnMut(&mut UllbcStatementTransformCtx, &mut ullbc_ast::Terminator),
571    ) {
572        if let Some(body) = self.body.as_unstructured_mut() {
573            let mut ctx = UllbcStatementTransformCtx {
574                ctx,
575                params: &self.generics,
576                locals: &mut body.locals,
577                span: self.item_meta.span,
578                statements: Vec::new(),
579            };
580            body.body.iter_mut().for_each(|block| {
581                ctx.span = block.terminator.span;
582                ctx.statements = mem::take(&mut block.statements);
583                f(&mut ctx, &mut block.terminator);
584                block.statements = mem::take(&mut ctx.statements);
585            });
586        }
587    }
588
589    pub fn transform_ullbc_operands(
590        &mut self,
591        ctx: &mut TransformCtx,
592        mut f: impl FnMut(&mut UllbcStatementTransformCtx, &mut Operand),
593    ) {
594        self.transform_ullbc_statements(ctx, |ctx, st| {
595            st.kind.dyn_visit_in_body_mut(|op: &mut Operand| f(ctx, op));
596        });
597        self.transform_ullbc_terminators(ctx, |ctx, st| {
598            st.kind.dyn_visit_in_body_mut(|op: &mut Operand| f(ctx, op));
599        });
600    }
601
602    pub fn transform_llbc_statements(
603        &mut self,
604        ctx: &mut TransformCtx,
605        mut f: impl FnMut(&mut LlbcStatementTransformCtx, &mut llbc_ast::Statement),
606    ) {
607        if let Some(body) = self.body.as_structured_mut() {
608            let mut ctx = LlbcStatementTransformCtx {
609                ctx,
610                locals: &mut body.locals,
611                statements: Vec::new(),
612                span: self.item_meta.span,
613                params: &self.generics,
614            };
615            body.body.visit_blocks_bwd(|block: &mut llbc_ast::Block| {
616                ctx.statements = Vec::with_capacity(block.statements.len());
617                for mut st in mem::take(&mut block.statements) {
618                    ctx.span = st.span;
619                    f(&mut ctx, &mut st);
620                    ctx.statements.push(st);
621                }
622                block.statements = mem::take(&mut ctx.statements)
623            })
624        }
625    }
626}