Wasm function can only trap (#29)

* Introduce Trap struct.

* get_local can't fail.

* Add MemoryOutOfBounds trap.

* from_little_endian use slice instead of vec.

* MemoryAccessOutOfBounds for mem get and set.

* from_little_endian conversion can't fail.

* call_indirect traps.

* DivisionByZero and InvalidConversionToInt

* Use traps in value to convey an error

* select: int condition on stack top

* if: int condition on stack top

* Assert pops.

* Another protions of assert pops

* Introduce ValueStack

Also, hide FunctionContext and remove some stale code

* Traps in execution

* Make it compile.

* Check args before invoke.

* Document RuntimeArgs.

* Update host.rs

* Add rustdoc for Trap.
This commit is contained in:
Sergey Pepyakin 2018-02-01 14:59:21 +03:00 committed by Nikolay Volf
parent 1d87fb09dc
commit c96735d6d6
11 changed files with 764 additions and 567 deletions

View File

@ -8,7 +8,7 @@ use wasmi::{
Error as InterpreterError, ModuleInstance, ModuleRef, Error as InterpreterError, ModuleInstance, ModuleRef,
Externals, RuntimeValue, FuncRef, ModuleImportResolver, Externals, RuntimeValue, FuncRef, ModuleImportResolver,
FuncInstance, HostError, ImportsBuilder, Signature, ValueType, FuncInstance, HostError, ImportsBuilder, Signature, ValueType,
RuntimeArgs, RuntimeArgs, Trap,
}; };
#[derive(Debug)] #[derive(Debug)]
@ -149,15 +149,15 @@ impl<'a> Externals for Runtime<'a> {
&mut self, &mut self,
index: usize, index: usize,
args: RuntimeArgs, args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, InterpreterError> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
SET_FUNC_INDEX => { SET_FUNC_INDEX => {
let idx: i32 = args.nth(0)?; let idx: i32 = args.nth(0);
self.game.set(idx, self.player)?; self.game.set(idx, self.player)?;
Ok(None) Ok(None)
} }
GET_FUNC_INDEX => { GET_FUNC_INDEX => {
let idx: i32 = args.nth(0)?; let idx: i32 = args.nth(0);
let val: i32 = tictactoe::Player::into_i32(self.game.get(idx)?); let val: i32 = tictactoe::Player::into_i32(self.game.get(idx)?);
Ok(Some(val.into())) Ok(Some(val.into()))
} }

View File

@ -13,13 +13,14 @@ use wasmi::{
GlobalInstance, GlobalRef, ImportResolver, ImportsBuilder, GlobalInstance, GlobalRef, ImportResolver, ImportsBuilder,
MemoryInstance, MemoryRef, ModuleImportResolver, ModuleInstance, MemoryInstance, MemoryRef, ModuleImportResolver, ModuleInstance,
ModuleRef, RuntimeValue, TableInstance, TableRef, ValueType, ModuleRef, RuntimeValue, TableInstance, TableRef, ValueType,
Module, Signature, MemoryDescriptor, Module, Signature, MemoryDescriptor, Trap,
TableDescriptor, GlobalDescriptor, FuncInstance, RuntimeArgs, TableDescriptor, GlobalDescriptor, FuncInstance, RuntimeArgs,
}; };
#[derive(Debug)] #[derive(Debug)]
enum Error { enum Error {
Load(String), Load(String),
Start(Trap),
Interpreter(InterpreterError), Interpreter(InterpreterError),
} }
@ -58,7 +59,7 @@ impl Externals for SpecModule {
&mut self, &mut self,
index: usize, index: usize,
args: RuntimeArgs, args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, InterpreterError> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
PRINT_FUNC_INDEX => { PRINT_FUNC_INDEX => {
println!("print: {:?}", args); println!("print: {:?}", args);
@ -243,7 +244,9 @@ fn try_load(
) -> Result<(), Error> { ) -> Result<(), Error> {
let module = try_load_module(base_dir, module_path)?; let module = try_load_module(base_dir, module_path)?;
let instance = ModuleInstance::new(&module, &ImportsBuilder::default())?; let instance = ModuleInstance::new(&module, &ImportsBuilder::default())?;
instance.run_start(spec_driver.spec_module())?; instance
.run_start(spec_driver.spec_module())
.map_err(|trap| Error::Start(trap))?;
Ok(()) Ok(())
} }

View File

@ -28,13 +28,6 @@ pub struct StackWithLimit<T> where T: Clone {
} }
impl<T> StackWithLimit<T> where T: Clone { impl<T> StackWithLimit<T> where T: Clone {
pub fn with_data<D: IntoIterator<Item=T>>(data: D, limit: usize) -> Self {
StackWithLimit {
values: data.into_iter().collect(),
limit: limit
}
}
pub fn with_limit(limit: usize) -> Self { pub fn with_limit(limit: usize) -> Self {
StackWithLimit { StackWithLimit {
values: VecDeque::new(), values: VecDeque::new(),

View File

@ -2,13 +2,11 @@ use std::rc::{Rc, Weak};
use std::fmt; use std::fmt;
use std::collections::HashMap; use std::collections::HashMap;
use parity_wasm::elements::{Local, Opcodes}; use parity_wasm::elements::{Local, Opcodes};
use {Error, Signature}; use {Trap, Signature};
use host::Externals; use host::Externals;
use runner::{prepare_function_args, FunctionContext, Interpreter}; use runner::{check_function_args, Interpreter};
use value::RuntimeValue; use value::RuntimeValue;
use module::ModuleInstance; use module::ModuleInstance;
use common::stack::StackWithLimit;
use common::{DEFAULT_FRAME_STACK_LIMIT, DEFAULT_VALUE_STACK_LIMIT};
/// Reference to a function (See [`FuncInstance`] for details). /// Reference to a function (See [`FuncInstance`] for details).
/// ///
@ -134,37 +132,16 @@ impl FuncInstance {
func: &FuncRef, func: &FuncRef,
args: &[RuntimeValue], args: &[RuntimeValue],
externals: &mut E, externals: &mut E,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
enum InvokeKind<'a> { debug_assert!(check_function_args(func.signature(), &args).is_ok());
Internal(FunctionContext), match *func.as_internal() {
Host(usize, &'a [RuntimeValue]), FuncInstanceInternal::Internal { .. } => {
} let mut interpreter = Interpreter::new(externals);
interpreter.start_execution(func, args)
let result = match *func.as_internal() {
FuncInstanceInternal::Internal { ref signature, .. } => {
let mut stack =
StackWithLimit::with_data(args.into_iter().cloned(), DEFAULT_VALUE_STACK_LIMIT);
let args = prepare_function_args(signature, &mut stack)?;
let context = FunctionContext::new(
func.clone(),
DEFAULT_VALUE_STACK_LIMIT,
DEFAULT_FRAME_STACK_LIMIT,
signature,
args,
);
InvokeKind::Internal(context)
} }
FuncInstanceInternal::Host { ref host_func_index, .. } => { FuncInstanceInternal::Host { ref host_func_index, .. } => {
InvokeKind::Host(*host_func_index, &*args) externals.invoke_index(*host_func_index, args.into())
} }
};
match result {
InvokeKind::Internal(ctx) => {
let mut interpreter = Interpreter::new(externals);
interpreter.run_function(ctx)
}
InvokeKind::Host(host_func, args) => externals.invoke_index(host_func, args.into()),
} }
} }
} }

View File

@ -1,6 +1,6 @@
use std::any::TypeId; use std::any::TypeId;
use value::{RuntimeValue, TryInto}; use value::{RuntimeValue, TryInto};
use Error; use {Error, Trap};
/// Safe wrapper for list of arguments. /// Safe wrapper for list of arguments.
#[derive(Debug)] #[derive(Debug)]
@ -13,19 +13,37 @@ impl<'a> From<&'a [RuntimeValue]> for RuntimeArgs<'a> {
} }
impl<'a> RuntimeArgs<'a> { impl<'a> RuntimeArgs<'a> {
/// Extract argument by index `idx` returning error if cast is invalid or not enough arguments /// Extract argument by index `idx`.
pub fn nth<T>(&self, idx: usize) -> Result<T, Error> where RuntimeValue: TryInto<T, Error> { ///
Ok(self.nth_value(idx)?.try_into().map_err(|_| Error::Value("Invalid argument cast".to_owned()))?) /// # Errors
///
/// Returns `Err` if cast is invalid or not enough arguments.
pub fn nth_checked<T>(&self, idx: usize) -> Result<T, Error> where RuntimeValue: TryInto<T, Error> {
Ok(self.nth_value_checked(idx)?.try_into().map_err(|_| Error::Value("Invalid argument cast".to_owned()))?)
} }
/// Extract argument as a runtime value by index `idx` returning error is not enough arguments /// Extract argument as a [`RuntimeValue`] by index `idx`.
pub fn nth_value(&self, idx: usize) -> Result<RuntimeValue, Error> { ///
/// # Errors
///
/// Returns `Err` if this list has not enough arguments.
pub fn nth_value_checked(&self, idx: usize) -> Result<RuntimeValue, Error> {
if self.0.len() <= idx { if self.0.len() <= idx {
return Err(Error::Value("Invalid argument index".to_owned())); return Err(Error::Value("Invalid argument index".to_owned()));
} }
Ok(self.0[idx]) Ok(self.0[idx])
} }
/// Extract argument by index `idx`.
///
/// # Panics
///
/// Panics if cast is invalid or not enough arguments.
pub fn nth<T>(&self, idx: usize) -> T where RuntimeValue: TryInto<T, Error> {
let value = self.nth_value_checked(idx).expect("Invalid argument index");
value.try_into().expect("Unexpected argument type")
}
/// Total number of arguments /// Total number of arguments
pub fn len(&self) -> usize { pub fn len(&self) -> usize {
self.0.len() self.0.len()
@ -104,7 +122,7 @@ impl HostError {
/// ```rust /// ```rust
/// use wasmi::{ /// use wasmi::{
/// Externals, RuntimeValue, RuntimeArgs, Error, ModuleImportResolver, /// Externals, RuntimeValue, RuntimeArgs, Error, ModuleImportResolver,
/// FuncRef, ValueType, Signature, FuncInstance /// FuncRef, ValueType, Signature, FuncInstance, Trap,
/// }; /// };
/// ///
/// struct HostExternals { /// struct HostExternals {
@ -118,11 +136,11 @@ impl HostError {
/// &mut self, /// &mut self,
/// index: usize, /// index: usize,
/// args: RuntimeArgs, /// args: RuntimeArgs,
/// ) -> Result<Option<RuntimeValue>, Error> { /// ) -> Result<Option<RuntimeValue>, Trap> {
/// match index { /// match index {
/// ADD_FUNC_INDEX => { /// ADD_FUNC_INDEX => {
/// let a: u32 = args.nth(0).unwrap(); /// let a: u32 = args.nth(0);
/// let b: u32 = args.nth(1).unwrap(); /// let b: u32 = args.nth(1);
/// let result = a + b; /// let result = a + b;
/// ///
/// Ok(Some(RuntimeValue::I32(result as i32))) /// Ok(Some(RuntimeValue::I32(result as i32)))
@ -173,7 +191,7 @@ pub trait Externals {
&mut self, &mut self,
index: usize, index: usize,
args: RuntimeArgs, args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error>; ) -> Result<Option<RuntimeValue>, Trap>;
} }
/// Implementation of [`Externals`] that just traps on [`invoke_index`]. /// Implementation of [`Externals`] that just traps on [`invoke_index`].
@ -187,8 +205,8 @@ impl Externals for NopExternals {
&mut self, &mut self,
_index: usize, _index: usize,
_args: RuntimeArgs, _args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
Err(Error::Trap("invoke index on no-op externals".into())) Err(Trap::Unreachable)
} }
} }
@ -201,14 +219,14 @@ mod tests {
#[test] #[test]
fn i32_runtime_args() { fn i32_runtime_args() {
let args: RuntimeArgs = (&[RuntimeValue::I32(0)][..]).into(); let args: RuntimeArgs = (&[RuntimeValue::I32(0)][..]).into();
let val: i32 = args.nth(0).unwrap(); let val: i32 = args.nth_checked(0).unwrap();
assert_eq!(val, 0); assert_eq!(val, 0);
} }
#[test] #[test]
fn i64_invalid_arg_cast() { fn i64_invalid_arg_cast() {
let args: RuntimeArgs = (&[RuntimeValue::I64(90534534545322)][..]).into(); let args: RuntimeArgs = (&[RuntimeValue::I64(90534534545322)][..]).into();
assert!(args.nth::<i32>(0).is_err()); assert!(args.nth_checked::<i32>(0).is_err());
} }
// Tests that `HostError` trait is object safe. // Tests that `HostError` trait is object safe.

View File

@ -106,6 +106,78 @@ use std::fmt;
use std::error; use std::error;
use std::collections::HashMap; use std::collections::HashMap;
/// Error type which can happen at execution time.
///
/// Under some conditions, wasm execution may produce a `Trap`, which immediately aborts execution.
/// Traps can't be handled by WebAssembly code, but are reported to the embedder.
#[derive(Debug)]
pub enum Trap {
/// Wasm code executed `unreachable` opcode.
///
/// `unreachable` is a special opcode which always traps upon execution.
/// This opcode have a similar purpose as `ud2` in x86.
Unreachable,
/// Attempt to load or store at the address which
/// lies outside of bounds of the memory.
///
/// Since addresses are interpreted as unsigned integers, out of bounds access
/// can't happen with negative addresses (i.e. they will always wrap).
MemoryAccessOutOfBounds,
/// Attempt to access table element at index which
/// lies outside of bounds.
///
/// This typically can happen when `call_indirect` is executed
/// with index that lies out of bounds.
///
/// Since indexes are interpreted as unsinged integers, out of bounds access
/// can't happen with negative indexes (i.e. they will always wrap).
TableAccessOutOfBounds,
/// Attempt to access table element which is uninitialized (i.e. `None`).
///
/// This typically can happen when `call_indirect` is executed.
ElemUninitialized,
/// Attempt to `call_indirect` function with mismatched [signature][`Signature`].
///
/// `call_indirect` always specifies the expected signature of function.
/// If `call_indirect` is executed with index that points on function with
/// signature different that is expected by this `call_indirect`, this trap is raised.
///
/// [`Signature`]: struct.Signature.html
ElemSignatureMismatch,
/// Attempt to divide by zero.
///
/// This trap typically can happen if `div` or `rem` is executed with
/// zero as divider.
DivisionByZero,
/// Attempt to make a conversion to an int failed.
///
/// This can happen when:
///
/// - trying to do signed division (or get the remainder) -2<sup>N-1</sup> over -1. This is
/// because the result +2<sup>N-1</sup> isn't representable as a N-bit signed integer.
/// - trying to truncate NaNs, infinity, or value for which the result is out of range into an integer.
InvalidConversionToInt,
/// Stack overflow.
///
/// This is likely caused by some infinite or very deep recursion.
/// Extensive inlining might also be the cause of stack overflow.
StackOverflow,
/// Error specified by the host.
///
/// Typically returned from an implementation of [`Externals`].
///
/// [`Externals`]: trait.Externals.html
Host(Box<host::HostError>),
}
/// Internal interpreter error. /// Internal interpreter error.
#[derive(Debug)] #[derive(Debug)]
pub enum Error { pub enum Error {
@ -127,7 +199,7 @@ pub enum Error {
/// Value-level error. /// Value-level error.
Value(String), Value(String),
/// Trap. /// Trap.
Trap(String), Trap(Trap),
/// Custom embedder error. /// Custom embedder error.
Host(Box<host::HostError>), Host(Box<host::HostError>),
} }
@ -143,7 +215,7 @@ impl Into<String> for Error {
Error::Global(s) => s, Error::Global(s) => s,
Error::Stack(s) => s, Error::Stack(s) => s,
Error::Value(s) => s, Error::Value(s) => s,
Error::Trap(s) => format!("trap: {}", s), Error::Trap(s) => format!("trap: {:?}", s),
Error::Host(e) => format!("user: {}", e), Error::Host(e) => format!("user: {}", e),
} }
} }
@ -160,7 +232,7 @@ impl fmt::Display for Error {
Error::Global(ref s) => write!(f, "Global: {}", s), Error::Global(ref s) => write!(f, "Global: {}", s),
Error::Stack(ref s) => write!(f, "Stack: {}", s), Error::Stack(ref s) => write!(f, "Stack: {}", s),
Error::Value(ref s) => write!(f, "Value: {}", s), Error::Value(ref s) => write!(f, "Value: {}", s),
Error::Trap(ref s) => write!(f, "Trap: {}", s), Error::Trap(ref s) => write!(f, "Trap: {:?}", s),
Error::Host(ref e) => write!(f, "User: {}", e), Error::Host(ref e) => write!(f, "User: {}", e),
} }
} }
@ -179,7 +251,7 @@ impl error::Error for Error {
Error::Global(ref s) => s, Error::Global(ref s) => s,
Error::Stack(ref s) => s, Error::Stack(ref s) => s,
Error::Value(ref s) => s, Error::Value(ref s) => s,
Error::Trap(ref s) => s, Error::Trap(_) => "Trap",
Error::Host(_) => "Host error", Error::Host(_) => "Host error",
} }
} }
@ -192,6 +264,18 @@ impl<U> From<U> for Error where U: host::HostError + Sized {
} }
} }
impl<U> From<U> for Trap where U: host::HostError + Sized {
fn from(e: U) -> Self {
Trap::Host(Box::new(e))
}
}
impl From<Trap> for Error {
fn from(e: Trap) -> Error {
Error::Trap(e)
}
}
impl From<validation::Error> for Error { impl From<validation::Error> for Error {
fn from(e: validation::Error) -> Error { fn from(e: validation::Error) -> Error {
Error::Validation(e.to_string()) Error::Validation(e.to_string())

View File

@ -1,3 +1,5 @@
use runner::check_function_args;
use Trap;
use std::rc::Rc; use std::rc::Rc;
use std::cell::RefCell; use std::cell::RefCell;
use std::fmt; use std::fmt;
@ -575,7 +577,9 @@ impl ModuleInstance {
} }
}; };
check_function_args(func_instance.signature(), &args)?;
FuncInstance::invoke(&func_instance, args, externals) FuncInstance::invoke(&func_instance, args, externals)
.map_err(|t| Error::Trap(t))
} }
/// Find export by a name. /// Find export by a name.
@ -608,7 +612,7 @@ impl<'a> NotStartedModuleRef<'a> {
&self.instance &self.instance
} }
pub fn run_start<E: Externals>(self, state: &mut E) -> Result<ModuleRef, Error> { pub fn run_start<E: Externals>(self, state: &mut E) -> Result<ModuleRef, Trap> {
if let Some(start_fn_idx) = self.loaded_module.module().start_section() { if let Some(start_fn_idx) = self.loaded_module.module().start_section() {
let start_func = self.instance.func_by_index(start_fn_idx).expect( let start_func = self.instance.func_by_index(start_fn_idx).expect(
"Due to validation start function should exists", "Due to validation start function should exists",

File diff suppressed because it is too large Load Diff

View File

@ -107,7 +107,7 @@ impl TableInstance {
} }
/// Get the specific value in the table /// Get the specific value in the table
pub fn get(&self, offset: u32) -> Result<FuncRef, Error> { pub fn get(&self, offset: u32) -> Result<Option<FuncRef>, Error> {
let buffer = self.buffer.borrow(); let buffer = self.buffer.borrow();
let buffer_len = buffer.len(); let buffer_len = buffer.len();
let table_elem = buffer.get(offset as usize).cloned().ok_or_else(|| let table_elem = buffer.get(offset as usize).cloned().ok_or_else(||
@ -117,10 +117,7 @@ impl TableInstance {
buffer_len buffer_len
)), )),
)?; )?;
Ok(table_elem.ok_or_else(|| Error::Table(format!( Ok(table_elem)
"trying to read uninitialized element on index {}",
offset
)))?)
} }
/// Set the table element to the specified function. /// Set the table element to the specified function.

View File

@ -1,7 +1,7 @@
use { use {
Error, Signature, Externals, FuncInstance, FuncRef, HostError, ImportsBuilder, Error, Signature, Externals, FuncInstance, FuncRef, HostError, ImportsBuilder,
MemoryInstance, MemoryRef, TableInstance, TableRef, ModuleImportResolver, ModuleInstance, ModuleRef, MemoryInstance, MemoryRef, TableInstance, TableRef, ModuleImportResolver, ModuleInstance, ModuleRef,
RuntimeValue, RuntimeArgs, Module, TableDescriptor, MemoryDescriptor, RuntimeValue, RuntimeArgs, Module, TableDescriptor, MemoryDescriptor, Trap,
}; };
use types::ValueType; use types::ValueType;
use wabt::wat2wasm; use wabt::wat2wasm;
@ -79,23 +79,23 @@ impl Externals for TestHost {
&mut self, &mut self,
index: usize, index: usize,
args: RuntimeArgs, args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
SUB_FUNC_INDEX => { SUB_FUNC_INDEX => {
let a: i32 = args.nth(0)?; let a: i32 = args.nth(0);
let b: i32 = args.nth(1)?; let b: i32 = args.nth(1);
let result: RuntimeValue = (a - b).into(); let result: RuntimeValue = (a - b).into();
Ok(Some(result)) Ok(Some(result))
} }
ERR_FUNC_INDEX => { ERR_FUNC_INDEX => {
let error_code: u32 = args.nth(0)?; let error_code: u32 = args.nth(0);
let error = HostErrorWithCode { error_code }; let error = HostErrorWithCode { error_code };
Err(Error::Host(Box::new(error))) Err(Trap::Host(Box::new(error)))
} }
INC_MEM_FUNC_INDEX => { INC_MEM_FUNC_INDEX => {
let ptr: u32 = args.nth(0)?; let ptr: u32 = args.nth(0);
let memory = self.memory.as_ref().expect( let memory = self.memory.as_ref().expect(
"Function 'inc_mem' expects attached memory", "Function 'inc_mem' expects attached memory",
@ -108,7 +108,7 @@ impl Externals for TestHost {
Ok(None) Ok(None)
} }
GET_MEM_FUNC_INDEX => { GET_MEM_FUNC_INDEX => {
let ptr: u32 = args.nth(0)?; let ptr: u32 = args.nth(0);
let memory = self.memory.as_ref().expect( let memory = self.memory.as_ref().expect(
"Function 'get_mem' expects attached memory", "Function 'get_mem' expects attached memory",
@ -119,7 +119,7 @@ impl Externals for TestHost {
Ok(Some(RuntimeValue::I32(buf[0] as i32))) Ok(Some(RuntimeValue::I32(buf[0] as i32)))
} }
RECURSE_FUNC_INDEX => { RECURSE_FUNC_INDEX => {
let val = args.nth_value(0)?; let val = args.nth_value_checked(0).expect("Exactly one argument expected");
let instance = self.instance let instance = self.instance
.as_ref() .as_ref()
@ -131,7 +131,7 @@ impl Externals for TestHost {
.expect("expected to be Some"); .expect("expected to be Some");
if val.value_type() != result.value_type() { if val.value_type() != result.value_type() {
return Err(Error::Host(Box::new(HostErrorWithCode { error_code: 123 }))); return Err(Trap::Host(Box::new(HostErrorWithCode { error_code: 123 })));
} }
Ok(Some(result)) Ok(Some(result))
} }
@ -263,7 +263,7 @@ fn host_err() {
); );
let host_error: Box<HostError> = match error { let host_error: Box<HostError> = match error {
Error::Host(err) => err, Error::Trap(Trap::Host(err)) => err,
err => panic!("Unexpected error {:?}", err), err => panic!("Unexpected error {:?}", err),
}; };
@ -464,10 +464,10 @@ fn defer_providing_externals() {
&mut self, &mut self,
index: usize, index: usize,
args: RuntimeArgs, args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
INC_FUNC_INDEX => { INC_FUNC_INDEX => {
let a = args.nth::<u32>(0)?; let a = args.nth::<u32>(0);
*self.acc += a; *self.acc += a;
Ok(None) Ok(None)
} }
@ -528,7 +528,7 @@ fn two_envs_one_externals() {
&mut self, &mut self,
index: usize, index: usize,
_args: RuntimeArgs, _args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
PRIVILEGED_FUNC_INDEX => { PRIVILEGED_FUNC_INDEX => {
println!("privileged!"); println!("privileged!");
@ -648,7 +648,7 @@ fn dynamically_add_host_func() {
&mut self, &mut self,
index: usize, index: usize,
_args: RuntimeArgs, _args: RuntimeArgs,
) -> Result<Option<RuntimeValue>, Error> { ) -> Result<Option<RuntimeValue>, Trap> {
match index { match index {
ADD_FUNC_FUNC_INDEX => { ADD_FUNC_FUNC_INDEX => {
// Allocate indicies for the new function. // Allocate indicies for the new function.
@ -661,7 +661,8 @@ fn dynamically_add_host_func() {
Signature::new(&[][..], Some(ValueType::I32)), Signature::new(&[][..], Some(ValueType::I32)),
host_func_index as usize, host_func_index as usize,
); );
self.table.set(table_index, Some(added_func))?; self.table.set(table_index, Some(added_func))
.map_err(|_| Trap::TableAccessOutOfBounds)?;
Ok(Some(RuntimeValue::I32(table_index as i32))) Ok(Some(RuntimeValue::I32(table_index as i32)))
} }

View File

@ -1,7 +1,7 @@
use std::{i32, i64, u32, u64, f32}; use std::{i32, i64, u32, u64, f32};
use std::io; use std::io;
use byteorder::{LittleEndian, ReadBytesExt, WriteBytesExt}; use byteorder::{LittleEndian, ReadBytesExt, WriteBytesExt};
use Error; use {Error, Trap};
use parity_wasm::elements::ValueType; use parity_wasm::elements::ValueType;
@ -53,7 +53,7 @@ pub trait LittleEndianConvert where Self: Sized {
/// Convert to little endian buffer. /// Convert to little endian buffer.
fn into_little_endian(self) -> Vec<u8>; fn into_little_endian(self) -> Vec<u8>;
/// Convert from little endian buffer. /// Convert from little endian buffer.
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error>; fn from_little_endian(buffer: &[u8]) -> Result<Self, Error>;
} }
/// Arithmetic operations. /// Arithmetic operations.
@ -65,7 +65,7 @@ pub trait ArithmeticOps<T> {
/// Multiply two values. /// Multiply two values.
fn mul(self, other: T) -> T; fn mul(self, other: T) -> T;
/// Divide two values. /// Divide two values.
fn div(self, other: T) -> Result<T, Error>; fn div(self, other: T) -> Result<T, Trap>;
} }
/// Integer value. /// Integer value.
@ -81,7 +81,7 @@ pub trait Integer<T>: ArithmeticOps<T> {
/// Get right bit rotation result. /// Get right bit rotation result.
fn rotr(self, other: T) -> T; fn rotr(self, other: T) -> T;
/// Get division remainder. /// Get division remainder.
fn rem(self, other: T) -> Result<T, Error>; fn rem(self, other: T) -> Result<T, Trap>;
} }
/// Float-point value. /// Float-point value.
@ -263,19 +263,19 @@ impl_wrap_into!(f64, f32);
macro_rules! impl_try_truncate_into { macro_rules! impl_try_truncate_into {
($from: ident, $into: ident) => { ($from: ident, $into: ident) => {
impl TryTruncateInto<$into, Error> for $from { impl TryTruncateInto<$into, Trap> for $from {
fn try_truncate_into(self) -> Result<$into, Error> { fn try_truncate_into(self) -> Result<$into, Trap> {
// Casting from a float to an integer will round the float towards zero // Casting from a float to an integer will round the float towards zero
// NOTE: currently this will cause Undefined Behavior if the rounded value cannot be represented by the // NOTE: currently this will cause Undefined Behavior if the rounded value cannot be represented by the
// target integer type. This includes Inf and NaN. This is a bug and will be fixed. // target integer type. This includes Inf and NaN. This is a bug and will be fixed.
if self.is_nan() || self.is_infinite() { if self.is_nan() || self.is_infinite() {
return Err(Error::Value("invalid float value for this operation".into())); return Err(Trap::InvalidConversionToInt);
} }
// range check // range check
let result = self as $into; let result = self as $into;
if result as $from != self.trunc() { if result as $from != self.trunc() {
return Err(Error::Value("invalid float value for this operation".into())); return Err(Trap::InvalidConversionToInt);
} }
Ok(self as $into) Ok(self as $into)
@ -376,7 +376,7 @@ impl LittleEndianConvert for i8 {
vec![self as u8] vec![self as u8]
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
buffer.get(0) buffer.get(0)
.map(|v| *v as i8) .map(|v| *v as i8)
.ok_or_else(|| Error::Value("invalid little endian buffer".into())) .ok_or_else(|| Error::Value("invalid little endian buffer".into()))
@ -388,7 +388,7 @@ impl LittleEndianConvert for u8 {
vec![self] vec![self]
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
buffer.get(0) buffer.get(0)
.cloned() .cloned()
.ok_or_else(|| Error::Value("invalid little endian buffer".into())) .ok_or_else(|| Error::Value("invalid little endian buffer".into()))
@ -403,7 +403,7 @@ impl LittleEndianConvert for i16 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_i16::<LittleEndian>() io::Cursor::new(buffer).read_i16::<LittleEndian>()
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
} }
@ -417,7 +417,7 @@ impl LittleEndianConvert for u16 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_u16::<LittleEndian>() io::Cursor::new(buffer).read_u16::<LittleEndian>()
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
} }
@ -431,7 +431,7 @@ impl LittleEndianConvert for i32 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_i32::<LittleEndian>() io::Cursor::new(buffer).read_i32::<LittleEndian>()
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
} }
@ -445,7 +445,7 @@ impl LittleEndianConvert for u32 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_u32::<LittleEndian>() io::Cursor::new(buffer).read_u32::<LittleEndian>()
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
} }
@ -459,7 +459,7 @@ impl LittleEndianConvert for i64 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_i64::<LittleEndian>() io::Cursor::new(buffer).read_i64::<LittleEndian>()
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
} }
@ -473,7 +473,7 @@ impl LittleEndianConvert for f32 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_u32::<LittleEndian>() io::Cursor::new(buffer).read_u32::<LittleEndian>()
.map(f32_from_bits) .map(f32_from_bits)
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
@ -488,7 +488,7 @@ impl LittleEndianConvert for f64 {
vec vec
} }
fn from_little_endian(buffer: Vec<u8>) -> Result<Self, Error> { fn from_little_endian(buffer: &[u8]) -> Result<Self, Error> {
io::Cursor::new(buffer).read_u64::<LittleEndian>() io::Cursor::new(buffer).read_u64::<LittleEndian>()
.map(f64_from_bits) .map(f64_from_bits)
.map_err(|e| Error::Value(e.to_string())) .map_err(|e| Error::Value(e.to_string()))
@ -537,12 +537,14 @@ macro_rules! impl_integer_arithmetic_ops {
fn add(self, other: $type) -> $type { self.wrapping_add(other) } fn add(self, other: $type) -> $type { self.wrapping_add(other) }
fn sub(self, other: $type) -> $type { self.wrapping_sub(other) } fn sub(self, other: $type) -> $type { self.wrapping_sub(other) }
fn mul(self, other: $type) -> $type { self.wrapping_mul(other) } fn mul(self, other: $type) -> $type { self.wrapping_mul(other) }
fn div(self, other: $type) -> Result<$type, Error> { fn div(self, other: $type) -> Result<$type, Trap> {
if other == 0 { Err(Error::Value("Division by zero".to_owned())) } if other == 0 {
Err(Trap::DivisionByZero)
}
else { else {
let (result, overflow) = self.overflowing_div(other); let (result, overflow) = self.overflowing_div(other);
if overflow { if overflow {
return Err(Error::Value("Attempt to divide with overflow".to_owned())); Err(Trap::InvalidConversionToInt)
} else { } else {
Ok(result) Ok(result)
} }
@ -563,7 +565,7 @@ macro_rules! impl_float_arithmetic_ops {
fn add(self, other: $type) -> $type { self + other } fn add(self, other: $type) -> $type { self + other }
fn sub(self, other: $type) -> $type { self - other } fn sub(self, other: $type) -> $type { self - other }
fn mul(self, other: $type) -> $type { self * other } fn mul(self, other: $type) -> $type { self * other }
fn div(self, other: $type) -> Result<$type, Error> { Ok(self / other) } fn div(self, other: $type) -> Result<$type, Trap> { Ok(self / other) }
} }
} }
} }
@ -579,8 +581,8 @@ macro_rules! impl_integer {
fn count_ones(self) -> $type { self.count_ones() as $type } fn count_ones(self) -> $type { self.count_ones() as $type }
fn rotl(self, other: $type) -> $type { self.rotate_left(other as u32) } fn rotl(self, other: $type) -> $type { self.rotate_left(other as u32) }
fn rotr(self, other: $type) -> $type { self.rotate_right(other as u32) } fn rotr(self, other: $type) -> $type { self.rotate_right(other as u32) }
fn rem(self, other: $type) -> Result<$type, Error> { fn rem(self, other: $type) -> Result<$type, Trap> {
if other == 0 { Err(Error::Value("Division by zero".to_owned())) } if other == 0 { Err(Trap::DivisionByZero) }
else { Ok(self.wrapping_rem(other)) } else { Ok(self.wrapping_rem(other)) }
} }
} }