Path: blob/main/cranelift/interpreter/src/interpreter.rs
1692 views
//! Cranelift IR interpreter.1//!2//! This module partially contains the logic for interpreting Cranelift IR.34use crate::address::{Address, AddressFunctionEntry, AddressRegion, AddressSize};5use crate::environment::{FuncIndex, FunctionStore};6use crate::frame::Frame;7use crate::instruction::DfgInstructionContext;8use crate::state::{InterpreterFunctionRef, MemoryError, State};9use crate::step::{ControlFlow, CraneliftTrap, StepError, step};10use crate::value::{DataValueExt, ValueError};11use cranelift_codegen::data_value::DataValue;12use cranelift_codegen::ir::{13ArgumentPurpose, Block, Endianness, ExternalName, FuncRef, Function, GlobalValue,14GlobalValueData, LibCall, MemFlags, StackSlot, Type,15};16use log::trace;17use smallvec::SmallVec;18use std::fmt::Debug;19use std::iter;20use thiserror::Error;2122/// The Cranelift interpreter; this contains some high-level functions to control the interpreter's23/// flow. The interpreter state is defined separately (see [InterpreterState]) as the execution24/// semantics for each Cranelift instruction (see [step]).25pub struct Interpreter<'a> {26state: InterpreterState<'a>,27fuel: Option<u64>,28}2930impl<'a> Interpreter<'a> {31pub fn new(state: InterpreterState<'a>) -> Self {32Self { state, fuel: None }33}3435/// The `fuel` mechanism sets a number of instructions that36/// the interpreter can execute before stopping. If this37/// value is `None` (the default), no limit is imposed.38pub fn with_fuel(self, fuel: Option<u64>) -> Self {39Self { fuel, ..self }40}4142/// Call a function by name; this is a helpful proxy for [Interpreter::call_by_index].43pub fn call_by_name(44&mut self,45func_name: &str,46arguments: &[DataValue],47) -> Result<ControlFlow<'a>, InterpreterError> {48let index = self49.state50.functions51.index_of(func_name)52.ok_or_else(|| InterpreterError::UnknownFunctionName(func_name.to_string()))?;53self.call_by_index(index, arguments)54}5556/// Call a function by its index in the [FunctionStore]; this is a proxy for57/// `Interpreter::call`.58pub fn call_by_index(59&mut self,60index: FuncIndex,61arguments: &[DataValue],62) -> Result<ControlFlow<'a>, InterpreterError> {63match self.state.functions.get_by_index(index) {64None => Err(InterpreterError::UnknownFunctionIndex(index)),65Some(func) => self.call(func, arguments),66}67}6869/// Interpret a call to a [Function] given its [DataValue] arguments.70fn call(71&mut self,72function: &'a Function,73arguments: &[DataValue],74) -> Result<ControlFlow<'a>, InterpreterError> {75trace!("Call: {}({:?})", function.name, arguments);76let first_block = function77.layout78.blocks()79.next()80.expect("to have a first block");81let parameters = function.dfg.block_params(first_block);82self.state.push_frame(function);83self.state84.current_frame_mut()85.set_all(parameters, arguments.to_vec());8687self.block(first_block)88}8990/// Interpret a [Block] in a [Function]. This drives the interpretation over sequences of91/// instructions, which may continue in other blocks, until the function returns.92fn block(&mut self, block: Block) -> Result<ControlFlow<'a>, InterpreterError> {93trace!("Block: {block}");94let function = self.state.current_frame_mut().function();95let layout = &function.layout;96let mut maybe_inst = layout.first_inst(block);97while let Some(inst) = maybe_inst {98if self.consume_fuel() == FuelResult::Stop {99return Err(InterpreterError::FuelExhausted);100}101102let inst_context = DfgInstructionContext::new(inst, &function.dfg);103match step(&mut self.state, inst_context)? {104ControlFlow::Assign(values) => {105self.state106.current_frame_mut()107.set_all(function.dfg.inst_results(inst), values.to_vec());108maybe_inst = layout.next_inst(inst)109}110ControlFlow::Continue => maybe_inst = layout.next_inst(inst),111ControlFlow::ContinueAt(block, block_arguments) => {112trace!("Block: {block}");113self.state114.current_frame_mut()115.set_all(function.dfg.block_params(block), block_arguments.to_vec());116maybe_inst = layout.first_inst(block)117}118ControlFlow::Call(called_function, arguments) => {119match self.call(called_function, &arguments)? {120ControlFlow::Return(rets) => {121self.state122.current_frame_mut()123.set_all(function.dfg.inst_results(inst), rets.to_vec());124maybe_inst = layout.next_inst(inst)125}126ControlFlow::Trap(trap) => return Ok(ControlFlow::Trap(trap)),127cf => {128panic!("invalid control flow after call: {cf:?}")129}130}131}132ControlFlow::ReturnCall(callee, args) => {133self.state.pop_frame();134135return match self.call(callee, &args)? {136ControlFlow::Return(rets) => Ok(ControlFlow::Return(rets)),137ControlFlow::Trap(trap) => Ok(ControlFlow::Trap(trap)),138cf => {139panic!("invalid control flow after return_call: {cf:?}")140}141};142}143ControlFlow::Return(returned_values) => {144self.state.pop_frame();145return Ok(ControlFlow::Return(returned_values));146}147ControlFlow::Trap(trap) => return Ok(ControlFlow::Trap(trap)),148}149}150Err(InterpreterError::Unreachable)151}152153fn consume_fuel(&mut self) -> FuelResult {154match self.fuel {155Some(0) => FuelResult::Stop,156Some(ref mut n) => {157*n -= 1;158FuelResult::Continue159}160161// We do not have fuel enabled, so unconditionally continue162None => FuelResult::Continue,163}164}165}166167#[derive(Debug, PartialEq, Clone)]168/// The result of consuming fuel. Signals if the caller should stop or continue.169pub enum FuelResult {170/// We still have `fuel` available and should continue execution.171Continue,172/// The available `fuel` has been exhausted, we should stop now.173Stop,174}175176/// The ways interpretation can fail.177#[derive(Error, Debug)]178pub enum InterpreterError {179#[error("failed to interpret instruction")]180StepError(#[from] StepError),181#[error("reached an unreachable statement")]182Unreachable,183#[error("unknown function index (has it been added to the function store?): {0}")]184UnknownFunctionIndex(FuncIndex),185#[error("unknown function with name (has it been added to the function store?): {0}")]186UnknownFunctionName(String),187#[error("value error")]188ValueError(#[from] ValueError),189#[error("fuel exhausted")]190FuelExhausted,191}192193pub type LibCallValues = SmallVec<[DataValue; 1]>;194pub type LibCallHandler = fn(LibCall, LibCallValues) -> Result<LibCallValues, CraneliftTrap>;195196/// Maintains the [Interpreter]'s state, implementing the [State] trait.197pub struct InterpreterState<'a> {198pub functions: FunctionStore<'a>,199pub libcall_handler: LibCallHandler,200pub frame_stack: Vec<Frame<'a>>,201/// Number of bytes from the bottom of the stack where the current frame's stack space is202pub frame_offset: usize,203pub stack: Vec<u8>,204pub pinned_reg: DataValue,205pub native_endianness: Endianness,206}207208impl Default for InterpreterState<'_> {209fn default() -> Self {210let native_endianness = if cfg!(target_endian = "little") {211Endianness::Little212} else {213Endianness::Big214};215Self {216functions: FunctionStore::default(),217libcall_handler: |_, _| Err(CraneliftTrap::UnreachableCodeReached),218frame_stack: vec![],219frame_offset: 0,220stack: Vec::with_capacity(1024),221pinned_reg: DataValue::I64(0),222native_endianness,223}224}225}226227impl<'a> InterpreterState<'a> {228pub fn with_function_store(self, functions: FunctionStore<'a>) -> Self {229Self { functions, ..self }230}231232/// Registers a libcall handler233pub fn with_libcall_handler(mut self, handler: LibCallHandler) -> Self {234self.libcall_handler = handler;235self236}237}238239impl<'a> State<'a> for InterpreterState<'a> {240fn get_function(&self, func_ref: FuncRef) -> Option<&'a Function> {241self.functions242.get_from_func_ref(func_ref, self.frame_stack.last().unwrap().function())243}244fn get_current_function(&self) -> &'a Function {245self.current_frame().function()246}247248fn get_libcall_handler(&self) -> LibCallHandler {249self.libcall_handler250}251252fn push_frame(&mut self, function: &'a Function) {253if let Some(frame) = self.frame_stack.iter().last() {254self.frame_offset += frame.function().fixed_stack_size() as usize;255}256257// Grow the stack by the space necessary for this frame258self.stack259.extend(iter::repeat(0).take(function.fixed_stack_size() as usize));260261self.frame_stack.push(Frame::new(function));262}263fn pop_frame(&mut self) {264if let Some(frame) = self.frame_stack.pop() {265// Shorten the stack after exiting the frame266self.stack267.truncate(self.stack.len() - frame.function().fixed_stack_size() as usize);268269// Reset frame_offset to the start of this function270if let Some(frame) = self.frame_stack.iter().last() {271self.frame_offset -= frame.function().fixed_stack_size() as usize;272}273}274}275276fn current_frame_mut(&mut self) -> &mut Frame<'a> {277let num_frames = self.frame_stack.len();278match num_frames {2790 => panic!("unable to retrieve the current frame because no frames were pushed"),280_ => &mut self.frame_stack[num_frames - 1],281}282}283284fn current_frame(&self) -> &Frame<'a> {285let num_frames = self.frame_stack.len();286match num_frames {2870 => panic!("unable to retrieve the current frame because no frames were pushed"),288_ => &self.frame_stack[num_frames - 1],289}290}291292fn stack_address(293&self,294size: AddressSize,295slot: StackSlot,296offset: u64,297) -> Result<Address, MemoryError> {298let stack_slots = &self.get_current_function().sized_stack_slots;299let stack_slot = &stack_slots[slot];300301// offset must be `0 <= Offset < sizeof(SS)`302if offset >= stack_slot.size as u64 {303return Err(MemoryError::InvalidOffset {304offset,305max: stack_slot.size as u64,306});307}308309// Calculate the offset from the current frame to the requested stack slot310let slot_offset: u64 = stack_slots311.keys()312.filter(|k| k < &slot)313.map(|k| stack_slots[k].size as u64)314.sum();315316let final_offset = self.frame_offset as u64 + slot_offset + offset;317Address::from_parts(size, AddressRegion::Stack, 0, final_offset)318}319320fn checked_load(321&self,322addr: Address,323ty: Type,324mem_flags: MemFlags,325) -> Result<DataValue, MemoryError> {326let load_size = ty.bytes() as usize;327let addr_start = addr.offset as usize;328let addr_end = addr_start + load_size;329330let src = match addr.region {331AddressRegion::Stack => {332if addr_end > self.stack.len() {333return Err(MemoryError::OutOfBoundsLoad {334addr,335load_size,336mem_flags,337});338}339340&self.stack[addr_start..addr_end]341}342_ => unimplemented!(),343};344345// Aligned flag is set and address is not aligned for the given type346if mem_flags.aligned() && addr_start % load_size != 0 {347return Err(MemoryError::MisalignedLoad { addr, load_size });348}349350Ok(match mem_flags.endianness(self.native_endianness) {351Endianness::Big => DataValue::read_from_slice_be(src, ty),352Endianness::Little => DataValue::read_from_slice_le(src, ty),353})354}355356fn checked_store(357&mut self,358addr: Address,359v: DataValue,360mem_flags: MemFlags,361) -> Result<(), MemoryError> {362let store_size = v.ty().bytes() as usize;363let addr_start = addr.offset as usize;364let addr_end = addr_start + store_size;365366let dst = match addr.region {367AddressRegion::Stack => {368if addr_end > self.stack.len() {369return Err(MemoryError::OutOfBoundsStore {370addr,371store_size,372mem_flags,373});374}375376&mut self.stack[addr_start..addr_end]377}378_ => unimplemented!(),379};380381// Aligned flag is set and address is not aligned for the given type382if mem_flags.aligned() && addr_start % store_size != 0 {383return Err(MemoryError::MisalignedStore { addr, store_size });384}385386Ok(match mem_flags.endianness(self.native_endianness) {387Endianness::Big => v.write_to_slice_be(dst),388Endianness::Little => v.write_to_slice_le(dst),389})390}391392fn function_address(393&self,394size: AddressSize,395name: &ExternalName,396) -> Result<Address, MemoryError> {397let curr_func = self.get_current_function();398let (entry, index) = match name {399ExternalName::User(username) => {400let ext_name = &curr_func.params.user_named_funcs()[*username];401402// TODO: This is not optimal since we are looking up by string name403let index = self.functions.index_of(&ext_name.to_string()).unwrap();404405(AddressFunctionEntry::UserFunction, index.as_u32())406}407408ExternalName::TestCase(testname) => {409// TODO: This is not optimal since we are looking up by string name410let index = self.functions.index_of(&testname.to_string()).unwrap();411412(AddressFunctionEntry::UserFunction, index.as_u32())413}414ExternalName::LibCall(libcall) => {415// We don't properly have a "libcall" store, but we can use `LibCall::all()`416// and index into that.417let index = LibCall::all_libcalls()418.iter()419.position(|lc| lc == libcall)420.unwrap();421422(AddressFunctionEntry::LibCall, index as u32)423}424_ => unimplemented!("function_address: {:?}", name),425};426427Address::from_parts(size, AddressRegion::Function, entry as u64, index as u64)428}429430fn get_function_from_address(&self, address: Address) -> Option<InterpreterFunctionRef<'a>> {431let index = address.offset as u32;432if address.region != AddressRegion::Function {433return None;434}435436match AddressFunctionEntry::from(address.entry) {437AddressFunctionEntry::UserFunction => self438.functions439.get_by_index(FuncIndex::from_u32(index))440.map(InterpreterFunctionRef::from),441442AddressFunctionEntry::LibCall => LibCall::all_libcalls()443.get(index as usize)444.copied()445.map(InterpreterFunctionRef::from),446}447}448449/// Non-Recursively resolves a global value until its address is found450fn resolve_global_value(&self, gv: GlobalValue) -> Result<DataValue, MemoryError> {451// Resolving a Global Value is a "pointer" chasing operation that lends itself to452// using a recursive solution. However, resolving this in a recursive manner453// is a bad idea because its very easy to add a bunch of global values and454// blow up the call stack.455//456// Adding to the challenges of this, is that the operations possible with GlobalValues457// mean that we cannot use a simple loop to resolve each global value, we must keep458// a pending list of operations.459460// These are the possible actions that we can perform461#[derive(Debug)]462enum ResolveAction {463Resolve(GlobalValue),464/// Perform an add on the current address465Add(DataValue),466/// Load From the current address and replace it with the loaded value467Load {468/// Offset added to the base pointer before doing the load.469offset: i32,470471/// Type of the loaded value.472global_type: Type,473},474}475476let func = self.get_current_function();477478// We start with a sentinel value that will fail if we try to load / add to it479// without resolving the base GV First.480let mut current_val = DataValue::I8(0);481let mut action_stack = vec![ResolveAction::Resolve(gv)];482483loop {484match action_stack.pop() {485Some(ResolveAction::Resolve(gv)) => match func.global_values[gv] {486GlobalValueData::VMContext => {487// Fetch the VMContext value from the values of the first block in the function488let index = func489.signature490.params491.iter()492.enumerate()493.find(|(_, p)| p.purpose == ArgumentPurpose::VMContext)494.map(|(i, _)| i)495// This should be validated by the verifier496.expect("No VMCtx argument was found, but one is referenced");497498let first_block =499func.layout.blocks().next().expect("to have a first block");500let vmctx_value = func.dfg.block_params(first_block)[index];501current_val = self.current_frame().get(vmctx_value).clone();502}503GlobalValueData::Load {504base,505offset,506global_type,507..508} => {509action_stack.push(ResolveAction::Load {510offset: offset.into(),511global_type,512});513action_stack.push(ResolveAction::Resolve(base));514}515GlobalValueData::IAddImm {516base,517offset,518global_type,519} => {520let offset: i64 = offset.into();521let dv = DataValue::int(offset as i128, global_type)522.map_err(|_| MemoryError::InvalidAddressType(global_type))?;523action_stack.push(ResolveAction::Add(dv));524action_stack.push(ResolveAction::Resolve(base));525}526GlobalValueData::Symbol { .. } => unimplemented!(),527GlobalValueData::DynScaleTargetConst { .. } => unimplemented!(),528},529Some(ResolveAction::Add(dv)) => {530current_val = current_val531.add(dv.clone())532.map_err(|_| MemoryError::InvalidAddress(dv))?;533}534Some(ResolveAction::Load {535offset,536global_type,537}) => {538let mut addr = Address::try_from(current_val)?;539let mem_flags = MemFlags::trusted();540// We can forego bounds checking here since its performed in `checked_load`541addr.offset += offset as u64;542current_val = self.checked_load(addr, global_type, mem_flags)?;543}544545// We are done resolving this, return the current value546None => return Ok(current_val),547}548}549}550551fn get_pinned_reg(&self) -> DataValue {552self.pinned_reg.clone()553}554555fn set_pinned_reg(&mut self, v: DataValue) {556self.pinned_reg = v;557}558}559560#[cfg(test)]561mod tests {562use super::*;563use crate::step::CraneliftTrap;564use cranelift_codegen::ir::TrapCode;565use cranelift_codegen::ir::immediates::Ieee32;566use cranelift_reader::parse_functions;567use smallvec::smallvec;568569// Most interpreter tests should use the more ergonomic `test interpret` filetest but this570// unit test serves as a sanity check that the interpreter still works without all of the571// filetest infrastructure.572#[test]573fn sanity() {574let code = "function %test() -> i8 {575block0:576v0 = iconst.i32 1577v1 = iadd_imm v0, 1578v2 = irsub_imm v1, 44 ; 44 - 2 == 42 (see irsub_imm's semantics)579v3 = icmp_imm eq v2, 42580return v3581}";582583let func = parse_functions(code).unwrap().into_iter().next().unwrap();584let mut env = FunctionStore::default();585env.add(func.name.to_string(), &func);586let state = InterpreterState::default().with_function_store(env);587let result = Interpreter::new(state).call_by_name("%test", &[]).unwrap();588589assert_eq!(result, ControlFlow::Return(smallvec![DataValue::I8(1)]));590}591592// We don't have a way to check for traps with the current filetest infrastructure593#[test]594fn udiv_by_zero_traps() {595let code = "function %test() -> i32 {596block0:597v0 = iconst.i32 1598v1 = udiv_imm.i32 v0, 0599return v1600}";601602let func = parse_functions(code).unwrap().into_iter().next().unwrap();603let mut env = FunctionStore::default();604env.add(func.name.to_string(), &func);605let state = InterpreterState::default().with_function_store(env);606let trap = Interpreter::new(state).call_by_name("%test", &[]).unwrap();607608assert_eq!(609trap,610ControlFlow::Trap(CraneliftTrap::User(TrapCode::INTEGER_DIVISION_BY_ZERO))611);612}613614#[test]615fn sdiv_min_by_neg_one_traps_with_overflow() {616let code = "function %test() -> i8 {617block0:618v0 = iconst.i32 -2147483648619v1 = sdiv_imm.i32 v0, -1620return v1621}";622623let func = parse_functions(code).unwrap().into_iter().next().unwrap();624let mut env = FunctionStore::default();625env.add(func.name.to_string(), &func);626let state = InterpreterState::default().with_function_store(env);627let result = Interpreter::new(state).call_by_name("%test", &[]).unwrap();628629match result {630ControlFlow::Trap(CraneliftTrap::User(TrapCode::INTEGER_OVERFLOW)) => {}631_ => panic!("Unexpected ControlFlow: {result:?}"),632}633}634635// This test verifies that functions can refer to each other using the function store. A double indirection is636// required, which is tricky to get right: a referenced function is a FuncRef when called but a FuncIndex inside the637// function store. This test would preferably be a CLIF filetest but the filetest infrastructure only looks at a638// single function at a time--we need more than one function in the store for this test.639#[test]640fn function_references() {641let code = "642function %child(i32) -> i32 {643block0(v0: i32):644v1 = iadd_imm v0, -1645return v1646}647648function %parent(i32) -> i32 {649fn42 = %child(i32) -> i32650block0(v0: i32):651v1 = iadd_imm v0, 1652v2 = call fn42(v1)653return v2654}";655656let mut env = FunctionStore::default();657let funcs = parse_functions(code).unwrap().to_vec();658funcs.iter().for_each(|f| env.add(f.name.to_string(), f));659660let state = InterpreterState::default().with_function_store(env);661let result = Interpreter::new(state)662.call_by_name("%parent", &[DataValue::I32(0)])663.unwrap();664665assert_eq!(result, ControlFlow::Return(smallvec![DataValue::I32(0)]));666}667668#[test]669fn fuel() {670let code = "function %test() -> i8 {671block0:672v0 = iconst.i32 1673v1 = iadd_imm v0, 1674return v1675}";676677let func = parse_functions(code).unwrap().into_iter().next().unwrap();678let mut env = FunctionStore::default();679env.add(func.name.to_string(), &func);680681// The default interpreter should not enable the fuel mechanism682let state = InterpreterState::default().with_function_store(env.clone());683let result = Interpreter::new(state).call_by_name("%test", &[]).unwrap();684685assert_eq!(result, ControlFlow::Return(smallvec![DataValue::I32(2)]));686687// With 2 fuel, we should execute the iconst and iadd, but not the return thus giving a688// fuel exhausted error689let state = InterpreterState::default().with_function_store(env.clone());690let result = Interpreter::new(state)691.with_fuel(Some(2))692.call_by_name("%test", &[]);693match result {694Err(InterpreterError::FuelExhausted) => {}695_ => panic!("Expected Err(FuelExhausted), but got {result:?}"),696}697698// With 3 fuel, we should be able to execute the return instruction, and complete the test699let state = InterpreterState::default().with_function_store(env.clone());700let result = Interpreter::new(state)701.with_fuel(Some(3))702.call_by_name("%test", &[])703.unwrap();704705assert_eq!(result, ControlFlow::Return(smallvec![DataValue::I32(2)]));706}707708// Verifies that writing to the stack on a called function does not overwrite the parents709// stack slots.710#[test]711fn stack_slots_multi_functions() {712let code = "713function %callee(i64, i64) -> i64 {714ss0 = explicit_slot 8715ss1 = explicit_slot 8716717block0(v0: i64, v1: i64):718stack_store.i64 v0, ss0719stack_store.i64 v1, ss1720v2 = stack_load.i64 ss0721v3 = stack_load.i64 ss1722v4 = iadd.i64 v2, v3723return v4724}725726function %caller(i64, i64, i64, i64) -> i64 {727fn0 = %callee(i64, i64) -> i64728ss0 = explicit_slot 8729ss1 = explicit_slot 8730731block0(v0: i64, v1: i64, v2: i64, v3: i64):732stack_store.i64 v0, ss0733stack_store.i64 v1, ss1734735v4 = call fn0(v2, v3)736737v5 = stack_load.i64 ss0738v6 = stack_load.i64 ss1739740v7 = iadd.i64 v4, v5741v8 = iadd.i64 v7, v6742743return v8744}";745746let mut env = FunctionStore::default();747let funcs = parse_functions(code).unwrap().to_vec();748funcs.iter().for_each(|f| env.add(f.name.to_string(), f));749750let state = InterpreterState::default().with_function_store(env);751let result = Interpreter::new(state)752.call_by_name(753"%caller",754&[755DataValue::I64(3),756DataValue::I64(5),757DataValue::I64(7),758DataValue::I64(11),759],760)761.unwrap();762763assert_eq!(result, ControlFlow::Return(smallvec![DataValue::I64(26)]))764}765766#[test]767fn out_of_slot_write_traps() {768let code = "769function %stack_write() {770ss0 = explicit_slot 8771772block0:773v0 = iconst.i64 10774stack_store.i64 v0, ss0+8775return776}";777778let func = parse_functions(code).unwrap().into_iter().next().unwrap();779let mut env = FunctionStore::default();780env.add(func.name.to_string(), &func);781let state = InterpreterState::default().with_function_store(env);782let trap = Interpreter::new(state)783.call_by_name("%stack_write", &[])784.unwrap();785786assert_eq!(787trap,788ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))789);790}791792#[test]793fn partial_out_of_slot_write_traps() {794let code = "795function %stack_write() {796ss0 = explicit_slot 8797798block0:799v0 = iconst.i64 10800stack_store.i64 v0, ss0+4801return802}";803804let func = parse_functions(code).unwrap().into_iter().next().unwrap();805let mut env = FunctionStore::default();806env.add(func.name.to_string(), &func);807let state = InterpreterState::default().with_function_store(env);808let trap = Interpreter::new(state)809.call_by_name("%stack_write", &[])810.unwrap();811812assert_eq!(813trap,814ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))815);816}817818#[test]819fn out_of_slot_read_traps() {820let code = "821function %stack_load() {822ss0 = explicit_slot 8823824block0:825v0 = stack_load.i64 ss0+8826return827}";828829let func = parse_functions(code).unwrap().into_iter().next().unwrap();830let mut env = FunctionStore::default();831env.add(func.name.to_string(), &func);832let state = InterpreterState::default().with_function_store(env);833let trap = Interpreter::new(state)834.call_by_name("%stack_load", &[])835.unwrap();836837assert_eq!(838trap,839ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))840);841}842843#[test]844fn partial_out_of_slot_read_traps() {845let code = "846function %stack_load() {847ss0 = explicit_slot 8848849block0:850v0 = stack_load.i64 ss0+4851return852}";853854let func = parse_functions(code).unwrap().into_iter().next().unwrap();855let mut env = FunctionStore::default();856env.add(func.name.to_string(), &func);857let state = InterpreterState::default().with_function_store(env);858let trap = Interpreter::new(state)859.call_by_name("%stack_load", &[])860.unwrap();861862assert_eq!(863trap,864ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))865);866}867868#[test]869fn partial_out_of_slot_read_by_addr_traps() {870let code = "871function %stack_load() {872ss0 = explicit_slot 8873874block0:875v0 = stack_addr.i64 ss0876v1 = iconst.i64 4877v2 = iadd.i64 v0, v1878v3 = load.i64 v2879return880}";881882let func = parse_functions(code).unwrap().into_iter().next().unwrap();883let mut env = FunctionStore::default();884env.add(func.name.to_string(), &func);885let state = InterpreterState::default().with_function_store(env);886let trap = Interpreter::new(state)887.call_by_name("%stack_load", &[])888.unwrap();889890assert_eq!(891trap,892ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))893);894}895896#[test]897fn partial_out_of_slot_write_by_addr_traps() {898let code = "899function %stack_store() {900ss0 = explicit_slot 8901902block0:903v0 = stack_addr.i64 ss0904v1 = iconst.i64 4905v2 = iadd.i64 v0, v1906store.i64 v1, v2907return908}";909910let func = parse_functions(code).unwrap().into_iter().next().unwrap();911let mut env = FunctionStore::default();912env.add(func.name.to_string(), &func);913let state = InterpreterState::default().with_function_store(env);914let trap = Interpreter::new(state)915.call_by_name("%stack_store", &[])916.unwrap();917918assert_eq!(919trap,920ControlFlow::Trap(CraneliftTrap::User(TrapCode::HEAP_OUT_OF_BOUNDS))921);922}923924#[test]925fn srem_trap() {926let code = "function %test() -> i64 {927block0:928v0 = iconst.i64 0x8000_0000_0000_0000929v1 = iconst.i64 -1930v2 = srem.i64 v0, v1931return v2932}";933934let func = parse_functions(code).unwrap().into_iter().next().unwrap();935let mut env = FunctionStore::default();936env.add(func.name.to_string(), &func);937let state = InterpreterState::default().with_function_store(env);938let trap = Interpreter::new(state).call_by_name("%test", &[]).unwrap();939940assert_eq!(941trap,942ControlFlow::Trap(CraneliftTrap::User(TrapCode::INTEGER_OVERFLOW))943);944}945946#[test]947fn libcall() {948let code = "function %test() -> i64 {949fn0 = colocated %CeilF32 (f32) -> f32 fast950block0:951v1 = f32const 0x0.5952v2 = call fn0(v1)953return v2954}";955956let func = parse_functions(code).unwrap().into_iter().next().unwrap();957let mut env = FunctionStore::default();958env.add(func.name.to_string(), &func);959let state = InterpreterState::default()960.with_function_store(env)961.with_libcall_handler(|libcall, args| {962Ok(smallvec![match (libcall, &args[..]) {963(LibCall::CeilF32, [DataValue::F32(a)]) => DataValue::F32(a.ceil()),964_ => panic!("Unexpected args"),965}])966});967968let result = Interpreter::new(state).call_by_name("%test", &[]).unwrap();969970assert_eq!(971result,972ControlFlow::Return(smallvec![DataValue::F32(Ieee32::with_float(1.0))])973)974}975976#[test]977fn misaligned_store_traps() {978let code = "979function %test() {980ss0 = explicit_slot 16981982block0:983v0 = stack_addr.i64 ss0984v1 = iconst.i64 1985store.i64 aligned v1, v0+2986return987}";988989let func = parse_functions(code).unwrap().into_iter().next().unwrap();990let mut env = FunctionStore::default();991env.add(func.name.to_string(), &func);992let state = InterpreterState::default().with_function_store(env);993let trap = Interpreter::new(state).call_by_name("%test", &[]).unwrap();994995assert_eq!(trap, ControlFlow::Trap(CraneliftTrap::HeapMisaligned));996}997998#[test]999fn misaligned_load_traps() {1000let code = "1001function %test() {1002ss0 = explicit_slot 1610031004block0:1005v0 = stack_addr.i64 ss01006v1 = iconst.i64 11007store.i64 aligned v1, v01008v2 = load.i64 aligned v0+21009return1010}";10111012let func = parse_functions(code).unwrap().into_iter().next().unwrap();1013let mut env = FunctionStore::default();1014env.add(func.name.to_string(), &func);1015let state = InterpreterState::default().with_function_store(env);1016let trap = Interpreter::new(state).call_by_name("%test", &[]).unwrap();10171018assert_eq!(trap, ControlFlow::Trap(CraneliftTrap::HeapMisaligned));1019}10201021// When a trap occurs in a function called by another function, the trap was not being propagated1022// correctly. Instead the interpterer panicked with a invalid control flow state.1023// See this issue for more details: https://github.com/bytecodealliance/wasmtime/issues/61551024#[test]1025fn trap_across_call_propagates_correctly() {1026let code = "1027function %u2() -> f32 system_v {1028ss0 = explicit_slot 691029ss1 = explicit_slot 691030ss2 = explicit_slot 6910311032block0:1033v0 = f32const -0x1.434342p-601034v1 = stack_addr.i64 ss2+241035store notrap aligned v0, v11036return v01037}10381039function %u1() -> f32 system_v {1040sig0 = () -> f32 system_v1041fn0 = colocated %u2 sig010421043block0:1044v57 = call fn0()1045return v571046}";10471048let mut env = FunctionStore::default();10491050let funcs = parse_functions(code).unwrap();1051for func in &funcs {1052env.add(func.name.to_string(), func);1053}10541055let state = InterpreterState::default().with_function_store(env);1056let trap = Interpreter::new(state).call_by_name("%u1", &[]).unwrap();10571058// Ensure that the correct trap was propagated.1059assert_eq!(trap, ControlFlow::Trap(CraneliftTrap::HeapMisaligned));1060}1061}106210631064