Merge pull request #2262 from AleoHQ/feat/function-inlining

[Feature] Function Inlining
This commit is contained in:
d0cd 2023-02-14 16:38:38 -08:00 committed by GitHub
commit b4f3568009
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
449 changed files with 3048 additions and 1588 deletions

View File

@ -17,8 +17,8 @@
pub mod annotation;
pub use annotation::*;
pub mod call_type;
pub use call_type::*;
pub mod variant;
pub use variant::*;
pub mod external;
pub use external::*;
@ -26,11 +26,11 @@ pub use external::*;
pub mod finalize;
pub use finalize::*;
pub mod function_input;
pub use function_input::*;
pub mod input;
pub use input::*;
pub mod function_output;
pub use function_output::*;
pub mod output;
pub use output::*;
pub mod mode;
pub use mode::*;
@ -47,7 +47,7 @@ pub struct Function {
/// Annotations on the function.
pub annotations: Vec<Annotation>,
/// Is this function a transition, inlined, or a regular function?.
pub call_type: CallType,
pub variant: Variant,
/// The function identifier, e.g., `foo` in `function foo(...) { ... }`.
pub identifier: Identifier,
/// The function's input parameters.
@ -77,7 +77,7 @@ impl Function {
#[allow(clippy::too_many_arguments)]
pub fn new(
annotations: Vec<Annotation>,
call_type: CallType,
variant: Variant,
identifier: Identifier,
input: Vec<Input>,
output: Vec<Output>,
@ -99,7 +99,7 @@ impl Function {
Function {
annotations,
call_type,
variant,
identifier,
input,
output,
@ -123,10 +123,10 @@ impl Function {
/// Private formatting method used for optimizing [fmt::Debug] and [fmt::Display] implementations.
///
fn format(&self, f: &mut fmt::Formatter) -> fmt::Result {
match self.call_type {
CallType::Inline => write!(f, "inline ")?,
CallType::Standard => write!(f, "function ")?,
CallType::Transition => write!(f, "transition ")?,
match self.variant {
Variant::Inline => write!(f, "inline ")?,
Variant::Standard => write!(f, "function ")?,
Variant::Transition => write!(f, "transition ")?,
}
write!(f, "{}", self.identifier)?;

View File

@ -16,12 +16,12 @@
use serde::{Deserialize, Serialize};
/// An enum declaring how the function is invoked.
/// Functions are always one of three variants.
/// A transition function is permitted the ability to manipulate records.
/// A regular function is not permitted to manipulate records.
/// An inline function is directly copied at the call site.
#[derive(Copy, Clone, Debug, Serialize, Deserialize, PartialEq, Eq)]
pub enum CallType {
pub enum Variant {
Inline,
Standard,
Transition,

View File

@ -378,7 +378,7 @@ pub trait ProgramReconstructor: StatementReconstructor {
fn reconstruct_function(&mut self, input: Function) -> Function {
Function {
annotations: input.annotations,
call_type: input.call_type,
variant: input.variant,
identifier: input.identifier,
input: input.input,
output: input.output,

View File

@ -192,14 +192,27 @@ impl<'a> Compiler<'a> {
}
/// Runs the flattening pass.
pub fn flattening_pass(&mut self, symbol_table: &SymbolTable, assigner: Assigner) -> Result<()> {
self.ast = Flattener::do_pass((std::mem::take(&mut self.ast), symbol_table, assigner))?;
pub fn flattening_pass(&mut self, symbol_table: &SymbolTable, assigner: Assigner) -> Result<Assigner> {
let (ast, assigner) = Flattener::do_pass((std::mem::take(&mut self.ast), symbol_table, assigner))?;
self.ast = ast;
if self.output_options.flattened_ast {
self.write_ast_to_json("flattened_ast.json")?;
}
Ok(())
Ok(assigner)
}
/// Runs the function inlining pass.
pub fn function_inlining_pass(&mut self, call_graph: &CallGraph, assigner: Assigner) -> Result<Assigner> {
let (ast, assigner) = FunctionInliner::do_pass((std::mem::take(&mut self.ast), call_graph, assigner))?;
self.ast = ast;
if self.output_options.inlined_ast {
self.write_ast_to_json("inlined_ast.json")?;
}
Ok(assigner)
}
/// Runs the compiler stages.
@ -213,7 +226,9 @@ impl<'a> Compiler<'a> {
// TODO: Make this pass optional.
let assigner = self.static_single_assignment_pass(&st)?;
self.flattening_pass(&st, assigner)?;
let assigner = self.flattening_pass(&st, assigner)?;
let _ = self.function_inlining_pass(&call_graph, assigner)?;
Ok((st, struct_graph, call_graph))
}

View File

@ -14,6 +14,8 @@
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
// NOTE: If compiler passes are made optional, pass preconditions and invariants may not necessarily hold true.
#[derive(Clone, Default)]
pub struct OutputOptions {
/// Whether spans are enabled in the output ASTs.
@ -28,4 +30,6 @@ pub struct OutputOptions {
pub ssa_ast: bool,
/// If enabled writes the AST after flattening.
pub flattened_ast: bool,
/// If enabled writes the AST after inlining.
pub inlined_ast: bool,
}

View File

@ -53,6 +53,7 @@ struct CompileOutput {
pub unrolled_ast: String,
pub ssa_ast: String,
pub flattened_ast: String,
pub inlined_ast: String,
pub bytecode: String,
}
@ -74,7 +75,7 @@ fn run_test(test: Test, handler: &Handler) -> Result<Value, ()> {
handler.extend_if_error(package.get_process().map_err(LeoError::Anyhow))?;
// Hash the ast files.
let (initial_ast, unrolled_ast, ssa_ast, flattened_ast) = hash_asts();
let (initial_ast, unrolled_ast, ssa_ast, flattened_ast, inlined_ast) = hash_asts();
// Clean up the output directory.
if fs::read_dir("/tmp/output").is_ok() {
@ -86,6 +87,7 @@ fn run_test(test: Test, handler: &Handler) -> Result<Value, ()> {
unrolled_ast,
ssa_ast,
flattened_ast,
inlined_ast,
bytecode: hash_content(&bytecode),
};
Ok(serde_yaml::to_value(final_output).expect("serialization failed"))

View File

@ -58,6 +58,7 @@ struct ExecuteOutput {
pub unrolled_ast: String,
pub ssa_ast: String,
pub flattened_ast: String,
pub inlined_ast: String,
pub bytecode: String,
pub results: BTreeMap<String, Vec<BTreeMap<String, String>>>,
}
@ -145,7 +146,7 @@ fn run_test(test: Test, handler: &Handler, err_buf: &BufferEmitter) -> Result<Va
}
// Hash the ast files.
let (initial_ast, unrolled_ast, ssa_ast, flattened_ast) = hash_asts();
let (initial_ast, unrolled_ast, ssa_ast, flattened_ast, inlined_ast) = hash_asts();
// Clean up the output directory.
if fs::read_dir("/tmp/output").is_ok() {
@ -157,6 +158,7 @@ fn run_test(test: Test, handler: &Handler, err_buf: &BufferEmitter) -> Result<Va
unrolled_ast,
ssa_ast,
flattened_ast,
inlined_ast,
bytecode: hash_content(&bytecode),
results,
};

View File

@ -39,13 +39,14 @@ pub type Network = Testnet3;
#[allow(unused)]
pub type Aleo = snarkvm::circuit::AleoV0;
pub fn hash_asts() -> (String, String, String, String) {
pub fn hash_asts() -> (String, String, String, String, String) {
let initial_ast = hash_file("/tmp/output/test.initial_ast.json");
let unrolled_ast = hash_file("/tmp/output/test.unrolled_ast.json");
let ssa_ast = hash_file("/tmp/output/test.ssa_ast.json");
let flattened_ast = hash_file("/tmp/output/test.flattened_ast.json");
let inlined_ast = hash_file("/tmp/output/test.inlined_ast.json");
(initial_ast, unrolled_ast, ssa_ast, flattened_ast)
(initial_ast, unrolled_ast, ssa_ast, flattened_ast, inlined_ast)
}
pub fn get_cwd_option(test: &Test) -> Option<PathBuf> {
@ -100,6 +101,7 @@ pub fn new_compiler(handler: &Handler, main_file_path: PathBuf) -> Compiler<'_>
unrolled_ast: true,
ssa_ast: true,
flattened_ast: true,
inlined_ast: true,
}),
)
}
@ -183,11 +185,16 @@ pub fn temp_dir() -> PathBuf {
pub fn compile_and_process<'a>(parsed: &'a mut Compiler<'a>) -> Result<String, LeoError> {
let st = parsed.symbol_table_pass()?;
let (st, struct_graph, call_graph) = parsed.type_checker_pass(st)?;
let st = parsed.loop_unrolling_pass(st)?;
let assigner = parsed.static_single_assignment_pass(&st)?;
parsed.flattening_pass(&st, assigner)?;
let assigner = parsed.flattening_pass(&st, assigner)?;
let _ = parsed.function_inlining_pass(&call_graph, assigner)?;
// Compile Leo program to bytecode.
let bytecode = CodeGenerator::do_pass((&parsed.ast, &st, &struct_graph, &call_graph))?;

View File

@ -156,7 +156,7 @@ impl ParserContext<'_> {
let (id, mapping) = self.parse_mapping()?;
mappings.insert(id, mapping);
}
Token::At | Token::Function | Token::Transition => {
Token::At | Token::Function | Token::Transition | Token::Inline => {
let (id, function) = self.parse_function()?;
functions.insert(id, function);
}
@ -172,6 +172,7 @@ impl ParserContext<'_> {
Token::At,
Token::Function,
Token::Transition,
Token::Inline,
],
)
.into())
@ -423,11 +424,12 @@ impl ParserContext<'_> {
while self.look_ahead(0, |t| &t.token) == &Token::At {
annotations.push(self.parse_annotation()?)
}
// Parse `<call_type> IDENT`, where `<call_type>` is `function` or `transition`.
let (call_type, start) = match self.token.token {
Token::Function => (CallType::Standard, self.expect(&Token::Function)?),
Token::Transition => (CallType::Transition, self.expect(&Token::Transition)?),
_ => self.unexpected("'function', 'transition'")?,
// Parse `<variant> IDENT`, where `<variant>` is `function`, `transition`, or `inline`.
let (variant, start) = match self.token.token {
Token::Inline => (Variant::Inline, self.expect(&Token::Inline)?),
Token::Function => (Variant::Standard, self.expect(&Token::Function)?),
Token::Transition => (Variant::Transition, self.expect(&Token::Transition)?),
_ => self.unexpected("'function', 'transition', or 'inline'")?,
};
let name = self.expect_identifier()?;
@ -489,7 +491,7 @@ impl ParserContext<'_> {
let span = start + block.span;
Ok((
name.name,
Function::new(annotations, call_type, name, inputs, output, block, finalize, span),
Function::new(annotations, variant, name, inputs, output, block, finalize, span),
))
}
}

View File

@ -425,9 +425,10 @@ impl Token {
"i64" => Token::I64,
"i128" => Token::I128,
"if" => Token::If,
"import" => Token::Import,
"in" => Token::In,
"increment" => Token::Increment,
"import" => Token::Import,
"inline" => Token::Inline,
"let" => Token::Let,
"leo" => Token::Leo,
"mapping" => Token::Mapping,

View File

@ -103,6 +103,7 @@ mod tests {
i8
if
in
inline
input
let
mut
@ -167,7 +168,7 @@ mod tests {
assert_eq!(
output,
r#""test" "test{}test" "test{}" "{}test" "test{" "test}" "test{test" "test}test" "te{{}}" test_ident 12345 address assert assert_eq assert_neq async bool const else false field finalize for function group i128 i64 i32 i16 i8 if in input let mut private program public return scalar self string struct test then transition true u128 u64 u32 u16 u8 console ! != && ( ) * ** + , - -> => _ . .. / : ; < <= = == > >= [ ] { { } } || ? @ // test
r#""test" "test{}test" "test{}" "{}test" "test{" "test}" "test{test" "test}test" "te{{}}" test_ident 12345 address assert assert_eq assert_neq async bool const else false field finalize for function group i128 i64 i32 i16 i8 if in inline input let mut private program public return scalar self string struct test then transition true u128 u64 u32 u16 u8 console ! != && ( ) * ** + , - -> => _ . .. / : ; < <= = == > >= [ ] { { } } || ? @ // test
/* test */ // "#
);
});

View File

@ -127,6 +127,7 @@ pub enum Token {
Import,
In,
Increment,
Inline,
Let,
Mapping,
Private,
@ -177,6 +178,7 @@ pub const KEYWORD_TOKENS: &[Token] = &[
Token::Import,
Token::In,
Token::Increment,
Token::Inline,
Token::Let,
Token::Mapping,
Token::Private,
@ -231,9 +233,10 @@ impl Token {
Token::I64 => sym::i64,
Token::I128 => sym::i128,
Token::If => sym::If,
Token::Import => sym::import,
Token::In => sym::In,
Token::Increment => sym::increment,
Token::Import => sym::import,
Token::Inline => sym::inline,
Token::Let => sym::Let,
Token::Leo => sym::leo,
Token::Mapping => sym::mapping,
@ -361,6 +364,7 @@ impl fmt::Display for Token {
Import => write!(f, "import"),
In => write!(f, "in"),
Increment => write!(f, "increment"),
Inline => write!(f, "inline"),
Let => write!(f, "let"),
Mapping => write!(f, "mapping"),
Private => write!(f, "private"),

View File

@ -16,7 +16,7 @@
use crate::CodeGenerator;
use leo_ast::{functions, CallType, Function, Mapping, Mode, Program, ProgramScope, Struct, Type};
use leo_ast::{functions, Function, Mapping, Mode, Program, ProgramScope, Struct, Type, Variant};
use indexmap::IndexMap;
use itertools::Itertools;
@ -94,7 +94,7 @@ impl<'a> CodeGenerator<'a> {
let function = program_scope.functions.get(&function_name).unwrap();
// Set the `is_transition_function` flag.
self.is_transition_function = matches!(function.call_type, CallType::Transition);
self.is_transition_function = matches!(function.variant, Variant::Transition);
let function_string = self.visit_function(function);

View File

@ -20,7 +20,7 @@ use indexmap::{IndexMap, IndexSet};
use std::fmt::Debug;
use std::hash::Hash;
/// An struct dependency graph.
/// A struct dependency graph.
pub type StructGraph = DiGraph<Symbol>;
/// A call graph.
@ -211,4 +211,11 @@ mod test {
let expected = Vec::from([1u32, 2, 4, 1]);
assert_eq!(cycle, expected);
}
#[test]
fn test_unconnected_graph() {
let graph = DiGraph::<u32>::new(IndexSet::from([1, 2, 3, 4, 5]));
check_post_order(&graph, &[1, 2, 3, 4, 5]);
}
}

View File

@ -23,5 +23,8 @@ pub use graph::*;
pub mod rename_table;
pub use rename_table::*;
pub mod replacer;
pub use replacer::*;
pub mod symbol_table;
pub use symbol_table::*;

View File

@ -0,0 +1,51 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use leo_ast::{Expression, ExpressionReconstructor, Identifier, ProgramReconstructor, StatementReconstructor};
/// A `Replacer` applies `replacer` to all `Identifier`s in an AST.
/// `Replacer`s are used to rename identifiers.
/// `Replacer`s are used to interpolate function arguments.
pub struct Replacer<F>
where
F: Fn(&Identifier) -> Expression,
{
replace: F,
}
impl<F> Replacer<F>
where
F: Fn(&Identifier) -> Expression,
{
pub fn new(replace: F) -> Self {
Self { replace }
}
}
impl<F> ExpressionReconstructor for Replacer<F>
where
F: Fn(&Identifier) -> Expression,
{
type AdditionalOutput = ();
fn reconstruct_identifier(&mut self, input: Identifier) -> (Expression, Self::AdditionalOutput) {
((self.replace)(&input), Default::default())
}
}
impl<F> StatementReconstructor for Replacer<F> where F: Fn(&Identifier) -> Expression {}
impl<F> ProgramReconstructor for Replacer<F> where F: Fn(&Identifier) -> Expression {}

View File

@ -14,7 +14,7 @@
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use leo_ast::{CallType, Function, Input, Type};
use leo_ast::{Function, Input, Type, Variant};
use leo_span::Span;
use crate::SymbolTable;
@ -36,7 +36,7 @@ pub struct FunctionSymbol {
/// The output type of the function.
pub(crate) output_type: Type,
/// Is this function a transition, inlined, or a regular function?.
pub call_type: CallType,
pub variant: Variant,
/// The `Span` associated with the function.
pub(crate) _span: Span,
/// The inputs to the function.
@ -50,7 +50,7 @@ impl SymbolTable {
FunctionSymbol {
id,
output_type: func.output_type.clone(),
call_type: func.call_type,
variant: func.variant,
_span: func.span,
input: func.input.clone(),
finalize: func.finalize.as_ref().map(|finalize| FinalizeData {

View File

@ -201,12 +201,12 @@ impl ExpressionReconstructor for Flattener<'_> {
statements.extend(stmts);
// Create and accumulate an intermediate assignment statement for the ternary expression corresponding to the struct member.
let (identifier, statement) = self.unique_simple_assign_statement(expression);
let (result, statement) = self.unique_simple_assign_statement(expression);
statements.push(statement);
StructVariableInitializer {
identifier,
expression: Some(Expression::Identifier(identifier)),
identifier: *identifier,
expression: Some(Expression::Identifier(result)),
}
})
.collect();
@ -297,12 +297,12 @@ impl ExpressionReconstructor for Flattener<'_> {
statements.extend(stmts);
// Create and accumulate an intermediate assignment statement for the ternary expression corresponding to the struct member.
let (identifier, statement) = self.unique_simple_assign_statement(expression);
let (result, statement) = self.unique_simple_assign_statement(expression);
statements.push(statement);
StructVariableInitializer {
identifier,
expression: Some(Expression::Identifier(identifier)),
identifier: *identifier,
expression: Some(Expression::Identifier(result)),
}
})
.collect();

View File

@ -69,7 +69,7 @@ impl ProgramReconstructor for Flattener<'_> {
Function {
annotations: function.annotations,
call_type: function.call_type,
variant: function.variant,
identifier: function.identifier,
input: function.input,
output: function.output,

View File

@ -27,7 +27,7 @@ use indexmap::IndexMap;
pub struct Flattener<'a> {
/// The symbol table associated with the program.
pub(crate) symbol_table: &'a SymbolTable,
/// An struct used to construct (unique) assignment statements.
/// A struct used to construct (unique) assignment statements.
pub(crate) assigner: Assigner,
/// The set of variables that are structs.
pub(crate) structs: IndexMap<Symbol, Symbol>,

View File

@ -66,12 +66,12 @@ use leo_errors::Result;
impl<'a> Pass for Flattener<'a> {
type Input = (Ast, &'a SymbolTable, Assigner);
type Output = Result<Ast>;
type Output = Result<(Ast, Assigner)>;
fn do_pass((ast, st, assigner): Self::Input) -> Self::Output {
let mut reconstructor = Flattener::new(st, assigner);
let program = reconstructor.reconstruct_program(ast.into_repr());
Ok(Ast::new(program))
Ok((Ast::new(program), reconstructor.assigner))
}
}

View File

@ -0,0 +1,149 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use crate::{Assigner, RenameTable};
use leo_ast::{
AssignStatement, ConditionalStatement, ConsoleStatement, DefinitionStatement, Expression, ExpressionReconstructor,
Identifier, IterationStatement, ProgramReconstructor, Statement, StatementReconstructor, StructExpression,
StructVariableInitializer,
};
use leo_span::Symbol;
// TODO: Generalize the functionality of this reconstructor to be used in other passes.
/// An `AssignmentRenamer` renames the left-hand side of all assignment statements in an AST node.
/// The new names are propagated to all following identifiers.
pub struct AssignmentRenamer {
pub assigner: Assigner,
pub rename_table: RenameTable,
pub is_lhs: bool,
}
impl AssignmentRenamer {
/// Initialize a new `AssignmentRenamer`.
pub fn new(assigner: Assigner) -> Self {
Self {
assigner,
rename_table: RenameTable::new(None),
is_lhs: false,
}
}
/// Load the internal rename table with a set of entries.
pub fn load(&mut self, entries: impl Iterator<Item = (Symbol, Symbol)>) {
for (key, value) in entries {
self.rename_table.update(key, value);
}
}
/// Clear the internal rename table.
pub fn clear(&mut self) {
self.rename_table = RenameTable::new(None);
}
}
impl ExpressionReconstructor for AssignmentRenamer {
type AdditionalOutput = ();
/// Rename the identifier if it is the left-hand side of an assignment, otherwise look up for a new name in the internal rename table.
fn reconstruct_identifier(&mut self, input: Identifier) -> (Expression, Self::AdditionalOutput) {
let name = match self.is_lhs {
// If consuming the left-hand side of an assignment, a new unique name is introduced.
true => {
let new_name = self.assigner.unique_symbol(input.name, "$");
self.rename_table.update(input.name, new_name);
new_name
}
// Otherwise, we look up the previous name in the `RenameTable`.
// Note that we do not panic if the identifier is not found in the rename table.
// Variables that do not exist in the rename table are ones that have been introduced during the SSA pass.
// These variables are never re-assigned, and will never have an entry in the rename-table.
false => *self.rename_table.lookup(input.name).unwrap_or(&input.name),
};
(
Expression::Identifier(Identifier { name, span: input.span }),
Default::default(),
)
}
/// Rename the variable initializers in the struct expression.
fn reconstruct_struct_init(&mut self, input: StructExpression) -> (Expression, Self::AdditionalOutput) {
(
Expression::Struct(StructExpression {
name: input.name,
members: input
.members
.into_iter()
.map(|member| StructVariableInitializer {
identifier: member.identifier,
expression: match member.expression {
Some(expression) => Some(self.reconstruct_expression(expression).0),
None => unreachable!(
"SSA guarantees that all struct members are always of the form `<id> : <expr>`."
),
},
})
.collect(),
span: input.span,
}),
Default::default(),
)
}
}
impl StatementReconstructor for AssignmentRenamer {
/// Rename the left-hand side of the assignment statement.
fn reconstruct_assign(&mut self, input: AssignStatement) -> (Statement, Self::AdditionalOutput) {
// First rename the right-hand-side of the assignment.
let value = self.reconstruct_expression(input.value).0;
// Then assign a new unique name to the left-hand-side of the assignment.
// Note that this order is necessary to ensure that the right-hand-side uses the correct name when consuming a complex assignment.
self.is_lhs = true;
let place = self.reconstruct_expression(input.place).0;
self.is_lhs = false;
(
Statement::Assign(Box::new(AssignStatement {
place,
value,
span: input.span,
})),
Default::default(),
)
}
/// Flattening removes conditional statements from the program.
fn reconstruct_conditional(&mut self, _: ConditionalStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`ConditionalStatement`s should not be in the AST at this phase of compilation.")
}
/// Parsing guarantees that console statements are not present in the program.
fn reconstruct_console(&mut self, _: ConsoleStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`ConsoleStatement`s should not be in the AST at this phase of compilation.")
}
/// Static single assignment replaces definition statements with assignment statements.
fn reconstruct_definition(&mut self, _: DefinitionStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`DefinitionStatement`s should not exist in the AST at this phase of compilation.")
}
/// Loop unrolling unrolls and removes iteration statements from the program.
fn reconstruct_iteration(&mut self, _: IterationStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`IterationStatement`s should not be in the AST at this phase of compilation.");
}
}
impl ProgramReconstructor for AssignmentRenamer {}

View File

@ -0,0 +1,42 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use crate::{Assigner, AssignmentRenamer, CallGraph};
use leo_ast::Function;
use leo_span::Symbol;
use indexmap::IndexMap;
pub struct FunctionInliner<'a> {
/// The call graph for the program.
pub(crate) call_graph: &'a CallGraph,
/// A wrapper around an Assigner used to create unique variable assignments.
pub(crate) assignment_renamer: AssignmentRenamer,
/// A map of reconstructed functions in the current program scope.
pub(crate) reconstructed_functions: IndexMap<Symbol, Function>,
}
impl<'a> FunctionInliner<'a> {
/// Initializes a new `FunctionInliner`.
pub fn new(call_graph: &'a CallGraph, assigner: Assigner) -> Self {
Self {
call_graph,
assignment_renamer: AssignmentRenamer::new(assigner),
reconstructed_functions: Default::default(),
}
}
}

View File

@ -0,0 +1,93 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use crate::{FunctionInliner, Replacer};
use leo_ast::{
CallExpression, Expression, ExpressionReconstructor, Identifier, ReturnStatement, Statement,
StatementReconstructor, UnitExpression, Variant,
};
use indexmap::IndexMap;
use itertools::Itertools;
impl ExpressionReconstructor for FunctionInliner<'_> {
type AdditionalOutput = Vec<Statement>;
fn reconstruct_call(&mut self, input: CallExpression) -> (Expression, Self::AdditionalOutput) {
// Get the name of the callee function.
let function_name = match *input.function {
Expression::Identifier(identifier) => identifier.name,
_ => unreachable!("Parser guarantees that `input.function` is always an identifier."),
};
// Lookup the reconstructed callee function.
// Since this pass processes functions in post-order, the callee function is guaranteed to exist in `self.reconstructed_functions`
let callee = self.reconstructed_functions.get(&function_name).unwrap();
// Inline the callee function, if required, otherwise, return the call expression.
match callee.variant {
Variant::Transition | Variant::Standard => (Expression::Call(input), Default::default()),
Variant::Inline => {
// Construct a mapping from input variables of the callee function to arguments passed to the callee.
let parameter_to_argument = callee
.input
.iter()
.map(|input| input.identifier())
.zip_eq(input.arguments.into_iter())
.collect::<IndexMap<_, _>>();
// Initializer `self.assignment_renamer` with the function parameters.
self.assignment_renamer.load(
callee
.input
.iter()
.map(|input| (input.identifier().name, input.identifier().name)),
);
// Duplicate the body of the callee and create a unique assignment statement for each assignment in the body.
// This is necessary to ensure the inlined variables do not conflict with variables in the caller.
let unique_block = self.assignment_renamer.reconstruct_block(callee.block.clone()).0;
// Reset `self.assignment_renamer`.
self.assignment_renamer.clear();
// Replace each input variable with the appropriate parameter.
let replace = |identifier: &Identifier| match parameter_to_argument.get(identifier) {
Some(expression) => expression.clone(),
None => Expression::Identifier(*identifier),
};
let mut inlined_statements = Replacer::new(replace).reconstruct_block(unique_block).0.statements;
// If the inlined block returns a value, then use the value in place of the call expression, otherwise, use the unit expression.
let result = match inlined_statements.last() {
Some(Statement::Return(_)) => {
// Note that this unwrap is safe since we know that the last statement is a return statement.
match inlined_statements.pop().unwrap() {
Statement::Return(ReturnStatement { expression, .. }) => expression,
_ => unreachable!("This branch checks that the last statement is a return statement."),
}
}
_ => Expression::Unit(UnitExpression {
span: Default::default(),
}),
};
(result, inlined_statements)
}
}
}
}

View File

@ -0,0 +1,50 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use crate::FunctionInliner;
use leo_ast::{ProgramReconstructor, ProgramScope};
impl ProgramReconstructor for FunctionInliner<'_> {
fn reconstruct_program_scope(&mut self, mut input: ProgramScope) -> ProgramScope {
// Get the post-order ordering of the call graph.
// Note that the post-order always contains all nodes in the call graph.
// Note that the unwrap is safe since type checking guarantees that the call graph is acyclic.
let order = self.call_graph.post_order().unwrap();
// Reconstruct and accumulate each of the functions in post-order.
for function_name in order.into_iter() {
// Note that this unwrap is safe since type checking guarantees that all functions are declared.
let function = input.functions.remove(&function_name).unwrap();
// Reconstruct the function.
let reconstructed_function = self.reconstruct_function(function);
// Add the reconstructed function to the mapping.
self.reconstructed_functions
.insert(function_name, reconstructed_function);
}
// Note that this intentionally clears `self.reconstructed_functions` for the next program scope.
let functions = core::mem::take(&mut self.reconstructed_functions);
ProgramScope {
program_id: input.program_id,
structs: input.structs,
mappings: input.mappings,
functions,
span: input.span,
}
}
}

View File

@ -0,0 +1,114 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
use crate::FunctionInliner;
use leo_ast::{
AssignStatement, Block, ConditionalStatement, ConsoleStatement, DefinitionStatement, Expression,
ExpressionReconstructor, ExpressionStatement, IterationStatement, Statement, StatementReconstructor,
};
impl StatementReconstructor for FunctionInliner<'_> {
/// Reconstruct an assignment statement by inlining any function calls.
/// This function also segments tuple assignment statements into multiple assignment statements.
fn reconstruct_assign(&mut self, input: AssignStatement) -> (Statement, Self::AdditionalOutput) {
let (value, mut statements) = self.reconstruct_expression(input.value.clone());
match (input.place, value) {
// If the function call produces a tuple, we need to segment the tuple into multiple assignment statements.
(Expression::Tuple(left), Expression::Tuple(right)) if left.elements.len() == right.elements.len() => {
statements.extend(
left.elements
.into_iter()
.zip(right.elements.into_iter())
.map(|(lhs, rhs)| {
Statement::Assign(Box::new(AssignStatement {
place: lhs,
value: rhs,
span: Default::default(),
}))
}),
);
(Statement::dummy(Default::default()), statements)
}
(place, value) => (
Statement::Assign(Box::new(AssignStatement {
place,
value,
span: input.span,
})),
statements,
),
}
}
/// Reconstructs the statements inside a basic block, accumulating any statements produced by function inlining.
fn reconstruct_block(&mut self, block: Block) -> (Block, Self::AdditionalOutput) {
let mut statements = Vec::with_capacity(block.statements.len());
for statement in block.statements {
let (reconstructed_statement, additional_statements) = self.reconstruct_statement(statement);
statements.extend(additional_statements);
statements.push(reconstructed_statement);
}
(
Block {
span: block.span,
statements,
},
Default::default(),
)
}
/// Flattening removes conditional statements from the program.
fn reconstruct_conditional(&mut self, _: ConditionalStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`ConditionalStatement`s should not be in the AST at this phase of compilation.")
}
/// Parsing guarantees that console statements are not present in the program.
fn reconstruct_console(&mut self, _: ConsoleStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`ConsoleStatement`s should not be in the AST at this phase of compilation.")
}
/// Static single assignment replaces definition statements with assignment statements.
fn reconstruct_definition(&mut self, _: DefinitionStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`DefinitionStatement`s should not exist in the AST at this phase of compilation.")
}
/// Reconstructs expression statements by inlining any function calls.
fn reconstruct_expression_statement(&mut self, input: ExpressionStatement) -> (Statement, Self::AdditionalOutput) {
// Reconstruct the expression.
// Note that type checking guarantees that the expression is a function call.
let (expression, additional_statements) = self.reconstruct_expression(input.expression);
// If the resulting expression is a unit expression, return a dummy statement.
let statement = match expression {
Expression::Unit(_) => Statement::dummy(Default::default()),
_ => Statement::Expression(ExpressionStatement {
expression,
span: input.span,
}),
};
(statement, additional_statements)
}
/// Loop unrolling unrolls and removes iteration statements from the program.
fn reconstruct_iteration(&mut self, _: IterationStatement) -> (Statement, Self::AdditionalOutput) {
unreachable!("`IterationStatement`s should not be in the AST at this phase of compilation.");
}
}

View File

@ -0,0 +1,82 @@
// Copyright (C) 2019-2023 Aleo Systems Inc.
// This file is part of the Leo library.
// The Leo library is free software: you can redistribute it and/or modify
// it under the terms of the GNU General Public License as published by
// the Free Software Foundation, either version 3 of the License, or
// (at your option) any later version.
// The Leo library is distributed in the hope that it will be useful,
// but WITHOUT ANY WARRANTY; without even the implied warranty of
// MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
// GNU General Public License for more details.
// You should have received a copy of the GNU General Public License
// along with the Leo library. If not, see <https://www.gnu.org/licenses/>.
//! The Function Inlining pass traverses the AST and inlines function at their call site.
//! See https://en.wikipedia.org/wiki/Inline_expansion for more information.
//! The pass also reorders `Function`s in a reconstructed `ProgramScope` so that they are in a post-order of the call graph.
//! In other words, a callee function will appear before a caller function in the order.
//!
//! Consider the following flattened Leo code.
//! ```leo
//! function main(flag: u8, value: u8) -> u8 {
//! $var$0 = flag == 0u8;
//! $var$1 = foo(value);
//! value$2 = $var$1;
//! value$3 = $var$0 ? value$2 : value;
//! return value$3;
//! }
//!
//! inline foo(x: u8) -> u8 {
//! $var$4 = x * x;
//! return $var$4;
//! }
//! ```
//!
//! The inlining pass produces the following code.
//! ```leo
//! inline foo(x: u8) -> u8 {
//! $var$4 = x * x;
//! return $var$4;
//! }
//!
//! function main(flag: u8, value: u8) -> u8 {
//! $var$0 = flag == 0u8;
//! $var$4$5 = value * value;
//! $var$1 = $var$4$5;
//! value$2 = $var$1;
//! value$3 = $var$0 ? value$2 : value;
//! return value$3;
//! }
//! ```
pub mod assignment_renamer;
pub use assignment_renamer::*;
mod inline_expression;
mod inline_statement;
mod inline_program;
pub mod function_inliner;
pub use function_inliner::*;
use crate::{Assigner, CallGraph, Pass};
use leo_ast::{Ast, ProgramReconstructor};
use leo_errors::Result;
impl<'a> Pass for FunctionInliner<'a> {
type Input = (Ast, &'a CallGraph, Assigner);
type Output = Result<(Ast, Assigner)>;
fn do_pass((ast, call_graph, assigner): Self::Input) -> Self::Output {
let mut reconstructor = FunctionInliner::new(call_graph, assigner);
let program = reconstructor.reconstruct_program(ast.into_repr());
Ok((Ast::new(program), reconstructor.assignment_renamer.assigner))
}
}

View File

@ -26,6 +26,9 @@ pub use common::*;
pub mod flattening;
pub use flattening::*;
pub mod function_inlining;
pub use function_inlining::*;
pub mod loop_unrolling;
pub use self::loop_unrolling::*;

View File

@ -58,7 +58,7 @@ impl ProgramReconstructor for Unroller<'_> {
// Reconstruct the function block.
let reconstructed_function = Function {
annotations: function.annotations,
call_type: function.call_type,
variant: function.variant,
identifier: function.identifier,
input: function.input,
output: function.output,

View File

@ -110,7 +110,7 @@ impl FunctionConsumer for StaticSingleAssigner<'_> {
Function {
annotations: function.annotations,
call_type: function.call_type,
variant: function.variant,
identifier: function.identifier,
input: function.input,
output: function.output,

View File

@ -197,7 +197,7 @@ impl StatementConsumer for StaticSingleAssigner<'_> {
statements
}
/// Consumes the expressions in a `ConsoleStatement`, returning the list of simplified statements.
/// Parsing guarantees that console statements are not present in the program.
fn consume_console(&mut self, _: ConsoleStatement) -> Self::Output {
unreachable!("Parsing guarantees that console statements are not present in the program.")
}

View File

@ -23,7 +23,7 @@ pub struct StaticSingleAssigner<'a> {
pub(crate) rename_table: RenameTable,
/// A flag to determine whether or not the traversal is on the left-hand side of a definition or an assignment.
pub(crate) is_lhs: bool,
/// An struct used to construct (unique) assignment statements.
/// A struct used to construct (unique) assignment statements.
pub(crate) assigner: Assigner,
}

View File

@ -447,14 +447,17 @@ impl<'a> ExpressionVisitor<'a> for TypeChecker<'a> {
if let Some(func) = func {
// Check that the call is valid.
match self.is_transition_function {
// If the function is not a transition function, it cannot call any other functions.
false => {
self.emit_err(TypeCheckerError::cannot_invoke_call_from_standard_function(input.span));
// Note that this unwrap is safe since we always set the variant before traversing the body of the function.
match self.variant.unwrap() {
// If the function is not a transition function, it can only call "inline" functions.
Variant::Inline | Variant::Standard => {
if !matches!(func.variant, Variant::Inline) {
self.emit_err(TypeCheckerError::can_only_call_inline_function(input.span));
}
}
// If the function is a transition function, then check that the call is not to another local transition function.
true => {
if matches!(func.call_type, CallType::Transition) && input.external.is_none() {
Variant::Transition => {
if matches!(func.variant, Variant::Transition) && input.external.is_none() {
self.emit_err(TypeCheckerError::cannot_invoke_call_to_local_transition_function(
input.span,
));

View File

@ -70,7 +70,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
let mut transition_count = 0;
for function in input.functions.values() {
self.visit_function(function);
if matches!(function.call_type, CallType::Transition) {
if matches!(function.variant, Variant::Transition) {
transition_count += 1;
}
}
@ -181,7 +181,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
self.emit_err(TypeCheckerError::unknown_annotation(annotation, annotation.span))
}
self.is_transition_function = matches!(function.call_type, CallType::Transition);
self.variant = Some(function.variant);
// Lookup function metadata in the symbol table.
// Note that this unwrap is safe since function metadata is stored in a prior pass.
@ -216,13 +216,14 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
self.emit_err(TypeCheckerError::function_cannot_take_tuple_as_input(input_var.span()))
}
match self.is_transition_function {
// Note that this unwrap is safe since we assign to `self.variant` above.
match self.variant.unwrap() {
// If the function is a transition function, then check that the parameter mode is not a constant.
true if input_var.mode() == Mode::Const => self.emit_err(
Variant::Transition if input_var.mode() == Mode::Const => self.emit_err(
TypeCheckerError::transition_function_inputs_cannot_be_const(input_var.span()),
),
// If the function is not a transition function, then check that the parameters do not have an associated mode.
false if input_var.mode() != Mode::None => self.emit_err(
Variant::Standard | Variant::Inline if input_var.mode() != Mode::None => self.emit_err(
TypeCheckerError::regular_function_inputs_cannot_have_modes(input_var.span()),
),
_ => {} // Do nothing.
@ -248,7 +249,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
Output::External(external) => {
// If the function is not a transition function, then it cannot output a record.
// Note that an external output must always be a record.
if !self.is_transition_function {
if !matches!(function.variant, Variant::Transition) {
self.emit_err(TypeCheckerError::function_cannot_output_record(external.span()));
}
// Otherwise, do not type check external record function outputs.
@ -259,7 +260,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
self.assert_type_is_defined(&function_output.type_, function_output.span);
// If the function is not a transition function, then it cannot output a record.
if let Type::Identifier(identifier) = function_output.type_ {
if !self.is_transition_function
if !matches!(function.variant, Variant::Transition)
&& self
.symbol_table
.borrow()
@ -307,7 +308,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
self.has_finalize = false;
// Check that the function is a transition function.
if !self.is_transition_function {
if !matches!(function.variant, Variant::Transition) {
self.emit_err(TypeCheckerError::only_transition_functions_can_have_finalize(
finalize.span,
));
@ -393,7 +394,7 @@ impl<'a> ProgramVisitor<'a> for TypeChecker<'a> {
// Exit the function's scope.
self.exit_scope(function_index);
// Unset `is_transition_function` flag.
self.is_transition_function = false;
// Unset the `variant`.
self.variant = None;
}
}

View File

@ -16,7 +16,7 @@
use crate::{CallGraph, StructGraph, SymbolTable};
use leo_ast::{Identifier, IntegerType, Node, Type};
use leo_ast::{Identifier, IntegerType, Node, Type, Variant};
use leo_core::*;
use leo_errors::{emitter::Handler, TypeCheckerError};
use leo_span::{Span, Symbol};
@ -35,12 +35,13 @@ pub struct TypeChecker<'a> {
pub(crate) handler: &'a Handler,
/// The name of the function that we are currently traversing.
pub(crate) function: Option<Symbol>,
/// The variant of the function that we are currently traversing.
pub(crate) variant: Option<Variant>,
/// Whether or not the function that we are currently traversing has a return statement.
pub(crate) has_return: bool,
/// Whether or not the function that we are currently traversing invokes the finalize block.
pub(crate) has_finalize: bool,
/// Whether or not we are currently traversing a transition function.
pub(crate) is_transition_function: bool,
/// Whether or not we are currently traversing a finalize block.
pub(crate) is_finalize: bool,
/// Whether or not we are currently traversing an imported program.
@ -99,15 +100,16 @@ impl<'a> TypeChecker<'a> {
let function_names = symbol_table.functions.keys().cloned().collect();
// Note that the `struct_graph` and `call_graph` are initialized with their full node sets.
Self {
symbol_table: RefCell::new(symbol_table),
struct_graph: StructGraph::new(struct_names),
call_graph: CallGraph::new(function_names),
handler,
function: None,
variant: None,
has_return: false,
has_finalize: false,
is_transition_function: false,
is_finalize: false,
is_imported: false,
is_return: false,

View File

@ -199,6 +199,7 @@ symbols! {
In: "in",
import,
increment,
inline,
input,
Let: "let",
leo,

View File

@ -413,9 +413,9 @@ create_messages!(
}
@formatted
cannot_invoke_call_from_standard_function {
can_only_call_inline_function {
args: (),
msg: format!("Cannot call another function from a standard function."),
msg: format!("Only `inline` can be called from a `function` or `inline`."),
help: None,
}

View File

@ -58,6 +58,8 @@ pub struct BuildOptions {
pub enable_ssa_ast_snapshot: bool,
#[structopt(long, help = "Writes AST snapshot of the flattened AST.")]
pub enable_flattened_ast_snapshot: bool,
#[structopt(long, help = "Writes AST snapshot of the inlined AST.")]
pub enable_inlined_ast_snapshot: bool,
}
impl From<BuildOptions> for OutputOptions {
@ -69,6 +71,7 @@ impl From<BuildOptions> for OutputOptions {
unrolled_ast: options.enable_unrolled_ast_snapshot,
ssa_ast: options.enable_ssa_ast_snapshot,
flattened_ast: options.enable_flattened_ast_snapshot,
inlined_ast: options.enable_inlined_ast_snapshot,
};
if options.enable_all_ast_snapshots {
out_options.initial_input_ast = true;
@ -76,6 +79,7 @@ impl From<BuildOptions> for OutputOptions {
out_options.unrolled_ast = true;
out_options.ssa_ast = true;
out_options.flattened_ast = true;
out_options.inlined_ast = true;
}
out_options

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: da35eac4b0b82c923dc6075007bb615896efc249afffc57a197b129256f967e8
unrolled_ast: da35eac4b0b82c923dc6075007bb615896efc249afffc57a197b129256f967e8
ssa_ast: b257a4b1cf0e2aaa47df31a4845b9bb7d8293579c4b1ec4fd0f516d25c924dfa
flattened_ast: 1e009489c50036acad0df55be7cf779ec95bba86dcd40f3d08ff955a26424005
- initial_ast: af4bf950a53dee2547324d1d139661bc2f881f1a7e9941fe34a85b3745c6958d
unrolled_ast: af4bf950a53dee2547324d1d139661bc2f881f1a7e9941fe34a85b3745c6958d
ssa_ast: 71af510447b440ecf9517b244604dead0fb36905201d7205e1da396acd0de0fe
flattened_ast: 5b0842e447b4e1f92f4bcd22824ed5e12c51c8db145d1541763d10ad3dc1f37a
inlined_ast: 5b0842e447b4e1f92f4bcd22824ed5e12c51c8db145d1541763d10ad3dc1f37a
bytecode: eada90968195512a17847ed966d0bef43b7011c18ceee417a3cdb02a1190ca52

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: ab702ef192a0fcfae6024bd86f419dc6413e2655d593a66abf2abe58c169e4cd
unrolled_ast: ab702ef192a0fcfae6024bd86f419dc6413e2655d593a66abf2abe58c169e4cd
ssa_ast: 2a47e9b9cc25ee909648953549f2cfc1cc575b1787a88af5c77bc3ab487ea6ce
flattened_ast: d4c71cde2c8e2b908ca3f868b33eb2486424ddfc7aa000d177a27c7fe768841a
- initial_ast: b45e54f18036a13051f622f2d8230240aaee77231e39a1b7cdb196615fb4829e
unrolled_ast: b45e54f18036a13051f622f2d8230240aaee77231e39a1b7cdb196615fb4829e
ssa_ast: 0d44a08f0cace01d86fec36ea409e6424ff21f9ee8834f53062569af8c35579e
flattened_ast: c1b6954bff1ce18c0bb3be1cd6392a554a15989c90939c99e375221b1003e3b7
inlined_ast: c1b6954bff1ce18c0bb3be1cd6392a554a15989c90939c99e375221b1003e3b7
bytecode: b192f4b7f52da46a22cec3aec7e8c14b6e3fad7c40b9d0c0990255902fb596ef

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 2c04fdc69608512ca566b0eccf8d5e9fdd741b391a7937da4431293cfc3a046e
unrolled_ast: 2c04fdc69608512ca566b0eccf8d5e9fdd741b391a7937da4431293cfc3a046e
ssa_ast: 894efdfdb630a92cfa8a0a5101b681e15995e993b8c87d75dd738c8d7ce80d0c
flattened_ast: 139b39ec11943d222a18eaa61e504bc4aa46e7accc63c68755fafe861668f245
- initial_ast: ba44127bf5639ba2f19baafa22979156d0887ace3471003cd5fc924294760e65
unrolled_ast: ba44127bf5639ba2f19baafa22979156d0887ace3471003cd5fc924294760e65
ssa_ast: 31d5323289dd39759350c6e762e3902b6841ed099a621b556f20b4e7dbbe4af8
flattened_ast: f61d34c7740a32dfad5a2ddc2c48c502d6464258e16a79ab15c7a97f0724c0b9
inlined_ast: f61d34c7740a32dfad5a2ddc2c48c502d6464258e16a79ab15c7a97f0724c0b9
bytecode: 4903abf35d22e4264aae4bf26b908108d11d981d069c247793cea817dd8851a7

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 035627cd394d2b54892daeb2a84e3eff5ce02f06f96c688356a7e66585aa61ec
unrolled_ast: 035627cd394d2b54892daeb2a84e3eff5ce02f06f96c688356a7e66585aa61ec
ssa_ast: ca06bce101b6ec51b72c064357bafb094b7f4aafeec89e07e7ff58694eb032c8
flattened_ast: 9cf2f3f0e5391494304cedaf34d1bced7cc1f43599ee341d389fab7e050b574f
- initial_ast: 5287dfea1a34ec0beb84ab5d084e315e94ce0776495b278433c2d0d2bbdeca1d
unrolled_ast: 5287dfea1a34ec0beb84ab5d084e315e94ce0776495b278433c2d0d2bbdeca1d
ssa_ast: 3424c5f86b9123b7a6d71767e709d9e4ab979aa94f9e6be54824d3a8f002f1f6
flattened_ast: e54c1dbb5d765841bcc11fc9767f47528201e221e047991929b45e2f1688704e
inlined_ast: e54c1dbb5d765841bcc11fc9767f47528201e221e047991929b45e2f1688704e
bytecode: 5cbdf4a6a290f80540d2653153c57495eaf45432bc7ce44d52af2b5d0594951c

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 10f6791409c21336d6b777f66a6266b9b69fe761701a8a3c4a61ab59ef158db9
unrolled_ast: 10f6791409c21336d6b777f66a6266b9b69fe761701a8a3c4a61ab59ef158db9
ssa_ast: 1fa53964ef7d5eb8e8e7cd5af01a73fab123e423cb0b49e88d1c6de6ebf227d1
flattened_ast: 17e88106f92c0c25e20740e89304527e3044b636ebf31fdf0e34dacd86a0322a
- initial_ast: 800f83913bb57aac57e0574c67deda923814503eaa812fb82280a7ffd64f038f
unrolled_ast: 800f83913bb57aac57e0574c67deda923814503eaa812fb82280a7ffd64f038f
ssa_ast: e0015762d1fb228999fd2ef236fae8fcf8bd6e6bbd0ce37fad230a708ca063d2
flattened_ast: 4e7759584ade51a19ff90284e5ee1ac91af6dad5cd966568b708ead553a8a4bd
inlined_ast: 4e7759584ade51a19ff90284e5ee1ac91af6dad5cd966568b708ead553a8a4bd
bytecode: e3deaf24a91bcb77628f7af29d4ad6d0ba67215617d6cfe753168543123ce7d2

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: f5fd959d856025c2866f27b37e3711a49509b9a385a78d9186b1ea077ef5146d
unrolled_ast: f5fd959d856025c2866f27b37e3711a49509b9a385a78d9186b1ea077ef5146d
ssa_ast: a119ee86727dd82a8a82aaf459b14e13bc2fdc5177efdbced71bc0d5c3566f44
flattened_ast: 67e55340f8f5e864f376d989201bf85de15783ff43c438039bf4b5790e5b3081
- initial_ast: 080edd413ce668be563c96e2625ba86d935b529a25ff2d009a41c36d63e90867
unrolled_ast: 080edd413ce668be563c96e2625ba86d935b529a25ff2d009a41c36d63e90867
ssa_ast: 42925975f1f91dc72941e3c018d6c0595824086f50fa5e6398f21649a57c6661
flattened_ast: de891bab08a157399fdceeeccc7c3d4fd70cc3f75d1ca694a4fcd0344fdaac20
inlined_ast: de891bab08a157399fdceeeccc7c3d4fd70cc3f75d1ca694a4fcd0344fdaac20
bytecode: d0d3f79c32e6cb17c98afa2f1d4861d0f71d7f805a87712b3491ef0a9e1b4892

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 01d0412fda7ce54eec5f6ca193e4444ec47d2aa22a1425df65c6a2038a991fbb
unrolled_ast: 01d0412fda7ce54eec5f6ca193e4444ec47d2aa22a1425df65c6a2038a991fbb
ssa_ast: 5c2e90a07d8231a03f335bf762f5f660d3197f23b8f59b34d4b5319f9dd2cecb
flattened_ast: e68c35439afdde8d5587f8ae5289a422871b6942f3aba8b633f0549a2ca93b3a
- initial_ast: ad12566d0b8f3bef282b67823b427a74e56acbcc34acaa4f939097fb451ea7d9
unrolled_ast: ad12566d0b8f3bef282b67823b427a74e56acbcc34acaa4f939097fb451ea7d9
ssa_ast: 453e77387be9254ded9019b6c362721f766ebf5a5b2d3604e51ae81452fac4e8
flattened_ast: b39344c70e1a23869b236146ace198addf0801b348deedfb3e4ff1e3c4ace904
inlined_ast: b39344c70e1a23869b236146ace198addf0801b348deedfb3e4ff1e3c4ace904
bytecode: e742ac3b95a8971f2018963aba6d915ea53205c21443d0b11ad52a42ad443b97

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: bbfd6e25c73c1abe14cdd1dee14000551c1b4bdc5dc3e36c06870b7ac39ddc67
unrolled_ast: bbfd6e25c73c1abe14cdd1dee14000551c1b4bdc5dc3e36c06870b7ac39ddc67
ssa_ast: e3cf6c2a4fd04746714c7328cb24fa58ef15fb4e44c2be1c87e4116c62169ffa
flattened_ast: 633da269b8763b2eff9e443744d07cea72e592159171f02125bd6fbb504cdb71
- initial_ast: 5c885cc9e6d5df3602e09f7b53cb49ee4bca3a57d647044d4d321de32c4cdd90
unrolled_ast: 5c885cc9e6d5df3602e09f7b53cb49ee4bca3a57d647044d4d321de32c4cdd90
ssa_ast: 211f4122a90e6a117dc4fe2e7ca3c3e21bdc09a4c7992b212b6c34c283e896f6
flattened_ast: 84fd34b95b75f6d72b28164a9cb2ac80fa4149564c8c187b0ead1e14d2299a63
inlined_ast: 84fd34b95b75f6d72b28164a9cb2ac80fa4149564c8c187b0ead1e14d2299a63
bytecode: 1db874ad15d9bb70df7372ed3250cc6d0f65992e17788cd90c656ef1e1ceb63e

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 9a25a8b1eee1fb699e503df05bb0f46ed2ca10c9bc7fafbb7c14a75674b09811
unrolled_ast: 9a25a8b1eee1fb699e503df05bb0f46ed2ca10c9bc7fafbb7c14a75674b09811
ssa_ast: 0fc0e6501d2a700776813dcc44051d19c7b7453cb326d3531c172384f016db56
flattened_ast: a59c0aacc4aff5883c09976c56621841297b435e693d4f217ef97a85c19e73c7
- initial_ast: 3fda93baba12b9f280ffad75a662dfd16def0a7a1414de4cd29aa0e5afff85cc
unrolled_ast: 3fda93baba12b9f280ffad75a662dfd16def0a7a1414de4cd29aa0e5afff85cc
ssa_ast: 0ae9482705f95c26507f0040b972c76267a30eaa265f95764c758613d841932b
flattened_ast: 1e61c9d9ccdae7fb4aed4d7332538438839bef08a322f52fabcf46eac7bfc9c8
inlined_ast: 1e61c9d9ccdae7fb4aed4d7332538438839bef08a322f52fabcf46eac7bfc9c8
bytecode: 1a2170c46bb214eb8bedf2e98b58393ec0fa09051aeb52c3f59734a8da6ca5dc

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: d7b1adf3c0e7c05e1a58a2c76d2c0c675be0e04408adf3d30e7eab17cbc356bf
unrolled_ast: d7b1adf3c0e7c05e1a58a2c76d2c0c675be0e04408adf3d30e7eab17cbc356bf
ssa_ast: ecc229eb0204777ed540e09bc44119a400526ff9ca6ca8081fc5e2c511deda25
flattened_ast: 0d3dae7d567bcf36fd442b6658b9b217907935fc91c93a8b10bfe14225f57b71
- initial_ast: 3281347d18634932dba7502590f4ed0a45e15205fecdfb11846a1ac9de0a7c10
unrolled_ast: 3281347d18634932dba7502590f4ed0a45e15205fecdfb11846a1ac9de0a7c10
ssa_ast: f1ecffe7065e9782af5bf452b6ea547bfb5026a4c56e0c3105077c85ce196216
flattened_ast: cf5034c292702654bd282c10c8d1abafed8ed328f8e6cd0a01b286438809afd5
inlined_ast: cf5034c292702654bd282c10c8d1abafed8ed328f8e6cd0a01b286438809afd5
bytecode: e859520fd52dbdf69b14a3c3d9bad64bf6165084fb949912224eda3ccab9b638

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: b901b1b3790e751b788bc5ea0b8e006d3387dfa1924c85d32d4869b82e50ccde
unrolled_ast: b901b1b3790e751b788bc5ea0b8e006d3387dfa1924c85d32d4869b82e50ccde
ssa_ast: 496f4c109108408884fbe6b95c7f0a959fa636acae72e8ad8c792e4ec45892ed
flattened_ast: b5235d5705cd37b10f879e16db77cbc46b2b74f91285cef4a5611d4d02798c5d
- initial_ast: fd71e8fd519abdaf6d5abe40e167766800e80d7f9bb2523b13e2fa318598f9ef
unrolled_ast: fd71e8fd519abdaf6d5abe40e167766800e80d7f9bb2523b13e2fa318598f9ef
ssa_ast: d257cbd300cbeb2ad9d4246839296ce9a33751827bb7c2fddab7d07fdf077dd5
flattened_ast: 80ffd6de6e7429119f5098f8f0d8290fcfc1b5c3aa00f86a9153cb1471547af6
inlined_ast: 80ffd6de6e7429119f5098f8f0d8290fcfc1b5c3aa00f86a9153cb1471547af6
bytecode: fdc5659b97d4dbfea710ca848dcffa29bcd4da3a7a54739fb916e5292284a1a4

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: bcd5d9f7ba1b3d81d16310653b8531e39eeb6f6d35367dda89a77819f421d950
unrolled_ast: bcd5d9f7ba1b3d81d16310653b8531e39eeb6f6d35367dda89a77819f421d950
ssa_ast: d9f534bddd58e3fff333eadef39784a9c0f67f0b4a65fe7f7f4ae0fa4e808488
flattened_ast: f6e67ced6f8462f2cd655fa9b6a64a767ee22b7e123cb00d364ba0c7fea73e92
- initial_ast: c97efd0956a3c8d6a511b38d61f3f3bdd34d95ad2f78242d2816c723d1676997
unrolled_ast: c97efd0956a3c8d6a511b38d61f3f3bdd34d95ad2f78242d2816c723d1676997
ssa_ast: 0a690ca166cfd10c1b57d3df756032f10b003cc0d006bf27f41901b6af2ce95e
flattened_ast: 083a9af2e592de0c827b15230cd2307daae4b90e324e35714f474d50cbb59162
inlined_ast: 083a9af2e592de0c827b15230cd2307daae4b90e324e35714f474d50cbb59162
bytecode: 9006475518263541b3a855db6907377b638ef28f2a44caf4e26db7991c3b58ef

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 88ea51c2bf3abfa3c8810b8208c6e8428217a02543c584f890df6c1321ae2ec8
unrolled_ast: 88ea51c2bf3abfa3c8810b8208c6e8428217a02543c584f890df6c1321ae2ec8
ssa_ast: a4cd65095b0188f08dca7a5bee3c77cea0c038009e52bddac6a634641c447d54
flattened_ast: 50f2528808517ccb6f238e23c83f75666c391958386a86daac4929b5b1155602
- initial_ast: 4a8dc50cee056a918b491cba6cb3099695e1a76119b5aba7a31a820eb8a87d35
unrolled_ast: 4a8dc50cee056a918b491cba6cb3099695e1a76119b5aba7a31a820eb8a87d35
ssa_ast: 7237101f0bbd24e60a7b15718981024268b4db6d751b739cafb18c133b639713
flattened_ast: 7ae62df655e7b0450817324f4e45f680ec4e19978e2629c9be2b4f93aa0bd519
inlined_ast: 7ae62df655e7b0450817324f4e45f680ec4e19978e2629c9be2b4f93aa0bd519
bytecode: 65dcc91a4e07d98f73a0eb5b43d945f85859694417bd643b3ebba0d40494d001

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 1970d7b7a42e3b910ab4d910d41eb3f5ae96d509099ff0983d66dbeb1b7a8e54
unrolled_ast: 1970d7b7a42e3b910ab4d910d41eb3f5ae96d509099ff0983d66dbeb1b7a8e54
ssa_ast: 260c86f7c85b3cbddd66756cab69d3890b86951411472470f68f3ab9dd937a49
flattened_ast: ed477843f037999f6d1dae0c16ac75fb0c8e525ab44dee32e41fb4c1332d3b67
- initial_ast: eda5b3638cf36c8454468ab8eef2a6d0ebc61bb1a66b1cba1ae4f8315753e204
unrolled_ast: eda5b3638cf36c8454468ab8eef2a6d0ebc61bb1a66b1cba1ae4f8315753e204
ssa_ast: 3cc52869ffc8473ac37b723f6d1e996b06e7aa9c659f19720d15d0624ea59bf9
flattened_ast: c75e5f5d43494ea1c16ab46cddd1080bb0b56cc0ada1bc6b769fe2f4e5018e9e
inlined_ast: c75e5f5d43494ea1c16ab46cddd1080bb0b56cc0ada1bc6b769fe2f4e5018e9e
bytecode: 629677c0e48a743b2fc341de20c6b14ccc59d74c6ae08c317bdf597e0cc2f749

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 8916dd47b049835d250b12ccdcb00ab91fe7c0272671098871b5e8c926f93312
unrolled_ast: 8916dd47b049835d250b12ccdcb00ab91fe7c0272671098871b5e8c926f93312
ssa_ast: 0b7a62110d7d274f974a91b61c592f89edec6831b743008dfd038c464f4ec6d1
flattened_ast: 7b1400a9873fab71406b4330e864bcd8aecaa70eee0e4d96601ae0baad00f8a9
- initial_ast: 26911d8bca0e30e7bffd6abdde2ceef7485781af40e0707f1fb79d4461a7b9cd
unrolled_ast: 26911d8bca0e30e7bffd6abdde2ceef7485781af40e0707f1fb79d4461a7b9cd
ssa_ast: 44075583b8ebb223476e90b1b7bff6b1d6a65978ff2caae4ee2eb76f7c9aac6c
flattened_ast: 4ed800dd53990580a2a318e46932e29df53fce35a9f228750fc8bce335e3576e
inlined_ast: 4ed800dd53990580a2a318e46932e29df53fce35a9f228750fc8bce335e3576e
bytecode: a120b1e1d98948faf72473e55ec5ee1ea7ead4b7b5e4e79560053918dc1ff81b

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 5fe729d3d383d535452bd280076f64905a8f508849c19cf72ea538e93d8c25bf
unrolled_ast: 5fe729d3d383d535452bd280076f64905a8f508849c19cf72ea538e93d8c25bf
ssa_ast: 99239a050244e0741313b53301d7581117b126dca571361c568209e00db550b1
flattened_ast: 9fca40489b79b32cc320285c96c47f45fdbd14de1194aa7ff824508fb6e2b252
- initial_ast: c7ff62b1364b62c4350b33ceb0525ee7007cb3f969641d2eecc393efc6e5e5e4
unrolled_ast: c7ff62b1364b62c4350b33ceb0525ee7007cb3f969641d2eecc393efc6e5e5e4
ssa_ast: ca140fdf0902378ff3e3cc7160c8d4e843171bcaa9ae2389c478a2c73b983345
flattened_ast: 9be84865a93e187128e8c2fc1149abeb52f62b6704925ba25b49c02b16433418
inlined_ast: 9be84865a93e187128e8c2fc1149abeb52f62b6704925ba25b49c02b16433418
bytecode: 0098070069139200be104771fcb1807e52b331b50bab0dc82d3a16d451e4db97

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 22635a9f3ee2333185a07a9a3fb2a2f8c428cef4d00bd2072db527c49a665b0a
unrolled_ast: 22635a9f3ee2333185a07a9a3fb2a2f8c428cef4d00bd2072db527c49a665b0a
ssa_ast: d1b72ed1203491549451aade31456e4926166c6448cc2805e55fbf706594e7c3
flattened_ast: cc9c1479272915baf29b1de71b54341d06ff588a2a0816f310dbfeec10e96309
- initial_ast: 8c19f3a015926d3333017cb8d3311be2550d83e09c06c3c58176bd3c21716252
unrolled_ast: 8c19f3a015926d3333017cb8d3311be2550d83e09c06c3c58176bd3c21716252
ssa_ast: 992bfc5526a3849095f173452bc0046a0523c965fa367d6242e982613d0b23d2
flattened_ast: 178ad0e9092cd550d97360a8b8cef5f98e1c66c084a7df1e7ba930b12d83bb57
inlined_ast: 178ad0e9092cd550d97360a8b8cef5f98e1c66c084a7df1e7ba930b12d83bb57
bytecode: b9875b017a81226905d6fec1324bf41bc859bb4fca832de6b5309224ca343511

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 65a69d890ef0691d1cfb7396f45a8a8c8181fc3b958a537f8b0559512a91e00f
unrolled_ast: 65a69d890ef0691d1cfb7396f45a8a8c8181fc3b958a537f8b0559512a91e00f
ssa_ast: 2810947b50fb6df3b805a5a20b61ddb5812be742bdaf8775e4d73030f4745f44
flattened_ast: 53818a9ccc8a985d7990ad921ebfae74b68aed98cf016a37f01b9bd8127a186a
- initial_ast: 809cb6c8d6b69d3866a1b15c1d4caed220b0f63fdd83373b347ffd0dabbf873b
unrolled_ast: 809cb6c8d6b69d3866a1b15c1d4caed220b0f63fdd83373b347ffd0dabbf873b
ssa_ast: bceed5c819391bf5b52644c9acd3ad48485ad8c549f6e41592db3457664851c7
flattened_ast: 556b4b4c66f55375093c41eaf4e524e2b333b423672c538db60fd18a249068ad
inlined_ast: 556b4b4c66f55375093c41eaf4e524e2b333b423672c538db60fd18a249068ad
bytecode: 976c0daf1133bb687c763b552cf546d3c02ad6f2ba23def2a1aec0e56e5aff64

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 4a1e5c2818e99b3779669345fec31621c54768b926a6e246ba5edeb1eb123f21
unrolled_ast: 4a1e5c2818e99b3779669345fec31621c54768b926a6e246ba5edeb1eb123f21
ssa_ast: 2c0de477e742beb8a851d2e26b18f9342b94ae65893bb55518a689b31b038572
flattened_ast: 894437e501f829491e9e22617052c169541174a854aa2afedd5d3c912b71400d
- initial_ast: b3f5b26e60cf731201f24d6a0d24983196996478f878dadca2aa6b0bc9ec25fc
unrolled_ast: b3f5b26e60cf731201f24d6a0d24983196996478f878dadca2aa6b0bc9ec25fc
ssa_ast: 9d8bc97ef40df6abb744d278769607e6f22c86ccf03767dc36e1589fbf1c7e76
flattened_ast: eb1efd8a8e0fa18452cbb0aa136995b9823d3b53fe8e579c49b31f721d65cf53
inlined_ast: eb1efd8a8e0fa18452cbb0aa136995b9823d3b53fe8e579c49b31f721d65cf53
bytecode: d8c824aec550f73676651160b81bf3e546b5958ec8520629f779b9156b033032

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: fd09783877c99da74d6ba033e99e7bda8e82433c307d93e9c0a4e3914d6a109a
unrolled_ast: fd09783877c99da74d6ba033e99e7bda8e82433c307d93e9c0a4e3914d6a109a
ssa_ast: 74e79d5cbd5e7b80826ff017294a6542b55cfe9e0734e82c8d0b47c257d1f0a3
flattened_ast: b5d3d72262fec88ea316abd2f6dfaf90e199aceb959b7eda6eaeddcc419233b2
- initial_ast: df336b5d77ece6762c2c118cb15cc6f25a6f108b66df81d8986ff2ba69d45461
unrolled_ast: df336b5d77ece6762c2c118cb15cc6f25a6f108b66df81d8986ff2ba69d45461
ssa_ast: ab2119a0fea501f0cbe04cdf465230625ee6f447ee9e7fe110b64aa6fc3c6192
flattened_ast: 97b6619ac108948a86250e9498148db33ae70756677e6ed23474e8268f9c061c
inlined_ast: 97b6619ac108948a86250e9498148db33ae70756677e6ed23474e8268f9c061c
bytecode: 45f6f395a7abca14ca5a1e3d5dc3fb2ac5ea3946740553ca5aee7a5822a235ec

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 557130262ded1b479b139b15c095260d83a1c3bdd038d7cb61e6ad852f760fe5
unrolled_ast: 557130262ded1b479b139b15c095260d83a1c3bdd038d7cb61e6ad852f760fe5
ssa_ast: f9a7a186621cb450664b053f18293aeec444b650403adc294cf1572958b6033d
flattened_ast: 44461faaa8269377891df43705df40021a18e0cdf75e5daebca3503005a71ab4
- initial_ast: c57c36b9b59e05c91849e91969a79474d5fcb49d284be8140ffaac4746ae7191
unrolled_ast: c57c36b9b59e05c91849e91969a79474d5fcb49d284be8140ffaac4746ae7191
ssa_ast: 586af29b0b79edf5f19a4bd494d8428f22a87f851119d7f74eab9b535810468d
flattened_ast: f72217bcb7185ae66114addb89c1bcf61da6ff200ece88309aa50dc3f5ebefd9
inlined_ast: f72217bcb7185ae66114addb89c1bcf61da6ff200ece88309aa50dc3f5ebefd9
bytecode: e5e0c25f5c089802ae51be9f719ccd87b53adf4676bc463ddf6e6f63d6e1f977

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: bb21bfd7f94edff0912f41c0c530966e1ffe6d117ed3143e2ec005af8a2f4ab0
unrolled_ast: bb21bfd7f94edff0912f41c0c530966e1ffe6d117ed3143e2ec005af8a2f4ab0
ssa_ast: 4a5fbd468288f7c57f21d71246fe1aece51945b617ec91d217e734faa88b7571
flattened_ast: fa3e33dae39e487a23accfa895fafbe1d9fa2e909d17b74d71135f5f3ab26340
- initial_ast: c00a32cd00345e923e87dcd16427c9f2ec733022431e38cceefbb5ab6143b161
unrolled_ast: c00a32cd00345e923e87dcd16427c9f2ec733022431e38cceefbb5ab6143b161
ssa_ast: 29ecb3770403a15ee05a695260ebc6f7b8e313b4614e3a1f06de34b4d03ff147
flattened_ast: 9711e4d72e2e9e85b24e3b4b3e73cc939a05a5846733c0eb15dab5c5b54a054a
inlined_ast: 9711e4d72e2e9e85b24e3b4b3e73cc939a05a5846733c0eb15dab5c5b54a054a
bytecode: 9217044f6eb12f18c1298c2ce3217533eb27618e7c8c5ead76848d21935783d4

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 7add075b516b60285bc9d29c17f5139f685fd899d332a401a4204d6067327419
unrolled_ast: 7add075b516b60285bc9d29c17f5139f685fd899d332a401a4204d6067327419
ssa_ast: 42f2f811d77249415e50dbaf99585f8617f8d14f934b5043798ab0c5c8f67ec8
flattened_ast: 734d6c9e1008eab4451158ad667a9d8e34e2af6e960548d501057fa2558299bc
- initial_ast: 4eaa587c05eb4f4a65f33d94cd12a1ce47a55263726b7619578846c620735b3d
unrolled_ast: 4eaa587c05eb4f4a65f33d94cd12a1ce47a55263726b7619578846c620735b3d
ssa_ast: 3e1a3d03f465a60b2ceb2fc480551d9d498beb758a6b378ae6558117ec2955a7
flattened_ast: 2448c3c0819a99113f4acb509371009619592a984911b836f9787a386cf1e617
inlined_ast: 2448c3c0819a99113f4acb509371009619592a984911b836f9787a386cf1e617
bytecode: 6a07bdcfa9cc3f72be7acb20de65bed8983094471083dee01fc5a46d923e7506

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 4205da189bdc3b248ed6fb048073e333135e4f7a8af2700adcb5c65554063794
unrolled_ast: 4205da189bdc3b248ed6fb048073e333135e4f7a8af2700adcb5c65554063794
ssa_ast: 1af2511db5bcf2cf74638c72d6981ae8e16005436e8c3ebd216bc4788a48b6d1
flattened_ast: 538f7c5c282048564690bb02ea04066e4e1e6a67f2b7550e87bf7cb53c4aab0c
- initial_ast: 8693aac7091d75fe65a52352decb2ce24e225dae9396cb45809575f7c3cbd8d9
unrolled_ast: 8693aac7091d75fe65a52352decb2ce24e225dae9396cb45809575f7c3cbd8d9
ssa_ast: 5a33864a3f91c0d4a63171ed36ef709b7e75593a3181b4ed3f11be873ce2b9a2
flattened_ast: 8a13f93c69d995ea32ab518a4287d77dd9e37e4e1f15fd257361c58a0f853e7b
inlined_ast: 8a13f93c69d995ea32ab518a4287d77dd9e37e4e1f15fd257361c58a0f853e7b
bytecode: e893a23da89b538d6d95e87e9a97340f63c798fda07cf50166d09e8c4e07932b

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: b944463bfcb82f802952036416b0e8908866974395da5e0a0c26afc1b0f19b5f
unrolled_ast: b944463bfcb82f802952036416b0e8908866974395da5e0a0c26afc1b0f19b5f
ssa_ast: e426c41fd8d87d5990e6bbb8dd934be2baa95d665850ecdaf5e3109a3675b541
flattened_ast: 6947188abbeb502e1a4c8be8568df35f2d4739e1d7c15213757c7ef4e37205f9
- initial_ast: 8e58cb6efd83104f693ef4a01ff24d7c9975357b2c01a3f5bf332d55738e90cf
unrolled_ast: 8e58cb6efd83104f693ef4a01ff24d7c9975357b2c01a3f5bf332d55738e90cf
ssa_ast: 44567392e4c23c50c0e6be993952ce738f3b6315ebb825f8ffe292937b3a30ba
flattened_ast: ab0f2b3350a96338d70af7d1e3e27ea74273c95b83591e365724250a3633a406
inlined_ast: ab0f2b3350a96338d70af7d1e3e27ea74273c95b83591e365724250a3633a406
bytecode: b82322298b5a498c7a2a308c597a5e5d546becd9ff6d05d0c571247e09f1cb7d

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 9308c9c592608dcd11960a906c0753504372209da9f7b7669ac3f9adda28a592
unrolled_ast: 9308c9c592608dcd11960a906c0753504372209da9f7b7669ac3f9adda28a592
ssa_ast: c1f73a8466d54233516776dc587ff566345e879c0e358776a2774d32d0d154da
flattened_ast: 0b6a282a12ba27a9968282a45dccbc2b69754c56c6f719b30895de9d7ef1bec2
- initial_ast: e7b840ff0656282ab0864e77d5e3f4a3053e98c131ccf0934f1815477c00214f
unrolled_ast: e7b840ff0656282ab0864e77d5e3f4a3053e98c131ccf0934f1815477c00214f
ssa_ast: c06874c153b1936f3caaa78f246738326ece3b38597f6507d11acf05f1d70375
flattened_ast: f776dd25cd7be3d76ca22588625d64ef4df262e70aa069a3664e5045000f8bac
inlined_ast: f776dd25cd7be3d76ca22588625d64ef4df262e70aa069a3664e5045000f8bac
bytecode: 849a917a861f86e0a277f0a92a018a81f6f6d69762816e29e585452dd805a1c1

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 095d15c58d65096d77a62ff54af5ab129002522930e0a8ca83dd0c6de4467bec
unrolled_ast: 095d15c58d65096d77a62ff54af5ab129002522930e0a8ca83dd0c6de4467bec
ssa_ast: 368a16973432fa980ce1b15d598fe6c27287055c3aefc05cf6b575cc2f4e7447
flattened_ast: 0ffd6d24cffb39df46258a3a2a11c297fdf1ac650ba368378aa1906ededc7d8b
- initial_ast: fb480abe75542e3b109f037b1ffd53ea64bb92b21f2b33f8297e2d6f24dc1fc0
unrolled_ast: fb480abe75542e3b109f037b1ffd53ea64bb92b21f2b33f8297e2d6f24dc1fc0
ssa_ast: e213f529b4ff45bebaac280ee485c8bdc3af09463c349368ab236a26cf0a9370
flattened_ast: bfaa00a0b6158f6798f5aa5852a07aa9f437d80a9b0e4e811a192fe111111084
inlined_ast: bfaa00a0b6158f6798f5aa5852a07aa9f437d80a9b0e4e811a192fe111111084
bytecode: 8d921ede85807f033431e06b604875bb1f6712fb96957faa5a805fe02fe24245

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 7a5cb78dd349a26c2962daece4fc6a3eaa4f03fb67f31ace4a9d84a5ce658e34
unrolled_ast: 7a5cb78dd349a26c2962daece4fc6a3eaa4f03fb67f31ace4a9d84a5ce658e34
ssa_ast: a087b2cd64869c8d52a49d39d50aaccc710ec9153fe8f96be820de8d3dc02854
flattened_ast: 41087f4e5ebf8e00900384acacecb9d2bc66774d31c63fef66ab5c4b5f44024e
- initial_ast: 9845d8aaec21f191c73e190b478e592e3e910b0dfd071cf86e692082f7ff9f23
unrolled_ast: 9845d8aaec21f191c73e190b478e592e3e910b0dfd071cf86e692082f7ff9f23
ssa_ast: a24e603330c02f87b70ed3e3f6467fc471d6d9d032f17eb023f37df005ceff85
flattened_ast: bc4a52e6fb7998c2a8a454306e75598177546db8f32a5a79e95ead68abc72880
inlined_ast: bc4a52e6fb7998c2a8a454306e75598177546db8f32a5a79e95ead68abc72880
bytecode: 9a1e5bb7b8d932d4afd347a856bfb38db144771f49a0d9589ef14236338e3dcf

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 0ae25e3f48a005325105c1ae4859df780f56dceaa7ece1c9f7036f804709d30f
unrolled_ast: 0ae25e3f48a005325105c1ae4859df780f56dceaa7ece1c9f7036f804709d30f
ssa_ast: 843981d448200fe4b207a10fea71a9afe5d3f8f255122efb3108644122cd85a4
flattened_ast: df1e8e6fc3eb3558b3069137816491d8e6ccdeabf1e9934b70e18dfb3c1c6ebb
- initial_ast: 179113dfb723097a26a7181bbb64ed56e7ca3ebf57cf782077d9743b64595446
unrolled_ast: 179113dfb723097a26a7181bbb64ed56e7ca3ebf57cf782077d9743b64595446
ssa_ast: 499f15e97d7e5d5dcc0be12e85176e6f160cc2b65d66b0667f640a3d0e64f369
flattened_ast: 9c98dfcdcb403983efb0b1078246ca9e3c3e8fe913f6ceabbd9a87b63f3fc3a4
inlined_ast: 9c98dfcdcb403983efb0b1078246ca9e3c3e8fe913f6ceabbd9a87b63f3fc3a4
bytecode: 230d4f2bda3933eb4fafc4dda4ce0087e74e4cbd9c65349746da561cbb3f99da

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: c4bf7caa38d22e8398b21a08af4bb1c675160e8d03d4b96b392d6f0c595918c2
unrolled_ast: c4bf7caa38d22e8398b21a08af4bb1c675160e8d03d4b96b392d6f0c595918c2
ssa_ast: cf4f7d1a7e63e3979a8aca25d94719410856bbc8e022df52ef4811d83fec9ea6
flattened_ast: 892e661b9c5a87b5598ae3f9c405a5d8d04a82dd70cc06980fb05b5e661b8820
- initial_ast: 873e6714527c41b8cf2b3db3236b443e8ead62c4580b57b4088c46c378524598
unrolled_ast: 873e6714527c41b8cf2b3db3236b443e8ead62c4580b57b4088c46c378524598
ssa_ast: d451c529bc4b3207205083e58d6521f0ea5526d63d1f77c42b40854f917316cf
flattened_ast: 0840cf638ec3532c7702d10bbbfcf2fbfc8c8f7c54e532acb4ac46cbb7c8ed61
inlined_ast: 0840cf638ec3532c7702d10bbbfcf2fbfc8c8f7c54e532acb4ac46cbb7c8ed61
bytecode: fa960590c979aea4bdfe07b7d37060bb593f73f745974241e2db578bd7ba2ced

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 0a2a053e5bf78c3cc425dbb1600479aef43f652c58b34b8749aa36765720889e
unrolled_ast: 0a2a053e5bf78c3cc425dbb1600479aef43f652c58b34b8749aa36765720889e
ssa_ast: a51dea4bd834bde7d545b7226c1d94beb2aa67de016288c4a33d7bf6a11e5084
flattened_ast: bf525b12a2743590934d071970c062b6383680a16de9f12e3f0a62422b4e6e61
- initial_ast: 63d16c8101f9ede853a5be9d493bc61bbe57449e99499e42254a1e2d6448e3d2
unrolled_ast: 63d16c8101f9ede853a5be9d493bc61bbe57449e99499e42254a1e2d6448e3d2
ssa_ast: df4cad7af230e0feb2036b920bde4aa81ed297a9ee8269aa95ded280610bde49
flattened_ast: f1d531bbe1b2e0bf0f30a1f7e86cce88c834fee9eb4d06548508907ad5f2dd24
inlined_ast: f1d531bbe1b2e0bf0f30a1f7e86cce88c834fee9eb4d06548508907ad5f2dd24
bytecode: e8cc0536d26ff27b9fe9ff3ad45b575185b9f60c9d3910481ab66843af0f2171

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 39a8a1edc287de186596bad9afa19c71d2bdf424eb5a666075d262edadb582ab
unrolled_ast: 39a8a1edc287de186596bad9afa19c71d2bdf424eb5a666075d262edadb582ab
ssa_ast: 2d88cadc3172088d6282e7ee6353af69ba372f592038028ce52cf4db1abdf8d5
flattened_ast: 021d9f732dac254daf87cc96a1378bca92202b0ddeedb2ae9057638caadda0b1
- initial_ast: 5b9e236e49cd157cbbcdb0c600c17901cd6d51169e8c115a293a1dae6d519d28
unrolled_ast: 5b9e236e49cd157cbbcdb0c600c17901cd6d51169e8c115a293a1dae6d519d28
ssa_ast: fcaf1a28f73f1c528ce87d8ad3dd3b30344e16a9deff4d2785f3f485b68f5a48
flattened_ast: bd09931f20c9c72dce35ccb55bdf5fff1f5a4754fcf4dbff46a0a8c79ca88abc
inlined_ast: bd09931f20c9c72dce35ccb55bdf5fff1f5a4754fcf4dbff46a0a8c79ca88abc
bytecode: eeb44a4faf22686de577f93db551bd83246583158dcecb35d2dc454e0693e419

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 7fbe1dee1da36bbf500eac5fdb3c85096b160e10da9dc765e072bb3a3e6f8725
unrolled_ast: 7fbe1dee1da36bbf500eac5fdb3c85096b160e10da9dc765e072bb3a3e6f8725
ssa_ast: 14a61902cd2de8c44d6ad562a382c42bc63764b975f9c668f90eb2c9297258f7
flattened_ast: f8cb858f6f41179bed737e5b47b138291dde83ec0a5d30787c7e1cd6d750e7c1
- initial_ast: 6b2150ffe972e2a329964dd76f061f3af94a9b15ce821ad4bb1cedd6b0134483
unrolled_ast: 6b2150ffe972e2a329964dd76f061f3af94a9b15ce821ad4bb1cedd6b0134483
ssa_ast: 610db9a9ab1fe344961c36a0fe5170902b9ca4cf036094b0a5f6fc9d8cfa7b72
flattened_ast: 384e746fcbe1428be942f6ee12041562e0c1ae98b398c26da1d62fdb57181343
inlined_ast: 384e746fcbe1428be942f6ee12041562e0c1ae98b398c26da1d62fdb57181343
bytecode: 90662aea378f911f2798c1ece956f7a2566fd99d99a87d8285f1476edf468e43

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 66d929a26a59181cacf5631183d943ad5d90de0c011f00b5e7fd0a86fbca6aea
unrolled_ast: 66d929a26a59181cacf5631183d943ad5d90de0c011f00b5e7fd0a86fbca6aea
ssa_ast: 1e0bc0c103cf34a316892d929e5018d04157bd2b8fccf21a39d3ba34dffb3bb8
flattened_ast: 07518cb75989513b755211d766ac61f7263d96203dcdd48af48b0fb8af64e6d4
- initial_ast: cde2a91af921e65b79849153ed229de8c9a0af850ee62ac23363d0f8d8b82899
unrolled_ast: cde2a91af921e65b79849153ed229de8c9a0af850ee62ac23363d0f8d8b82899
ssa_ast: ebba08995e71307851655254c51deb67364ee12aa4320f9efa32c16668d26cf6
flattened_ast: 7111dab311b76ad61366abb7a6e40586f44e17da7f8784eb6f8431dd0c41bd42
inlined_ast: 7111dab311b76ad61366abb7a6e40586f44e17da7f8784eb6f8431dd0c41bd42
bytecode: 57bdcce5ea2ea7890a6a4786e4795f5c458da4b6b29f6295f86e15f11479f3e6

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 730c7aae808fdd0cf9b2e008612191ad36d29f7f4fe7770c9b72371a1146b65f
unrolled_ast: 730c7aae808fdd0cf9b2e008612191ad36d29f7f4fe7770c9b72371a1146b65f
ssa_ast: 6b20fc19648d776c5321df0b9b3b403456cae4ed0aa815366373a64f7810e0cf
flattened_ast: b4ce53a9ebedbfe05b96f67e0894f31a70d0948d371574772e76dc3adbeb19b0
- initial_ast: 6926fc7a56a99fb841bff97631e39b332f998908009c9c0f83e0f0b4d1a0b8f3
unrolled_ast: 6926fc7a56a99fb841bff97631e39b332f998908009c9c0f83e0f0b4d1a0b8f3
ssa_ast: 842449a3e29d8cd788deae538a1642bc89e326ed45768ee5121095e4293f553f
flattened_ast: c48dab53c83532533096307cda1460de12397de310e2b6622f644dcace4f4391
inlined_ast: c48dab53c83532533096307cda1460de12397de310e2b6622f644dcace4f4391
bytecode: 1bfceea51d0a0df233268cc281d300a3c15c291de63528a723a763eba97e9b93

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: ab79658672090ee72bf8485cdc370f980531343942292b54ff3f1b3ab9fce67d
unrolled_ast: ab79658672090ee72bf8485cdc370f980531343942292b54ff3f1b3ab9fce67d
ssa_ast: e2c1c59b1bb7a799df57044ab663d0c4b93f8c5e96c1927639abd8c0a9dc8488
flattened_ast: f75ea6e363e36e8123507e5a285072730e03621238ade5f5c96e53314ba2b1fe
- initial_ast: bdaaa81417cac5a4c216d7d4255e79cf0d29b1f911f689956f376b1a1e265c44
unrolled_ast: bdaaa81417cac5a4c216d7d4255e79cf0d29b1f911f689956f376b1a1e265c44
ssa_ast: 0decc436c2c4d95b24b526a58adff114601ed3ac8a4cff335ab5229aed730480
flattened_ast: 088aaeccfd897a2e6d323966090c3989859be3fdaf9ab3e19b54cdaebde926fc
inlined_ast: 088aaeccfd897a2e6d323966090c3989859be3fdaf9ab3e19b54cdaebde926fc
bytecode: 7540a269502febfe91bebfc15030891bde7667f921d5d8d9d22efbcf16410543

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: dc0f03e353714f93a4bdc39a64369a4bf6af4725dea2342523c0447e5666b63f
unrolled_ast: dc0f03e353714f93a4bdc39a64369a4bf6af4725dea2342523c0447e5666b63f
ssa_ast: ceb2eeb666f437f0adeb57cb689e501a3a0981ba14663b9e4f85bcf795aa5ecb
flattened_ast: dd5f38e0fa2531c81ad10c5aa2143388c47caca081cf7fc56a2e39cb0d07d98b
- initial_ast: 64dda97ee29caaad409d6138d57d8ed852caee9e40486539e03fcc570c7c3d1a
unrolled_ast: 64dda97ee29caaad409d6138d57d8ed852caee9e40486539e03fcc570c7c3d1a
ssa_ast: b893aa52ba3b0404bcfdcd8f9708df62cb91b70ba5e9417e1455fc7710c6ceb6
flattened_ast: bbf216c1e754d2012edb4ef4896499255d956bf4f39c0b9852ee45f75d914a0b
inlined_ast: bbf216c1e754d2012edb4ef4896499255d956bf4f39c0b9852ee45f75d914a0b
bytecode: ef0f05392652587de58875f041bb805a5a1172a153d96973638342d143798863

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: ce068059ab68c906175c70d259fc716165e8749596c80ec6e2be52c88d076335
unrolled_ast: ce068059ab68c906175c70d259fc716165e8749596c80ec6e2be52c88d076335
ssa_ast: 48b84a8784368f924f6b15526204f34921eee3224112cc0555a223c169047809
flattened_ast: be6f3fd5174704ed384b73de24f569b4ca3744480bd454b7501a52a4dddb98f6
- initial_ast: 1ea7ba67b3f77b976b2ec3493afe6958ea64055e10f7a15f33235ff62dd488ee
unrolled_ast: 1ea7ba67b3f77b976b2ec3493afe6958ea64055e10f7a15f33235ff62dd488ee
ssa_ast: dd0df684331375510fdd96af15bd2aadb8932f3eff2fabb9d1b8dba199728b48
flattened_ast: 60cf4f7e83d3ffc10b362b701749b0d5afcf8307e099bc5c7908c9ccb4df3efc
inlined_ast: 60cf4f7e83d3ffc10b362b701749b0d5afcf8307e099bc5c7908c9ccb4df3efc
bytecode: b65dba415908458745a14bfc52abda70a0899732f807ba22f56776ab3fcbf589

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 56080d8ee17c7eedbe41d281e986a8b09fa1d0e33bb1c7876fc93599c194f0fa
unrolled_ast: 56080d8ee17c7eedbe41d281e986a8b09fa1d0e33bb1c7876fc93599c194f0fa
ssa_ast: 56080d8ee17c7eedbe41d281e986a8b09fa1d0e33bb1c7876fc93599c194f0fa
flattened_ast: c49b193e77aad777550c12ac43fdb3fed1d33ef1a841af37294c43902fef9ee8
- initial_ast: 6fa465e63f2b8e880d621cb1758b3d1c0edfa9ce09e6d4f0f28bbe6c2ca2b955
unrolled_ast: 6fa465e63f2b8e880d621cb1758b3d1c0edfa9ce09e6d4f0f28bbe6c2ca2b955
ssa_ast: 6fa465e63f2b8e880d621cb1758b3d1c0edfa9ce09e6d4f0f28bbe6c2ca2b955
flattened_ast: 4b8969d1adf68074bc7a8458a9146e128041bf929f2f6a4e76a16ad2769b81b1
inlined_ast: 4b8969d1adf68074bc7a8458a9146e128041bf929f2f6a4e76a16ad2769b81b1
bytecode: 39aa8516297ece27331b633a72466d2ff0122d36beca663a48bc07589e2d3e15

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: ab5dac20c924339789f9c0ac6fcd035b58ea7b9d3b2084a5151cba3c14f7939b
unrolled_ast: ab5dac20c924339789f9c0ac6fcd035b58ea7b9d3b2084a5151cba3c14f7939b
ssa_ast: 83186e2623d9152de399d213137d7813c027e1e76e607295b66b8c03dbe32a4d
flattened_ast: 26b244650007d2841b4dec172241e8eae1091483054e8c82ac61df52fd63f91f
- initial_ast: 9ffd4f34e261ee315a2ee29353e707b35c300e928aca532743142a538957c4ce
unrolled_ast: 9ffd4f34e261ee315a2ee29353e707b35c300e928aca532743142a538957c4ce
ssa_ast: 9ea8f3743b9bcf1584319472ca0a80707f117616d687d4c401b34fb10b44703a
flattened_ast: a4eca8b80af9863d59ebfb837fa5dae061fca7d52315d3a9f5778e6dc4b75716
inlined_ast: a4eca8b80af9863d59ebfb837fa5dae061fca7d52315d3a9f5778e6dc4b75716
bytecode: 6db857dc2b80ea257d141b3980404e050024771f95c5f9b74f899145b2001432

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 521891369f880815ee6ac54d87588c147feff8ec4cb62128c44afdf792e30e76
unrolled_ast: 521891369f880815ee6ac54d87588c147feff8ec4cb62128c44afdf792e30e76
ssa_ast: fe601539c9d1cc61da7a956647e4712edacc7928d3cce050ade833e2fb9da0d8
flattened_ast: 5c1a7f7f8356f0de01c7cd183adb76570fa9f9e80f3f5853c761293cb6ca355f
- initial_ast: 809ffb6b3ce9e51d58f18a841bbbe79e87bf3b5c0ac2d82d226b3cb66427f235
unrolled_ast: 809ffb6b3ce9e51d58f18a841bbbe79e87bf3b5c0ac2d82d226b3cb66427f235
ssa_ast: efeab621ea3b6113ae3ef1f326cbd75668ce034b81a1bb09a55c9a671f62c127
flattened_ast: 06fea09c85a281be025d66565aa362f80f2036c88c284fbfb5f9b874a605916b
inlined_ast: 06fea09c85a281be025d66565aa362f80f2036c88c284fbfb5f9b874a605916b
bytecode: 9f1144202f6b114409c379f7ecc4b480dd81daaf0f6f8b244efd20c520f7b76c

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 1fc29cb05bcd2261b761b2342516525cf7cf27f897cce03d13ea456362411b9d
unrolled_ast: 1fc29cb05bcd2261b761b2342516525cf7cf27f897cce03d13ea456362411b9d
ssa_ast: 1fc29cb05bcd2261b761b2342516525cf7cf27f897cce03d13ea456362411b9d
flattened_ast: ad48f1a8f2207a90e1707665b0a1047b313647d83a38a413f7fb5bc1751f22ef
- initial_ast: 19378936c22e4e747e16e132bbc727115598dfbd17068349cb300525cde35556
unrolled_ast: 19378936c22e4e747e16e132bbc727115598dfbd17068349cb300525cde35556
ssa_ast: 19378936c22e4e747e16e132bbc727115598dfbd17068349cb300525cde35556
flattened_ast: c55a0edeb6a52dd728e5500ff5b1d387186321c8a3d68f2d0638628bbb05696e
inlined_ast: c55a0edeb6a52dd728e5500ff5b1d387186321c8a3d68f2d0638628bbb05696e
bytecode: 49afa4d378578bc680308083733b31b8272f9c952fe8dbc133398676e3f0d2ba

View File

@ -0,0 +1,5 @@
---
namespace: Compile
expectation: Fail
outputs:
- "Error [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:8:9\n |\n 8 | increment(values, 0u8, 1u8);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:9:9\n |\n 9 | increment(account, self.caller, 1u64);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:13:9\n |\n 13 | decrement(values, 0u8, 1u8);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:14:9\n |\n 14 | decrement(account, self.caller, 1u64);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372044]: Function must contain a `finalize` statement on all execution paths.\n --> compiler-test:12:5\n |\n 12 | inline bar() {\n 13 | decrement(values, 0u8, 1u8);\n 14 | decrement(account, self.caller, 1u64);\n 15 | }\n | ^\nError [ETYC0372031]: Only transition functions can have a `finalize` block.\n --> compiler-test:17:5\n |\n 17 | finalize finalize_no_params() {\n 18 | foo();\n 19 | bar();\n 20 | }\n | ^\n |\n = Remove the `finalize` block or use the keyword `transition` instead of `function`.\nError [ETYC0372045]: `finalize` name `bar` does not match function name `finalize_no_params`\n --> compiler-test:17:5\n |\n 17 | finalize finalize_no_params() {\n 18 | foo();\n 19 | bar();\n 20 | }\n | ^\nError [ETYC0372066]: Cyclic dependency between functions: `bar` --> `bar`\n"

View File

@ -0,0 +1,5 @@
---
namespace: Compile
expectation: Fail
outputs:
- "Error [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:8:9\n |\n 8 | increment(values, 0u8, 1u8);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:9:9\n |\n 9 | increment(account, self.caller, 1u64);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:13:9\n |\n 13 | decrement(values, 0u8, 1u8);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372035]: `increment` or `decrement` statements must be inside a finalize block.\n --> compiler-test:14:9\n |\n 14 | decrement(account, self.caller, 1u64);\n | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^\nError [ETYC0372044]: Function must contain a `finalize` statement on all execution paths.\n --> compiler-test:12:5\n |\n 12 | inline bar() {\n 13 | decrement(values, 0u8, 1u8);\n 14 | decrement(account, self.caller, 1u64);\n 15 | }\n | ^\nError [ETYC0372031]: Only transition functions can have a `finalize` block.\n --> compiler-test:17:5\n |\n 17 | finalize finalize_no_params() {\n 18 | foo();\n 19 | bar();\n 20 | }\n | ^\n |\n = Remove the `finalize` block or use the keyword `transition` instead of `function`.\nError [ETYC0372045]: `finalize` name `bar` does not match function name `finalize_no_params`\n --> compiler-test:17:5\n |\n 17 | finalize finalize_no_params() {\n 18 | foo();\n 19 | bar();\n 20 | }\n | ^\nError [ETYC0372066]: Cyclic dependency between functions: `bar` --> `bar`\n"

View File

@ -0,0 +1,10 @@
---
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 66c08e3fa6b22b7663e949d6ea72455bca2c26a573ba094f084310a086a4cf0b
unrolled_ast: 66c08e3fa6b22b7663e949d6ea72455bca2c26a573ba094f084310a086a4cf0b
ssa_ast: b86f0471ac1a09039a4ce147e87b0c41b9c3378dba0cb560b4a2fe41533781c2
flattened_ast: 52f744a6cf805739c77995f73c2f626ad2403301c5dc6e007b9c2092869f5224
inlined_ast: b7ef966c924c7fd055476fd974079af70e235aed43d8cbff30c7eadad8e342c7
bytecode: 614bd1f31a52b5176e1c8a7b8a7fc17ced721ae71e50d33daa09ebc65b5fff20

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: fcd02f88db2176cd7e0befcda00311776bff3f8518b272c256c0e19de8eb7dd9
unrolled_ast: fcd02f88db2176cd7e0befcda00311776bff3f8518b272c256c0e19de8eb7dd9
ssa_ast: ba628027acf56c96c1047a38f2adb820e4dd16e7e82a77e210f60cbee2265f13
flattened_ast: 6d53ef7e6768f1eb0b230e73c743a0d07ef5eca2e962ae8666c3a88677a33b03
- initial_ast: 654954e49f42b92d7405fb08be0dcd80810abbdd366623ed6bb36ba99ddbaa58
unrolled_ast: 654954e49f42b92d7405fb08be0dcd80810abbdd366623ed6bb36ba99ddbaa58
ssa_ast: 45cca036e205fdb706c442f8d507038e0ec813cfa4de443b8c69c8485d5256e9
flattened_ast: 8a8a44500158d7b38ccf354a7d14c4292c3cc302889489d0fe3761c1917befed
inlined_ast: 8a8a44500158d7b38ccf354a7d14c4292c3cc302889489d0fe3761c1917befed
bytecode: 1da5a78fcb6f77bd197de7dce1e7e94e7a9d30a6ec26703a645b25ab7c65cc08

View File

@ -2,4 +2,4 @@
namespace: Compile
expectation: Fail
outputs:
- "Error [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:7:16\n |\n 7 | return two(n);\n | ^^^^^^\nError [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:11:16\n |\n 11 | return three(n) + four(n);\n | ^^^^^^^^\nError [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:11:27\n |\n 11 | return three(n) + four(n);\n | ^^^^^^^\nError [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:15:16\n |\n 15 | return one(n);\n | ^^^^^^\nError [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:19:16\n |\n 19 | return one(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:23:16\n |\n 23 | return six(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:27:16\n |\n 27 | return seven(n) + eight(n);\n | ^^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:27:27\n |\n 27 | return seven(n) + eight(n);\n | ^^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:31:16\n |\n 31 | return five(n);\n | ^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:35:16\n |\n 35 | return five(n);\n | ^^^^^^^\nError [ETYC0372066]: Cyclic dependency between functions: `one` --> `two` --> `three` --> `one`\n"
- "Error [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:7:16\n |\n 7 | return two(n);\n | ^^^^^^\nError [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:11:16\n |\n 11 | return three(n) + four(n);\n | ^^^^^^^^\nError [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:11:27\n |\n 11 | return three(n) + four(n);\n | ^^^^^^^\nError [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:15:16\n |\n 15 | return one(n);\n | ^^^^^^\nError [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:19:16\n |\n 19 | return one(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:23:16\n |\n 23 | return six(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:27:16\n |\n 27 | return seven(n) + eight(n);\n | ^^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:27:27\n |\n 27 | return seven(n) + eight(n);\n | ^^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:31:16\n |\n 31 | return five(n);\n | ^^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:35:16\n |\n 35 | return five(n);\n | ^^^^^^^\nError [ETYC0372066]: Cyclic dependency between functions: `one` --> `two` --> `three` --> `one`\n"

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: f54d04dcf65f240192da3a67407a1ab8c97942603bb031269a3b6e20e450c364
unrolled_ast: f54d04dcf65f240192da3a67407a1ab8c97942603bb031269a3b6e20e450c364
ssa_ast: f7a38580c6da1a0cb128799ff979843f913b87af368920ed5dbbd72b64a8f67e
flattened_ast: b31dbb46a29aec9009f7aff2a8f6f206cfed5fd837174aa66d450b38adc4df21
- initial_ast: 01eca8ffd19d37c08eee234033181ccb72873bc1fff02a90e1863e24c9e2d446
unrolled_ast: 01eca8ffd19d37c08eee234033181ccb72873bc1fff02a90e1863e24c9e2d446
ssa_ast: 9e151de216e44ca801adec05a7b534cbe347c3a64f31d570a9f33591a90af191
flattened_ast: 9f1b62847c7b725e934fd72fb9a6ab076a6d1c778957bb93e6d2e4c7c0910c58
inlined_ast: 9f1b62847c7b725e934fd72fb9a6ab076a6d1c778957bb93e6d2e4c7c0910c58
bytecode: 434d585ff5cbe799cf645514abda7bc7ad069563501ded68fc716e583390fefa

View File

@ -0,0 +1,10 @@
---
namespace: Compile
expectation: Pass
outputs:
- initial_ast: d9f4c437725505b182331f81158f4d4fe868ac4c51bc09e7adb207668875c1d3
unrolled_ast: d9f4c437725505b182331f81158f4d4fe868ac4c51bc09e7adb207668875c1d3
ssa_ast: 52bde71da7f559a45cb7dff0b9ec4634cbca56d1d02dad9ade11cd733590d5e1
flattened_ast: 47f2df48295953cd15343288c0ce7de1394e82e6c34734992397df74c1b2138a
inlined_ast: 419b798e293216e127d3a3ed8775e39a5ab0c6cb2b0bfbcd69bbeffed7d0f113
bytecode: d84ad59c8407d4d733886e7442e47de32cc3025f6e411d2cce0816ce30f0b488

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: f2670ba11429088c8d6ca1498d168df09a9706ac2d20d71cd29af5dd867d35e7
unrolled_ast: f2670ba11429088c8d6ca1498d168df09a9706ac2d20d71cd29af5dd867d35e7
ssa_ast: 16b1a88055e17a218086950e9ff90c81c167872faf26d901814952f3a3c57091
flattened_ast: 814492938dd116ec6cd0db1f084a22379a12127f123b7630efecd668452a33d9
- initial_ast: ef14a0a65f3e5450826ec37f7e2369f0f2b8d3314ed00da3ec900d4a51ed4eb9
unrolled_ast: ef14a0a65f3e5450826ec37f7e2369f0f2b8d3314ed00da3ec900d4a51ed4eb9
ssa_ast: 608b9c426b08b8b97dcf7f6a14c7dd0ebca594a7e17f79580cd1274757c4dd3c
flattened_ast: 7df8d7e0db134613ff120ee03cffc2fef159cf8204273668222e41d2b0efd2d1
inlined_ast: 7df8d7e0db134613ff120ee03cffc2fef159cf8204273668222e41d2b0efd2d1
bytecode: 2a939858f2f71f1bbe25bd039899cdb71254e56acc203eb6d60dbb5c191a4224

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: a21e39f581430c0a9a90d57dc00d6ba3baa09ec5a210d9c5021cd1b11072e91f
unrolled_ast: a21e39f581430c0a9a90d57dc00d6ba3baa09ec5a210d9c5021cd1b11072e91f
ssa_ast: 7b49247c500c5c179b892dd9fd7eff60104a1e04ffd00b9b53b5fbbc493b92af
flattened_ast: 99341a91c75455d056b346f1a35ecb55af04b18ddaecbbe4ec154d4d3512b0f0
- initial_ast: 374f33d7c3329a997d7740cbbc7e79d2c52b4ab70f14c72e1e77e01789e4932c
unrolled_ast: 374f33d7c3329a997d7740cbbc7e79d2c52b4ab70f14c72e1e77e01789e4932c
ssa_ast: a4166d418cadd9c9994b6d882f94cabbc80d302439c3e738c19d8603e3e56af4
flattened_ast: bf07e16742127065efb1989aef9582c08c8a6251057c5ae8e422901aaf46a3fa
inlined_ast: bf07e16742127065efb1989aef9582c08c8a6251057c5ae8e422901aaf46a3fa
bytecode: 27556a268723e0d8ffc4210290babab1ad098d9c8a77ad2dc84195d98059deac

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 5db5302effae454776641d39a58e71fb54e7a0dd8f9abb33aad033eae6bfd2dc
unrolled_ast: 5db5302effae454776641d39a58e71fb54e7a0dd8f9abb33aad033eae6bfd2dc
ssa_ast: 380883028c4da15e42781a4488f67821bc0a24da459de5a96fe424ed0296b09b
flattened_ast: 0689f093ce2cc6ca0be1ea491a268329d67ecc1251025a5d398ab174288e2c9c
- initial_ast: ed7e4dbd69ff4231ff48a6527734c7837b872f12a444668b071298b5cdd15d99
unrolled_ast: ed7e4dbd69ff4231ff48a6527734c7837b872f12a444668b071298b5cdd15d99
ssa_ast: e5afe402b1d3eeb7f1465a197952931597c2a8147aa624a15c985cad460479ce
flattened_ast: 8583c9afe3f11178c11fd4ff9e7babd3ed3a74719484084d1af353b7844dddc8
inlined_ast: 303725110ce5af01a45222f852d1096172a8aba363ef3b192bc026d2144fdf7c
bytecode: 713ce56eafa3f358be317894fd3ddf287a03422f855a304ee64becfcbd1f8590

View File

@ -0,0 +1,10 @@
---
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 42426e586a4fe69f06466e4621a8cf4547408ad38c108a85bf0beaf812859a66
unrolled_ast: 42426e586a4fe69f06466e4621a8cf4547408ad38c108a85bf0beaf812859a66
ssa_ast: 4b33649beef6673aafce25c35c16c034fb41ec33fdf562751b7e51d712caf309
flattened_ast: e2ee2ee7c6e1c187de5977f44b0a39e11d946243990eb48df92b3a3e939b3906
inlined_ast: aae0684cc29719b75d1af553fae21f9cafd04d15fd211eeb4d8cebbf913bcd6c
bytecode: da94a3a78e8ddb3883dafdc52da31472986f7e4a89348e03acaa1f7d1de5bb31

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 4d3f6c6d7473e2d981f1b66cdc0aa6c8c838b0cd5119749f88dd78173d948429
unrolled_ast: 4d3f6c6d7473e2d981f1b66cdc0aa6c8c838b0cd5119749f88dd78173d948429
ssa_ast: fa0c588d9ee1505dd06a0b90526b09e62f4336d4ce9f2e7be420235cd93bc9ab
flattened_ast: de929e7ddaf5ac129f46d064d2581b79cad2a4ff2d91243af77c07dd1d2d9146
- initial_ast: 4b891db2ba8f6ee0f2299622ecd387e6565c8b97045ff75808e559d7ee34c074
unrolled_ast: 4b891db2ba8f6ee0f2299622ecd387e6565c8b97045ff75808e559d7ee34c074
ssa_ast: f67f11e78a102b6eba419bb394cdf10b6bf8b1f2245c2fd022aa03db7750535a
flattened_ast: eddea7cbfc19bcdd6880741cc2446dd88b2d070c334dcf7e3780e71061c803f4
inlined_ast: 9220160485ed7aa28cd86b329b9805a7da3e1fb03d9526881ea356c9bc4f3f03
bytecode: a190851c7a73c1068c1c5819c3e064535d56273dffbc007874376c094399cd9e

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: ea45c4127e9cb23abb3761d287db4841f7d4bcde901104d33447cf5f8a40a964
unrolled_ast: ea45c4127e9cb23abb3761d287db4841f7d4bcde901104d33447cf5f8a40a964
ssa_ast: 3c660861f07ab818e4fc7ac137e9fcc2d8b39d5174a2fd307b0a478869bf8a83
flattened_ast: 5d5bc19468ea4393bb7a5f3737724c999030eb04c71d25e8d53e995706e2b762
- initial_ast: a1949ce7cd298b1eae7d1e68bd82735cb5430fbf3d98c846a742af495162c674
unrolled_ast: a1949ce7cd298b1eae7d1e68bd82735cb5430fbf3d98c846a742af495162c674
ssa_ast: 7af2a8837e274f8fad34cc663a822d4d901d71cda15268baf9b2b3131fad1b0e
flattened_ast: 66bc317eaff6320b4eb9eb8bdfad3dec14040cbeb81c9c73f3f7632ae699408e
inlined_ast: 66bc317eaff6320b4eb9eb8bdfad3dec14040cbeb81c9c73f3f7632ae699408e
bytecode: 56875e297f05e4c60762445a3ac97b57e4a0f12d69180bb7207ef62f950b0b25

View File

@ -2,4 +2,4 @@
namespace: Compile
expectation: Fail
outputs:
- "Error [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:5:16\n |\n 5 | return bar(n);\n | ^^^^^^\nError [ETYC0372047]: Cannot call another function from a standard function.\n --> compiler-test:9:16\n |\n 9 | return foo(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:13:16\n |\n 13 | return bax(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:17:16\n |\n 17 | return baz(n);\n | ^^^^^^\nError [ETYC0372066]: Cyclic dependency between functions: `foo` --> `bar` --> `foo`\n"
- "Error [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:5:16\n |\n 5 | return bar(n);\n | ^^^^^^\nError [ETYC0372047]: Only `inline` can be called from a `function` or `inline`.\n --> compiler-test:9:16\n |\n 9 | return foo(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:13:16\n |\n 13 | return bax(n);\n | ^^^^^^\nError [ETYC0372048]: Cannot call a local transition function from a transition function.\n --> compiler-test:17:16\n |\n 17 | return baz(n);\n | ^^^^^^\nError [ETYC0372066]: Cyclic dependency between functions: `foo` --> `bar` --> `foo`\n"

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 3d476147549e1c400f8d99b7868d4bb55873f63a8935b64335cdb8ebb2836121
unrolled_ast: 3d476147549e1c400f8d99b7868d4bb55873f63a8935b64335cdb8ebb2836121
ssa_ast: f36b1fa15ab9001afe65592d2c3f5a790301ff47c68fb3ef3766c75251cdbe75
flattened_ast: 161c5aa7b09fc1d3fba97998c52055855426cbe34ca5c05739f7831e8dcf30a5
- initial_ast: 2d1182494129f414a987c520aeb8fcfaaa9019d688f810d8ba5589accbb11747
unrolled_ast: 2d1182494129f414a987c520aeb8fcfaaa9019d688f810d8ba5589accbb11747
ssa_ast: 7ae7272d7babd64cc7845463755decb6073b095b04ae52e76197c68bc081cdf6
flattened_ast: c39c24be2f2e4792f87bf1e8dcd123064e1a9f31fec6923f8daf7800e6b9cd2a
inlined_ast: c39c24be2f2e4792f87bf1e8dcd123064e1a9f31fec6923f8daf7800e6b9cd2a
bytecode: 6d5fea51d9eec1cf3a5037b123147f9d532855197e3891ff870fbe700dd08d3f

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: 6d33052975f037a698138d1e3d7bb06ac22669c9f92997b1ececac62a8361fe4
unrolled_ast: 6d33052975f037a698138d1e3d7bb06ac22669c9f92997b1ececac62a8361fe4
ssa_ast: e0f62e9ed02da031b11dbffb2e0e928ea2fbda7c2f0eb28013068abdaab521dc
flattened_ast: bf48feece0821904ba575f8a4de972d1f72631e102f2e6968c69a364fd70ea1d
- initial_ast: 5a6e8de67b7ddf2a34dd7bf302b73001f4cdf9f64517806622ed3787d35eadcc
unrolled_ast: 5a6e8de67b7ddf2a34dd7bf302b73001f4cdf9f64517806622ed3787d35eadcc
ssa_ast: 4d554be19b3abc0c13bccd5225bc75011e8a77ec764de362788b9d221bae09c4
flattened_ast: faf9aeb6811760108ec7fd83ddf1d08314719c707e9818bebd9cf9b7e0adbff2
inlined_ast: faf9aeb6811760108ec7fd83ddf1d08314719c707e9818bebd9cf9b7e0adbff2
bytecode: 76a90286cb4903577bb9b0d219abe140fd8e2ef8a74df48a82d986e8efc4235d

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: df120b5592ffccee9259d4dd27c40c0d1b960abc37c342057aeb73e0c01ca8bf
unrolled_ast: df120b5592ffccee9259d4dd27c40c0d1b960abc37c342057aeb73e0c01ca8bf
ssa_ast: df120b5592ffccee9259d4dd27c40c0d1b960abc37c342057aeb73e0c01ca8bf
flattened_ast: df120b5592ffccee9259d4dd27c40c0d1b960abc37c342057aeb73e0c01ca8bf
- initial_ast: 6a343455b83835e5e2bc4760238b988d508dbf8a73b078f95bbd92c825f931bc
unrolled_ast: 6a343455b83835e5e2bc4760238b988d508dbf8a73b078f95bbd92c825f931bc
ssa_ast: 6a343455b83835e5e2bc4760238b988d508dbf8a73b078f95bbd92c825f931bc
flattened_ast: 6a343455b83835e5e2bc4760238b988d508dbf8a73b078f95bbd92c825f931bc
inlined_ast: 6a343455b83835e5e2bc4760238b988d508dbf8a73b078f95bbd92c825f931bc
bytecode: a26eca302425b77f7d017763631062a040d57f8557dd53a31bfe4d17584ab0e2

View File

@ -2,8 +2,9 @@
namespace: Compile
expectation: Pass
outputs:
- initial_ast: b7ac599b0560dd5b3282817ed32b8b22d0999026b0dfa78a8cd2b90691fad917
unrolled_ast: b7ac599b0560dd5b3282817ed32b8b22d0999026b0dfa78a8cd2b90691fad917
ssa_ast: b7ac599b0560dd5b3282817ed32b8b22d0999026b0dfa78a8cd2b90691fad917
flattened_ast: b7ac599b0560dd5b3282817ed32b8b22d0999026b0dfa78a8cd2b90691fad917
- initial_ast: 46c8f1a9f3ac3b544211d70cbf18eb2c30e659096e36309b268ecebfb8901047
unrolled_ast: 46c8f1a9f3ac3b544211d70cbf18eb2c30e659096e36309b268ecebfb8901047
ssa_ast: 46c8f1a9f3ac3b544211d70cbf18eb2c30e659096e36309b268ecebfb8901047
flattened_ast: 46c8f1a9f3ac3b544211d70cbf18eb2c30e659096e36309b268ecebfb8901047
inlined_ast: 46c8f1a9f3ac3b544211d70cbf18eb2c30e659096e36309b268ecebfb8901047
bytecode: 8f6238b1942bb3cf2eb7d0eed9745dffaf088c884c423992f0d23b989f3954ff

Some files were not shown because too many files have changed in this diff Show More