Merge pull request #1946 from AleoHQ/feat/ssa

Static Single Assignment
This commit is contained in:
Collin Chin 2022-07-28 14:15:37 -07:00 committed by GitHub
commit eb63f13cf7
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
352 changed files with 1374 additions and 189 deletions

View File

@ -194,8 +194,8 @@ pub trait StatementReconstructor: ExpressionReconstructor {
fn reconstruct_conditional(&mut self, input: ConditionalStatement) -> Statement {
Statement::Conditional(ConditionalStatement {
condition: self.reconstruct_expression(input.condition).0,
block: self.reconstruct_block(input.block),
next: input.next.map(|n| Box::new(self.reconstruct_statement(*n))),
then: self.reconstruct_block(input.then),
otherwise: input.otherwise.map(|n| Box::new(self.reconstruct_statement(*n))),
span: input.span,
})
}

View File

@ -140,8 +140,8 @@ pub trait StatementVisitor<'a>: ExpressionVisitor<'a> {
fn visit_conditional(&mut self, input: &'a ConditionalStatement) {
self.visit_expression(&input.condition, &Default::default());
self.visit_block(&input.block);
if let Some(stmt) = input.next.as_ref() {
self.visit_block(&input.then);
if let Some(stmt) = input.otherwise.as_ref() {
self.visit_statement(stmt);
}
}

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 crate::{Expression, Node};
use crate::{BinaryOperation, Expression, Node};
use leo_span::Span;
use serde::{Deserialize, Serialize};
@ -47,12 +47,34 @@ pub enum AssignOperation {
BitXor,
/// Shift right assignment.
Shr,
/// Signed shift right assignment.
ShrSigned,
// /// Signed shift right assignment.
// ShrSigned,
/// Shift left assignment.
Shl,
/// Modulus / remainder assignment.
Mod,
// /// Modulus / remainder assignment.
// Mod,
}
impl AssignOperation {
pub fn into_binary_operation(assign_op: AssignOperation) -> Option<BinaryOperation> {
match assign_op {
AssignOperation::Assign => None,
AssignOperation::Add => Some(BinaryOperation::Add),
AssignOperation::Sub => Some(BinaryOperation::Sub),
AssignOperation::Mul => Some(BinaryOperation::Mul),
AssignOperation::Div => Some(BinaryOperation::Div),
AssignOperation::Pow => Some(BinaryOperation::Pow),
AssignOperation::Or => Some(BinaryOperation::Or),
AssignOperation::And => Some(BinaryOperation::And),
AssignOperation::BitOr => Some(BinaryOperation::BitwiseOr),
AssignOperation::BitAnd => Some(BinaryOperation::BitwiseAnd),
AssignOperation::BitXor => Some(BinaryOperation::Xor),
AssignOperation::Shr => Some(BinaryOperation::Shr),
// AssignOperation::ShrSigned => Some(BinaryOperation::ShrSigned),
AssignOperation::Shl => Some(BinaryOperation::Shl),
// AssignOperation::Mod => Some(BinaryOperation::Mod),
}
}
}
impl AsRef<str> for AssignOperation {
@ -70,9 +92,9 @@ impl AsRef<str> for AssignOperation {
AssignOperation::BitAnd => "&=",
AssignOperation::BitXor => "^=",
AssignOperation::Shr => ">>=",
AssignOperation::ShrSigned => ">>>=",
// AssignOperation::ShrSigned => ">>>=",
AssignOperation::Shl => "<<=",
AssignOperation::Mod => "%=",
// AssignOperation::Mod => "%=",
}
}
}

View File

@ -26,17 +26,17 @@ pub struct ConditionalStatement {
/// The `bool`-typed condition deciding what to evaluate.
pub condition: Expression,
/// The block to evaluate in case `condition` yields `true`.
pub block: Block,
pub then: Block,
/// The statement, if any, to evaluate when `condition` yields `false`.
pub next: Option<Box<Statement>>,
pub otherwise: Option<Box<Statement>>,
/// The span from `if` to `next` or to `block`.
pub span: Span,
}
impl fmt::Display for ConditionalStatement {
fn fmt(&self, f: &mut fmt::Formatter) -> fmt::Result {
write!(f, "if ({}) {}", self.condition, self.block)?;
match self.next.clone() {
write!(f, "if ({}) {}", self.condition, self.then)?;
match self.otherwise.as_ref() {
Some(n_or_e) => write!(f, " else {}", n_or_e),
None => write!(f, ""),
}

View File

@ -95,20 +95,13 @@ impl<'a> Compiler<'a> {
let prg_sf = with_session_globals(|s| s.source_map.new_source(program_string, name));
// Use the parser to construct the abstract syntax tree (ast).
let mut ast: leo_ast::Ast = leo_parser::parse_ast(self.handler, &prg_sf.src, prg_sf.start_pos)?;
let mut ast: Ast = leo_parser::parse_ast(self.handler, &prg_sf.src, prg_sf.start_pos)?;
ast = ast.set_program_name(self.program_name.clone());
ast = ast.set_network(self.network.clone());
self.ast = ast.set_network(self.network.clone());
if self.output_options.initial_ast {
// Write the AST snapshot post parsing.
if self.output_options.spans_enabled {
ast.to_json_file(self.output_directory.clone(), "initial_ast.json")?;
} else {
ast.to_json_file_without_keys(self.output_directory.clone(), "initial_ast.json", &["span"])?;
self.write_ast_to_json("initial_ast.json")?;
}
}
self.ast = ast;
Ok(())
}
@ -171,6 +164,17 @@ impl<'a> Compiler<'a> {
Ok(symbol_table)
}
/// Runs the static single assignment pass.
pub fn static_single_assignment_pass(&mut self) -> Result<()> {
self.ast = StaticSingleAssigner::do_pass((std::mem::take(&mut self.ast), self.handler))?;
if self.output_options.ssa_ast {
self.write_ast_to_json("ssa_ast.json")?;
}
Ok(())
}
/// Runs the compiler stages.
pub fn compiler_stages(&mut self) -> Result<SymbolTable> {
let st = self.symbol_table_pass()?;
@ -178,6 +182,10 @@ impl<'a> Compiler<'a> {
// TODO: Make this pass optional.
let st = self.loop_unrolling_pass(st)?;
// TODO: Make this pass optional.
self.static_single_assignment_pass()?;
Ok(st)
}

View File

@ -24,4 +24,6 @@ pub struct OutputOptions {
pub initial_input_ast: bool,
/// If enabled writes the AST after loop unrolling.
pub unrolled_ast: bool,
/// If enabled writes the AST after static single assignment.
pub ssa_ast: bool,
}

View File

@ -51,6 +51,7 @@ fn new_compiler(handler: &Handler, main_file_path: PathBuf) -> Compiler<'_> {
initial_input_ast: true,
initial_ast: true,
unrolled_ast: true,
ssa_ast: true,
}),
)
}
@ -106,6 +107,7 @@ struct CompileOutput {
pub output: Vec<OutputItem>,
pub initial_ast: String,
pub unrolled_ast: String,
pub ssa_ast: String,
}
/// Get the path of the `input_file` given in `input` into `list`.
@ -139,6 +141,7 @@ fn compile_and_process<'a>(parsed: &'a mut Compiler<'a>) -> Result<SymbolTable,
let st = parsed.symbol_table_pass()?;
let st = parsed.type_checker_pass(st)?;
let st = parsed.loop_unrolling_pass(st)?;
parsed.static_single_assignment_pass()?;
Ok(st)
}
@ -219,6 +222,7 @@ fn run_test(test: Test, handler: &Handler, err_buf: &BufferEmitter) -> Result<Va
let initial_ast = hash_file("/tmp/output/initial_ast.json");
let unrolled_ast = hash_file("/tmp/output/unrolled_ast.json");
let ssa_ast = hash_file("/tmp/output/ssa_ast.json");
if fs::read_dir("/tmp/output").is_ok() {
fs::remove_dir_all(Path::new("/tmp/output")).expect("Error failed to clean up output dir.");
@ -228,6 +232,7 @@ fn run_test(test: Test, handler: &Handler, err_buf: &BufferEmitter) -> Result<Va
output: output_items,
initial_ast,
unrolled_ast,
ssa_ast,
};
Ok(serde_yaml::to_value(&final_output).expect("serialization failed"))
}

View File

@ -93,8 +93,8 @@ impl ParserContext<'_> {
Ok(ConditionalStatement {
span: start + next.as_ref().map(|x| x.span()).unwrap_or(body.span),
condition: expr,
block: body,
next,
then: body,
otherwise: next,
})
}

View File

@ -17,8 +17,8 @@
use crate::CodeGenerator;
use leo_ast::{
AssignStatement, Block, ConditionalStatement, ConsoleStatement, DefinitionStatement, IterationStatement,
ReturnStatement, Statement,
AssignStatement, Block, ConditionalStatement, ConsoleStatement, DefinitionStatement, Expression,
IterationStatement, ReturnStatement, Statement,
};
use itertools::Itertools;
@ -61,21 +61,27 @@ impl<'a> CodeGenerator<'a> {
expression_instructions
}
fn visit_assign(&mut self, _input: &'a AssignStatement) -> String {
// Note: SSA form requires that a variable is assigned to only once.
// Since we do not have a compiler pass that transforms the AST into SSA form,
// we will disallow `AssignStatement`s. This will only be the case for this prototype.
unimplemented!("Code generation is not implemented for `AssignStatement`s.")
fn visit_assign(&mut self, input: &'a AssignStatement) -> String {
// TODO: Once SSA is made optional, this should be made optional.
match &input.place {
Expression::Identifier(identifier) => {
let (operand, expression_instructions) = self.visit_expression(&input.value);
self.variable_mapping.insert(&identifier.name, operand);
expression_instructions
}
_ => unimplemented!(
"Code generation for the left-hand side of an assignment is only implemented for `Identifier`s."
),
}
}
fn visit_conditional(&mut self, _input: &'a ConditionalStatement) -> String {
// Note: This requires that the AST is in static-single assignment form.
// It is not possible to provide an input program with a conditional statement in SSA form as
// complete SSA has different semantics from source Leo programs.
unimplemented!("Code generation is not implemented for conditional statements.")
// TODO: Once SSA is made optional, create a Leo error informing the user to enable the SSA pass.
unreachable!("`ConditionalStatement`s should not be in the AST at this phase of compilation.")
}
fn visit_iteration(&mut self, _input: &'a IterationStatement) -> String {
// TODO: Once loop unrolling is made optional, create a Leo error informing the user to enable the loop unrolling pass..
unreachable!("`IterationStatement`s should not be in the AST at this phase of compilation.");
}

View File

@ -25,6 +25,9 @@ pub use self::pass::*;
pub mod loop_unrolling;
pub use self::loop_unrolling::*;
pub mod static_single_assignment;
pub use static_single_assignment::*;
pub mod symbol_table;
pub use symbol_table::*;

View File

@ -0,0 +1,76 @@
// Copyright (C) 2019-2022 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 Static Single Assignment pass traverses the AST and converts it into SSA form.
//! See https://en.wikipedia.org/wiki/Static_single-assignment_form for more information.
//! The pass also flattens `ConditionalStatement`s into a sequence of `AssignStatement`s.
//! The pass also rewrites complex `AssignStatement`s into simpler ones. e.g x += 1 -> x = x + 1.
//! The pass also rewrites `ReturnStatement`s into `AssignStatement`s and consolidates the returned values as a single `ReturnStatement` at the end of the function.
//!
//! Consider the following Leo code.
//! ```leo
//! function main(flag: u8, value: u8) -> u8 {
//! if (flag == 0u8) {
//! value += 1u8;
//! return value;
//! } else {
//! value += 2u8;
//! }
//! return value;
//! }
//! ```
//!
//! The SSA pass produces the following code.
//! ```leo
//! function main(flag: u8, value: u8) -> u8 {
//! let $cond$0 = flag == 0u8;
//! let value$1 = value + 1u8;
//! let $return$2 = value$1;
//! let value$3 = value + 2u8;
//! let value$4 = $cond$0 ? value$1 : value$4;
//! let $return$5 = value$4;
//! return $cond$0 ? $return$2 : $return$5;
//! ```
mod rename_expression;
mod rename_program;
mod rename_statement;
mod rename_table;
pub(crate) use rename_table::*;
pub mod static_single_assigner;
pub use static_single_assigner::*;
use crate::Pass;
use leo_ast::{Ast, ProgramReconstructor};
use leo_errors::{emitter::Handler, Result};
impl<'a> Pass for StaticSingleAssigner<'a> {
type Input = (Ast, &'a Handler);
type Output = Result<Ast>;
fn do_pass((ast, handler): Self::Input) -> Self::Output {
let mut reconstructor = StaticSingleAssigner::new(handler);
let program = reconstructor.reconstruct_program(ast.into_repr());
handler.last_err()?;
Ok(Ast::new(program))
}
}

View File

@ -0,0 +1,95 @@
// Copyright (C) 2019-2022 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::StaticSingleAssigner;
use leo_ast::{
CallExpression, CircuitExpression, CircuitVariableInitializer, Expression, ExpressionReconstructor, Identifier,
};
impl ExpressionReconstructor for StaticSingleAssigner<'_> {
type AdditionalOutput = ();
/// Reconstructs `CallExpression` without visiting the function name.
fn reconstruct_call(&mut self, input: CallExpression) -> (Expression, Self::AdditionalOutput) {
(
Expression::Call(CallExpression {
// Note that we do not rename the function name.
function: input.function,
arguments: input
.arguments
.into_iter()
.map(|arg| self.reconstruct_expression(arg).0)
.collect(),
span: input.span,
}),
Default::default(),
)
}
/// Produces a new `CircuitExpression` with renamed variables.
fn reconstruct_circuit_init(&mut self, input: CircuitExpression) -> (Expression, Self::AdditionalOutput) {
(
Expression::Circuit(CircuitExpression {
name: input.name,
span: input.span,
members: input
.members
.into_iter()
.map(|arg| CircuitVariableInitializer {
identifier: arg.identifier,
expression: Some(match &arg.expression.is_some() {
// If the expression is None, then `arg` is a `CircuitVariableInitializer` of the form `<id>,`.
// In this case, we must reconstruct the identifier and produce an initializer of the form `<id>: <renamed_id>`.
false => self.reconstruct_identifier(arg.identifier).0,
// If expression is `Some(..)`, then `arg is a `CircuitVariableInitializer` of the form `<id>: <expr>,`.
// In this case, we must reconstruct the expression.
true => self.reconstruct_expression(arg.expression.unwrap()).0,
}),
})
.collect(),
}),
Default::default(),
)
}
/// Produces a new `Identifier` with a unique name.
fn reconstruct_identifier(&mut self, identifier: Identifier) -> (Expression, Self::AdditionalOutput) {
let name = match self.is_lhs {
// If reconstructing the left-hand side of a definition or assignment, a new unique name is introduced.
true => {
let new_name = self.unique_symbol(identifier.name);
self.rename_table.update(identifier.name, new_name);
new_name
}
// Otherwise, we look up the previous name in the `RenameTable`.
false => *self.rename_table.lookup(identifier.name).unwrap_or_else(|| {
panic!(
"SSA Error: An entry in the `RenameTable` for {} should exist.",
identifier.name
)
}),
};
(
Expression::Identifier(Identifier {
name,
span: identifier.span,
}),
Default::default(),
)
}
}

View File

@ -0,0 +1,84 @@
// Copyright (C) 2019-2022 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::StaticSingleAssigner;
use leo_ast::{
Expression, Function, FunctionInput, ProgramReconstructor, ReturnStatement, Statement, StatementReconstructor,
TernaryExpression,
};
impl ProgramReconstructor for StaticSingleAssigner<'_> {
/// Reconstructs the `Function`s in the `Program`, while allocating the appropriate `RenameTable`s.
fn reconstruct_function(&mut self, function: Function) -> Function {
// Allocate a `RenameTable` for the function.
self.push();
// There is no need to reconstruct `function.inputs`.
// However, for each input, we must add each symbol to the rename table.
for input in function.input.iter() {
match input {
FunctionInput::Variable(function_input_variable) => {
self.rename_table.update(
function_input_variable.identifier.name,
function_input_variable.identifier.name,
);
}
}
}
let mut block = self.reconstruct_block(function.block);
// Add the `ReturnStatement` to the end of the block.
let mut returns = self.clear_early_returns();
// Type checking guarantees that there exists at least one return statement in the function body.
let (_, last_return_expression) = returns.pop().unwrap();
// Fold all return expressions into a single ternary expression.
let expression =
returns
.into_iter()
.rev()
.fold(last_return_expression, |acc, (guard, expression)| match guard {
None => unreachable!("All return statements except for the last one must have a guard."),
Some(guard) => Expression::Ternary(TernaryExpression {
condition: Box::new(guard),
if_true: Box::new(expression),
if_false: Box::new(acc),
span: Default::default(),
}),
});
// Add the `ReturnStatement` to the end of the block.
block.statements.push(Statement::Return(ReturnStatement {
expression,
span: Default::default(),
}));
// Remove the `RenameTable` for the function.
self.pop();
Function {
identifier: function.identifier,
input: function.input,
output: function.output,
core_mapping: function.core_mapping,
block,
span: function.span,
}
}
}

View File

@ -0,0 +1,225 @@
// Copyright (C) 2019-2022 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::{RenameTable, StaticSingleAssigner};
use leo_ast::{
AssignOperation, AssignStatement, BinaryExpression, BinaryOperation, Block, ConditionalStatement,
DefinitionStatement, Expression, ExpressionReconstructor, Identifier, Node, ReturnStatement, Statement,
StatementReconstructor, TernaryExpression, UnaryExpression, UnaryOperation,
};
use leo_span::Symbol;
use indexmap::IndexSet;
impl StatementReconstructor for StaticSingleAssigner<'_> {
/// Transforms a `ReturnStatement` into an `AssignStatement`,
/// storing the variable and the associated guard in `self.early_returns`.
///
/// Note that this pass assumes that there is at most one `ReturnStatement` in a block.
fn reconstruct_return(&mut self, input: ReturnStatement) -> Statement {
// Create a fresh name for the expression in the return statement.
let symbol = self.unique_symbol("$return");
self.rename_table.update(symbol, symbol);
// Initialize a new `AssignStatement` for the return expression.
let place = Expression::Identifier(Identifier::new(symbol));
// Add the variable and associated guard.
let guard = match self.condition_stack.is_empty() {
true => None,
false => {
let (first, rest) = self.condition_stack.split_first().unwrap();
Some(rest.iter().cloned().fold(first.clone(), |acc, condition| {
Expression::Binary(BinaryExpression {
op: BinaryOperation::And,
left: Box::new(acc),
right: Box::new(condition),
span: Default::default(),
})
}))
}
};
self.early_returns.push((guard, place.clone()));
Self::simple_assign_statement(place, self.reconstruct_expression(input.expression).0)
}
/// Reconstructs the `DefinitionStatement` into an `AssignStatement`, renaming the left-hand-side as appropriate.
fn reconstruct_definition(&mut self, definition: DefinitionStatement) -> Statement {
self.is_lhs = true;
let identifier = match self.reconstruct_identifier(definition.variable_name).0 {
Expression::Identifier(identifier) => identifier,
_ => unreachable!("`self.reconstruct_identifier` will always return an `Identifier`."),
};
self.is_lhs = false;
Self::simple_assign_statement(
Expression::Identifier(identifier),
self.reconstruct_expression(definition.value).0,
)
}
/// Transform all `AssignStatement`s to simple `AssignStatement`s.
/// For example,
/// `x += y * 3` becomes `x = x + (y * 3)`
/// `x &= y | 1` becomes `x = x & (y | 1)`
/// `x = y + 3` remains `x = y + 3`
fn reconstruct_assign(&mut self, assign: AssignStatement) -> Statement {
self.is_lhs = true;
let place = self.reconstruct_expression(assign.place).0;
self.is_lhs = false;
let value = match assign.operation {
AssignOperation::Assign => self.reconstruct_expression(assign.value).0,
// Note that all `AssignOperation`s except for the `Assign` variant have an equivalent `BinaryOperation`.
_ => Expression::Binary(BinaryExpression {
left: Box::new(place.clone()),
right: Box::new(self.reconstruct_expression(assign.value).0),
op: AssignOperation::into_binary_operation(assign.operation).unwrap(),
span: assign.span,
}),
};
Self::simple_assign_statement(place, value)
}
/// Reconstructs a `ConditionalStatement`, producing phi functions for variables written in the then-block and otherwise-block.
fn reconstruct_conditional(&mut self, conditional: ConditionalStatement) -> Statement {
let condition = self.reconstruct_expression(conditional.condition).0;
// Instantiate a `RenameTable` for the then-block.
self.push();
// Add condition to the condition stack.
self.condition_stack.push(condition.clone());
// Reconstruct the then-block.
let then = self.reconstruct_block(conditional.then);
// Remove condition from the condition stack.
self.condition_stack.pop();
// Remove the `RenameTable` for the then-block.
let if_table = self.pop();
// Instantiate a `RenameTable` for the otherwise-block.
self.push();
// Reconstruct the otherwise-block.
let otherwise = conditional.otherwise.map(|statement| {
// Add the negated condition to the condition stack.
self.condition_stack.push(Expression::Unary(UnaryExpression {
op: UnaryOperation::Not,
receiver: Box::new(condition.clone()),
span: condition.span(),
}));
let reconstructed_block = Box::new(match *statement {
// The `ConditionalStatement` must be reconstructed as a `Block` statement to ensure that appropriate statements are produced.
Statement::Conditional(stmt) => self.reconstruct_statement(Statement::Block(Block {
statements: vec![Statement::Conditional(stmt)],
span: Default::default(),
})),
Statement::Block(stmt) => self.reconstruct_statement(Statement::Block(stmt)),
_ => unreachable!(
"`ConditionalStatement`s next statement must be a `ConditionalStatement` or a `Block`."
),
});
// Remove the negated condition from the condition stack.
self.condition_stack.pop();
reconstructed_block
});
// Remove the `RenameTable` for the otherwise-block.
let else_table = self.pop();
// Compute the write set for the variables written in the then-block or otherwise-block.
let if_write_set: IndexSet<&Symbol> = IndexSet::from_iter(if_table.local_names());
let else_write_set: IndexSet<&Symbol> = IndexSet::from_iter(else_table.local_names());
let write_set = if_write_set.union(&else_write_set);
// For each variable in the write set, instantiate a phi function.
for symbol in write_set {
// Note that phi functions only need to be instantiated if the variable exists before the `ConditionalStatement`.
if self.rename_table.lookup(**symbol).is_some() {
// Helper to lookup a symbol and create an argument for the phi function.
let create_phi_argument = |table: &RenameTable, symbol: Symbol| {
let name = *table
.lookup(symbol)
.unwrap_or_else(|| panic!("Symbol {} should exist in the program.", symbol));
Box::new(Expression::Identifier(Identifier {
name,
span: Default::default(),
}))
};
// Create a new name for the variable written to in the `ConditionalStatement`.
let new_name = self.unique_symbol(symbol);
// Create a new `AssignStatement` for the phi function.
let assignment = Self::simple_assign_statement(
Expression::Identifier(Identifier {
name: new_name,
span: Default::default(),
}),
Expression::Ternary(TernaryExpression {
condition: Box::new(condition.clone()),
if_true: create_phi_argument(&if_table, **symbol),
if_false: create_phi_argument(&else_table, **symbol),
span: Default::default(),
}),
);
// Update the `RenameTable` with the new name of the variable.
self.rename_table.update(*(*symbol), new_name);
// Store the generate phi functions.
self.phi_functions.push(assignment);
}
}
// Note that we only produce
Statement::Conditional(ConditionalStatement {
condition,
then,
otherwise,
span: conditional.span,
})
}
/// Reconstructs a `Block`, flattening its constituent `ConditionalStatement`s.
fn reconstruct_block(&mut self, block: Block) -> Block {
let mut statements = Vec::with_capacity(block.statements.len());
// Reconstruct each statement in the block.
for statement in block.statements.into_iter() {
match statement {
Statement::Conditional(conditional_statement) => {
statements.extend(self.flatten_conditional_statement(conditional_statement))
}
_ => statements.push(self.reconstruct_statement(statement)),
}
}
Block {
statements,
span: block.span,
}
}
}

View File

@ -0,0 +1,60 @@
// Copyright (C) 2019-2022 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_span::Symbol;
use indexmap::IndexMap;
/// `RenameTable` tracks the names assigned by static single assignment in a single scope.
#[derive(Clone, Debug, Default, Eq, PartialEq)]
pub(crate) struct RenameTable {
/// The `RenameTable` of the parent scope.
pub(crate) parent: Option<Box<RenameTable>>,
/// The mapping from names in the original AST to new names in the renamed AST.
mapping: IndexMap<Symbol, Symbol>,
}
impl RenameTable {
/// Create a new `RenameTable` with the given parent.
pub(crate) fn new(parent: Option<Box<RenameTable>>) -> Self {
Self {
parent,
mapping: IndexMap::new(),
}
}
/// Returns the symbols that were renamed in the current scope.
pub(crate) fn local_names(&self) -> impl Iterator<Item = &Symbol> {
self.mapping.keys()
}
/// Updates `self.mapping` with the desired entry.
/// Creates a new entry if `symbol` is not already in `self.mapping`.
pub(crate) fn update(&mut self, symbol: Symbol, new_symbol: Symbol) {
self.mapping.insert(symbol, new_symbol);
}
/// Looks up the new name for `symbol`, recursively checking the parent if it is not found.
pub(crate) fn lookup(&self, symbol: Symbol) -> Option<&Symbol> {
if let Some(var) = self.mapping.get(&symbol) {
Some(var)
} else if let Some(parent) = &self.parent {
parent.lookup(symbol)
} else {
None
}
}
}

View File

@ -0,0 +1,165 @@
// Copyright (C) 2019-2022 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::RenameTable;
use std::fmt::Display;
use leo_ast::{
AssignOperation, AssignStatement, ConditionalStatement, Expression, ExpressionReconstructor, Identifier, Statement,
StatementReconstructor,
};
use leo_errors::emitter::Handler;
use leo_span::Symbol;
pub struct StaticSingleAssigner<'a> {
/// The `RenameTable` for the current basic block in the AST
pub(crate) rename_table: RenameTable,
/// An error handler used for any errors found during unrolling.
pub(crate) _handler: &'a Handler,
/// A strictly increasing counter, used to ensure that new variable names are unique.
pub(crate) counter: usize,
/// 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,
/// Phi functions produced by static single assignment.
pub(crate) phi_functions: Vec<Statement>,
/// A stack of condition `Expression`s visited up to the current point in the AST.
pub(crate) condition_stack: Vec<Expression>,
/// A list containing tuples of guards and expressions associated with early `ReturnStatement`s.
pub(crate) early_returns: Vec<(Option<Expression>, Expression)>,
}
impl<'a> StaticSingleAssigner<'a> {
pub(crate) fn new(handler: &'a Handler) -> Self {
Self {
rename_table: RenameTable::new(None),
_handler: handler,
counter: 0,
is_lhs: false,
phi_functions: Vec::new(),
condition_stack: Vec::new(),
early_returns: Vec::new(),
}
}
/// Return a new unique `Symbol` from a `&str`.
pub(crate) fn unique_symbol(&mut self, arg: impl Display) -> Symbol {
self.counter += 1;
Symbol::intern(&format!("{}${}", arg, self.counter - 1))
}
/// Constructs the assignment statement `place = expr;`.
pub(crate) fn simple_assign_statement(place: Expression, value: Expression) -> Statement {
Statement::Assign(Box::new(AssignStatement {
operation: AssignOperation::Assign,
place,
value,
span: Default::default(),
}))
}
/// Clears the `self.phi_functions`, returning the ones that were previously produced.
pub(crate) fn clear_phi_functions(&mut self) -> Vec<Statement> {
core::mem::take(&mut self.phi_functions)
}
/// Clears the state associated with `ReturnStatements`, returning the ones that were previously produced.
pub(crate) fn clear_early_returns(&mut self) -> Vec<(Option<Expression>, Expression)> {
core::mem::take(&mut self.early_returns)
}
/// Pushes a new scope, setting the current scope as the new scope's parent.
pub(crate) fn push(&mut self) {
let parent_table = core::mem::take(&mut self.rename_table);
self.rename_table = RenameTable::new(Some(Box::from(parent_table)));
}
/// If the RenameTable has a parent, then `self.rename_table` is set to the parent, otherwise it is set to a default `RenameTable`.
pub(crate) fn pop(&mut self) -> RenameTable {
let parent = self.rename_table.parent.clone().unwrap_or_default();
core::mem::replace(&mut self.rename_table, *parent)
}
/// Introduces a new `AssignStatement` for non-trivial expressions in the condition of `ConditionalStatement`s.
/// For example,
/// - `if x > 0 { x = x + 1 }` becomes `let $cond$0 = x > 0; if $cond$0 { x = x + 1; }`
/// - `if true { x = x + 1 }` remains the same.
/// - `if b { x = x + 1 }` remains the same.
/// And then reconstructs and flattens `ConditionalStatement`.
pub(crate) fn flatten_conditional_statement(
&mut self,
conditional_statement: ConditionalStatement,
) -> Vec<Statement> {
let mut statements = Vec::new();
// Reconstruct the `ConditionalStatement`.
let reconstructed_statement = match conditional_statement.condition {
Expression::Err(_) => {
unreachable!("Err expressions should not exist in the AST at this stage of compilation.")
}
Expression::Identifier(..) | Expression::Literal(..) => self.reconstruct_conditional(conditional_statement),
// If the condition is a complex expression, introduce a new `AssignStatement` for it.
Expression::Access(..)
| Expression::Call(..)
| Expression::Circuit(..)
| Expression::Tuple(..)
| Expression::Binary(..)
| Expression::Unary(..)
| Expression::Ternary(..) => {
// Create a fresh variable name for the condition.
let symbol = self.unique_symbol("$cond$");
self.rename_table.update(symbol, symbol);
// Initialize a new `AssignStatement` for the condition.
let place = Expression::Identifier(Identifier::new(symbol));
let assign_statement = Self::simple_assign_statement(
place.clone(),
self.reconstruct_expression(conditional_statement.condition).0,
);
let rewritten_conditional_statement = ConditionalStatement {
condition: place,
then: conditional_statement.then,
otherwise: conditional_statement.otherwise,
span: conditional_statement.span,
};
statements.push(assign_statement);
self.reconstruct_conditional(rewritten_conditional_statement)
}
};
// Flatten the reconstructed `ConditionalStatement`
// by lifting the statements in the "if" and "else" block into their parent block.
let mut conditional_statement = match reconstructed_statement {
Statement::Conditional(conditional_statement) => conditional_statement,
_ => unreachable!("`reconstruct_conditional` will always produce a `ConditionalStatement`"),
};
statements.append(&mut conditional_statement.then.statements);
if let Some(statement) = conditional_statement.otherwise {
match *statement {
// If we encounter a `BlockStatement`,
// we need to lift its constituent statements into the current `BlockStatement`.
Statement::Block(mut block) => statements.append(&mut block.statements),
_ => unreachable!(
"`self.reconstruct_conditional` will always produce a `BlockStatement` in the next block."
),
}
}
// Add all phi functions to the current block.
statements.append(&mut self.clear_phi_functions());
statements
}
}

View File

@ -27,7 +27,7 @@ use crate::{FunctionSymbol, VariableSymbol};
#[derive(Clone, Debug, Default)]
pub struct SymbolTable {
/// The parent scope if it exists.
/// For example if we are in a if block inside a function.
/// For example, the parent scope of a then-block is the scope containing the associated ConditionalStatement.
pub(crate) parent: Option<Box<SymbolTable>>,
/// Functions represents the name of each function mapped to the AST's function definition.
/// This field is populated at a first pass.

View File

@ -22,6 +22,23 @@ use leo_errors::TypeCheckerError;
use std::cell::RefCell;
impl<'a> StatementVisitor<'a> for TypeChecker<'a> {
fn visit_statement(&mut self, input: &'a Statement) {
if self.has_return {
self.emit_err(TypeCheckerError::unreachable_code_after_return(input.span()));
return;
}
match input {
Statement::Return(stmt) => self.visit_return(stmt),
Statement::Definition(stmt) => self.visit_definition(stmt),
Statement::Assign(stmt) => self.visit_assign(stmt),
Statement::Conditional(stmt) => self.visit_conditional(stmt),
Statement::Iteration(stmt) => self.visit_iteration(stmt),
Statement::Console(stmt) => self.visit_console(stmt),
Statement::Block(stmt) => self.visit_block(stmt),
}
}
fn visit_return(&mut self, input: &'a ReturnStatement) {
// we can safely unwrap all self.parent instances because
// statements should always have some parent block
@ -94,10 +111,30 @@ impl<'a> StatementVisitor<'a> for TypeChecker<'a> {
fn visit_conditional(&mut self, input: &'a ConditionalStatement) {
self.visit_expression(&input.condition, &Some(Type::Boolean));
self.visit_block(&input.block);
if let Some(s) = input.next.as_ref() {
self.visit_statement(s)
let mut then_block_has_return = false;
let mut otherwise_block_has_return = false;
// Set the `has_return` flag for the then-block.
let previous_has_return = core::mem::replace(&mut self.has_return, then_block_has_return);
self.visit_block(&input.then);
// Store the `has_return` flag for the then-block.
then_block_has_return = self.has_return;
if let Some(s) = input.otherwise.as_ref() {
// Set the `has_return` flag for the otherwise-block.
self.has_return = otherwise_block_has_return;
self.visit_statement(s);
// Store the `has_return` flag for the otherwise-block.
otherwise_block_has_return = self.has_return;
}
// Restore the previous `has_return` flag.
self.has_return = previous_has_return || (then_block_has_return && otherwise_block_has_return);
}
fn visit_iteration(&mut self, input: &'a IterationStatement) {
@ -124,8 +161,16 @@ impl<'a> StatementVisitor<'a> for TypeChecker<'a> {
self.handler.emit_err(err);
}
let prior_has_return = core::mem::take(&mut self.has_return);
input.block.statements.iter().for_each(|s| self.visit_statement(s));
if self.has_return {
self.emit_err(TypeCheckerError::loop_body_contains_return(input.span()));
}
self.has_return = prior_has_return;
// Restore the previous scope.
let prev_st = *self.symbol_table.borrow_mut().parent.take().unwrap();
self.symbol_table

View File

@ -259,4 +259,18 @@ create_messages!(
msg: format!("Tuples are only allowed as function return types."),
help: None,
}
@formatted
unreachable_code_after_return {
args: (),
msg: format!("Cannot reach the following statement."),
help: Some("Remove the unreachable code.".to_string()),
}
@formatted
loop_body_contains_return {
args: (),
msg: format!("Loop body contains a return statement or always returns."),
help: Some("Remove the code in the loop body that always returns.".to_string()),
}
);

View File

@ -52,6 +52,8 @@ pub struct BuildOptions {
pub enable_initial_ast_snapshot: bool,
#[structopt(long, help = "Writes AST snapshot of the unrolled AST.")]
pub enable_unrolled_ast_snapshot: bool,
#[structopt(long, help = "Writes AST snapshot of the SSA AST.")]
pub enable_ssa_ast_snapshot: bool,
// Note: This is currently made optional since code generation is just a prototype.
#[structopt(
long,
@ -67,11 +69,13 @@ impl From<BuildOptions> for OutputOptions {
initial_input_ast: options.enable_initial_input_ast_snapshot,
initial_ast: options.enable_initial_ast_snapshot,
unrolled_ast: options.enable_unrolled_ast_snapshot,
ssa_ast: options.enable_ssa_ast_snapshot,
};
if options.enable_all_ast_snapshots {
out_options.initial_input_ast = true;
out_options.initial_ast = true;
out_options.unrolled_ast = true;
out_options.ssa_ast = true;
}
out_options

View File

@ -0,0 +1,22 @@
/*
namespace: Compile
expectation: Fail
input_file: inputs/u32_3.in
*/
function main(x: u32) -> bool {
for i: u32 in 0u32..9u32 {
return false;
}
for i: u32 in 0u32..9u32 {
if (x == 0u32) {
return false;
} else {
return true;
}
}
return x == 1u32;
}

View File

@ -0,0 +1,12 @@
/*
namespace: Compile
expectation: Fail
input_file:
- inputs/u32_3.in
*/
function main(x: u32) -> u32 {
return x;
let double: u32 = x + x;
return double;
}

View File

@ -0,0 +1,16 @@
/*
namespace: Compile
expectation: Fail
input_file:
- inputs/u32_3.in
*/
function main(x: u32) -> bool {
if x == 3u32 {
return true;
} else {
return false;
}
let double: u32 = x + x;
return double;
}

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: fe880c907d0257c9fc8314b8b98cabd8a8282b587d2d618408cc3cd8e528fda5
initial_ast: 3c29b1703692d5095c7d346a8597e7bc96ab6b272d344d4873359aec19a90238
unrolled_ast: 3c29b1703692d5095c7d346a8597e7bc96ab6b272d344d4873359aec19a90238
ssa_ast: a9a6c516513ba1cc3ddb1c5d88c640f74014b2bff08dc460a473466cc57f9ac6

View File

@ -4,5 +4,6 @@ expectation: Pass
outputs:
- output:
- initial_input_ast: 00f5aba05e4efae5a125eb52f02f16400132085b8a34919d910aa40c6c405a22
initial_ast: 90154a6a01ee247ee0336cb3abcced38a628a9918e0554127dfe5ca7caa34dc7
unrolled_ast: 90154a6a01ee247ee0336cb3abcced38a628a9918e0554127dfe5ca7caa34dc7
initial_ast: a780252484ea83fe082982669ee2639ee1d5daab6a32a2638d6fbd47e8385f4d
unrolled_ast: a780252484ea83fe082982669ee2639ee1d5daab6a32a2638d6fbd47e8385f4d
ssa_ast: ec9190a9ddb0def4b88241078139781759c90db13ec5b4042febf02252d6744f

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 03e9df3bd1409f4af9e2a7f55130bc52f27d41f32a624ffa27f0ab114bf6fbf4
initial_ast: e0fde99efa3e6ec987044326b17bc96d16379992b2a30585ed85878285fb761c
unrolled_ast: e0fde99efa3e6ec987044326b17bc96d16379992b2a30585ed85878285fb761c
ssa_ast: 656cee892f013317582ee41fb9bad0d98de9ad9aa4853c2be4a7a16bab98a1ef

View File

@ -7,3 +7,4 @@ outputs:
- initial_input_ast: cb1d48114c10b2b732ad47a46fc8d05bf7a3e783da89e7f00065244bfc8d15c8
initial_ast: d36a542762bed92ce775751e99f5e47dbbc9014a0d4d00deb0fa8abb7af7f2e7
unrolled_ast: d36a542762bed92ce775751e99f5e47dbbc9014a0d4d00deb0fa8abb7af7f2e7
ssa_ast: 021c362f66fefa5b7507f2870872c8b286c0fe27f618d83b3670f1b1ebd59a16

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: a56b3f9908dec2acaed302691d4fe0c2cf046f0deb8f188f617e042e75502f71
initial_ast: 86e914325e67e129cc7e6738b7924fb1b9d72e49b008c8dd0c4da86d57d2132d
unrolled_ast: 86e914325e67e129cc7e6738b7924fb1b9d72e49b008c8dd0c4da86d57d2132d
ssa_ast: 3072bdb096990a6611e20c37908048ae20ab25247d1dc45b63f1a3e802d3864b

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: 650984ca5077d11a815889421656b7735b4c6bd320bdf68b4deb87dfc0f49388
initial_ast: fe582e7789f2199f81249864983873612bc4f24ae7086fac13f88515c2f89d81
unrolled_ast: fe582e7789f2199f81249864983873612bc4f24ae7086fac13f88515c2f89d81
ssa_ast: b70de2fb00bbd6124e529f323cdd10722d7538a7a345613e014c32f0ae1e74cf

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: a56b3f9908dec2acaed302691d4fe0c2cf046f0deb8f188f617e042e75502f71
initial_ast: e4dded56ee2ee6cbd2bbf301a3c87c09d35aafad9b66b8a7a8592f9d16aad47f
unrolled_ast: e4dded56ee2ee6cbd2bbf301a3c87c09d35aafad9b66b8a7a8592f9d16aad47f
ssa_ast: 5037ff9e6b508df19a5b1b3fe490353b88ab7e9ad65e62fbf1e3396694e169a0

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: a56b3f9908dec2acaed302691d4fe0c2cf046f0deb8f188f617e042e75502f71
initial_ast: 90bb6f3c3453976ec13875f39ad9ae5c852f3d24bd76e6c170e526f11f91b626
unrolled_ast: 90bb6f3c3453976ec13875f39ad9ae5c852f3d24bd76e6c170e526f11f91b626
ssa_ast: bb10ed5109f6b414cdded757644b2e1cdc1e23eebd17159170da60e8d46ffc38

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: d2fc1992beaf062678bbf6c3e862820dbbea39926589afcdc46c19c8669f0e37
initial_ast: be62eafa579fa7e239f5c9cd91bfea1e072b3db6192ff9a1c29be78dd9faf680
unrolled_ast: be62eafa579fa7e239f5c9cd91bfea1e072b3db6192ff9a1c29be78dd9faf680
ssa_ast: 4dd2fe14c60f053563a8dcc0cd647a02c17ad5587cfdfeb99202e59e36cd1acd

View File

@ -9,3 +9,4 @@ outputs:
- initial_input_ast: a56b3f9908dec2acaed302691d4fe0c2cf046f0deb8f188f617e042e75502f71
initial_ast: 907f4bebafce2bdcd67c2f8f00282e71b50b1fddc7144f5266e6012ef9f48bf5
unrolled_ast: 907f4bebafce2bdcd67c2f8f00282e71b50b1fddc7144f5266e6012ef9f48bf5
ssa_ast: 68e43f288ac90377cad0422763c6bc6c1cd5027dc1e567d942f4807d319059b5

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: f1af7e79dff9ede0d2a1c88d5d22801cb3dfe3a9fb34e93bca646e29a61e9f65
initial_ast: 7d8c6817927782cd181a5a3429f72139b452ec511ec5352dfe27b98a3b7b5955
unrolled_ast: 7d8c6817927782cd181a5a3429f72139b452ec511ec5352dfe27b98a3b7b5955
ssa_ast: dcde4a7790dc3be6a15ed832a887d894ccc6d242863c64d92ecbd211b2fe55da

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: d1b4db4f3be418df4f8b42d12e90491aa6fa4f68f07a3b0ac3b4ef6eb8091b17
unrolled_ast: d1b4db4f3be418df4f8b42d12e90491aa6fa4f68f07a3b0ac3b4ef6eb8091b17
ssa_ast: 66ced46da16574d00b77598212ae5c19c60458cbded763122fb3a976b08fe7a2

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 29f6139d908d390f890f04d8ee620757d29b7f71cd48c46ff65bc1e70aae840c
initial_ast: 930940759647baf7f6c4764b33bf6653fb3971ccdd5362466babf392788062bf
unrolled_ast: 930940759647baf7f6c4764b33bf6653fb3971ccdd5362466babf392788062bf
ssa_ast: cb48efda4262915f9b8583b29f018ad1de6cb929f726e6fff851604d2ee797de

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 15a1f00a6c0ca8141202e45e534b7afd196e9391c184a4efd94f0d0ccf04a59d
initial_ast: f8f5c01bf6a9bcaa57c44e48e4b1d873a526b7f24efce54928b5bd7d17568f54
unrolled_ast: f8f5c01bf6a9bcaa57c44e48e4b1d873a526b7f24efce54928b5bd7d17568f54
ssa_ast: 15bcf16d8ea9eea237ad103c1e2adb87bb7fbffa547d8b2e2cceed6e64cb7844

View File

@ -5,5 +5,6 @@ outputs:
- output:
- initial_input_ast: 8b94c0dbc84f44bd29c614b87947e625ad136549ea29ff18233ba5b01ce63c9b
- initial_input_ast: a62874e75304ab81d487909be1c6b1efa2e5756a2980b46e3bb1368586c3ee83
initial_ast: 4adda6e6c732a161071c21da51ba8902e7f351722971fd1834a560e5f672d2e8
unrolled_ast: 4adda6e6c732a161071c21da51ba8902e7f351722971fd1834a560e5f672d2e8
initial_ast: d37a2c8f119181a3c52d6d730f19a352aa4543dc1b8c346b42c35bffeff952ab
unrolled_ast: d37a2c8f119181a3c52d6d730f19a352aa4543dc1b8c346b42c35bffeff952ab
ssa_ast: 4a0a62e799fcf478a5007e660135bda9702d8e8f32882f62629794d8bee65d2a

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 14cd2c781b154a9037de84e945cfb348e9c587cef94d3e1f3be83e4306f92a0e
initial_ast: d8ad2aeb376197baa9983ae969d6e5d2a190675c8001a66a6062b7c3c9fda6b8
unrolled_ast: d8ad2aeb376197baa9983ae969d6e5d2a190675c8001a66a6062b7c3c9fda6b8
ssa_ast: 475f49a605d9267bdc6e2db77d27d7b6ef04c51ae5aaf6a9bc54bf7e531ca7a0

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: fd19d82c3aba921f01b37174e3eb7fb603438506fe511657e21235b9fb3647d2
initial_ast: 003f8bcd2065401f1ed21d8ca268a02a7da3c0f7dcbbfd1b145ba76b9cf9885d
unrolled_ast: 003f8bcd2065401f1ed21d8ca268a02a7da3c0f7dcbbfd1b145ba76b9cf9885d
ssa_ast: fa9d7e642da3e09e5407870565a7d58229911ac28411ae6f24bbad06d13ba12e

View File

@ -5,5 +5,6 @@ outputs:
- output:
- initial_input_ast: 12a0efa27e9b65c045088e471e6c254bb71c60cca4eb369f41e83a29301130cf
- initial_input_ast: 5622eb396c2aea656e3bfa6b1ad0d39fce6bc221978a13c9be4d750da46cfc48
initial_ast: 6c7e6756250ed03109d11b90db2cd72e3db761c0ef2d1f8467d8a069df29cfb2
unrolled_ast: 6c7e6756250ed03109d11b90db2cd72e3db761c0ef2d1f8467d8a069df29cfb2
initial_ast: 6cc5bae2d2eb431f10364f7b7fb793810168a938cf19286bc5b86dee628c649e
unrolled_ast: 6cc5bae2d2eb431f10364f7b7fb793810168a938cf19286bc5b86dee628c649e
ssa_ast: 159f247b80c041210de21b1727fcd8134d415d4a60838250561d5d806b3040b5

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 0961f603812e241567b6e3ef5adb458309f1829eb2c08a216efccb17bea89faf
initial_ast: 2b70c901581ffad2061e1a3ad8cef937b7fc3d251a7967704c28ce0613aa0217
unrolled_ast: 2b70c901581ffad2061e1a3ad8cef937b7fc3d251a7967704c28ce0613aa0217
ssa_ast: f21ddb2c05e4050c8f6fbe815263b5d7a160a558e0c887d0c1c8fda808e34ee9

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: f18a0e019ca4719c4c4ef5b7313f562c3bc9581819d161d84566e706f3765249
initial_ast: 64bbdbb801591bfdd54cf70019107a359ca2ef0d326be54c4d5622c848cb8de0
unrolled_ast: 64bbdbb801591bfdd54cf70019107a359ca2ef0d326be54c4d5622c848cb8de0
ssa_ast: 19b73b23e2d41012322994dcd56c19e6780b731023b5ee5e8e13e299b060294e

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 16910a94cf1f803ae6425ae6bee9422b01651c2c243b5e46807dc3191d169e64
initial_ast: 46c9559d438edf0a3ea4939dc33277a04ea67b88424d5aeb56de436635fd83ca
unrolled_ast: 46c9559d438edf0a3ea4939dc33277a04ea67b88424d5aeb56de436635fd83ca
ssa_ast: 1faf75a41e53b620176467eb9883e5dd8339e7bb791b7ed1b0c8b08835a420dc

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 8e1298749b5e6da6a3dcd789ad5a3a102952a00d88c38d4a693616679ef6d8d6
unrolled_ast: 8e1298749b5e6da6a3dcd789ad5a3a102952a00d88c38d4a693616679ef6d8d6
ssa_ast: 1d46e10ce7b7646f4e298961edf7da296fabeb69ae984923ec65dee62e290941

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: e72f4854cb2c9ea3063fee92b826df43185c6a7a0a33f4fe6d24adc0c3dd8367
unrolled_ast: e72f4854cb2c9ea3063fee92b826df43185c6a7a0a33f4fe6d24adc0c3dd8367
ssa_ast: 8a7c826142f437634101d01a4ce859ab4b4d8959c962f349b11e8b8d7a267bfb

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 76aa1559acd72ebe4cf473f29efb52b9196776dd8bf19e1eb936f86eae7c26ed
unrolled_ast: 76aa1559acd72ebe4cf473f29efb52b9196776dd8bf19e1eb936f86eae7c26ed
ssa_ast: b243d8c46f75b6fa82d84ebd9e9fb04c97b90de92a36ec9d332b1cc28da5e89a

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: b649852fa2fd7eda05bd0ba261f01dcee93b6b825d5d30fddb8dd5c5710081ca
initial_ast: 6b8b967664dbbb3e22420516aeac72957bf62b07a87851cf7ba803aac76fb752
unrolled_ast: 6b8b967664dbbb3e22420516aeac72957bf62b07a87851cf7ba803aac76fb752
ssa_ast: 62cc5bcd553392faf1e715afa9a9035afe0b40b4eb446eccfec2a27aca70353e

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 3f35e74d282a1e5281e7f283d1e572a3dc75dea1a5ef1a0f8c7f46412ef946a7
initial_ast: 9ac9a6615a1b0a46478b319e69f84de3a9a39123ddf9c8c7009931d7319967fc
unrolled_ast: 9ac9a6615a1b0a46478b319e69f84de3a9a39123ddf9c8c7009931d7319967fc
ssa_ast: 7f67cf55fd8ec65e9fcfe2c0e2990a432fc2a8c5408bc30df1f6eaeae95ad162

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 4e3882d83c8044e40258f8414966b09c715b00e08bc3383030cecf2c4a825c60
initial_ast: 78cd624d5a41679923d1802a284924d1f2de370a6ab422f1f9b04096707c8731
unrolled_ast: 78cd624d5a41679923d1802a284924d1f2de370a6ab422f1f9b04096707c8731
ssa_ast: f7e3b97fb2bc4ce149d15305e0038985c234c95dfbe360fe3b3ec0bcb3c8044f

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: eeba130bda3ee24f2a4bf92f67fb555ab849173910a647096e28729c2ebd71c2
initial_ast: 460afe0b1aaa06f664fcee198f09df7295b79e5abe384748ac533a1196b38e90
unrolled_ast: 460afe0b1aaa06f664fcee198f09df7295b79e5abe384748ac533a1196b38e90
ssa_ast: 8713711c24c866440d30162270e95a39ad6d0d72c4c959c9758d81111f4c48f9

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 3a510480221eb323713b4b10cc374ba357f130e8ac2b07bf1c69ad5d8c936f12
initial_ast: 103c76ecc5358ae585e4cea32287890acfafc3fcbe53e5c6326b5750eafedbe7
unrolled_ast: 103c76ecc5358ae585e4cea32287890acfafc3fcbe53e5c6326b5750eafedbe7
ssa_ast: 4319c54e2fed6103aff9824776b5a387c8741632f3f1e26886353c04b5fb8ee2

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 3f35e74d282a1e5281e7f283d1e572a3dc75dea1a5ef1a0f8c7f46412ef946a7
initial_ast: 214befce37c1b0c1be25f614a3d928bae22f6d1dc518288b7ed4d007eddec2a7
unrolled_ast: 214befce37c1b0c1be25f614a3d928bae22f6d1dc518288b7ed4d007eddec2a7
ssa_ast: a34c52d398c767b4879c1e0f692b780703af6dd4d13280509404dcc80da5f6d3

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 9206742d7f18345efbd4d9077cd1aca0855d43a2436be0697ec22954650e3737
initial_ast: 30a95deb7c12861dc98cbed1fb9ba37af688ad23504ac80bfedb74d00bcb1d32
unrolled_ast: 30a95deb7c12861dc98cbed1fb9ba37af688ad23504ac80bfedb74d00bcb1d32
ssa_ast: d008e9d11be8d38056c03b0a4ce9ebe26b7af71c438a300380153fc400cc874c

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 047866515f4dc74cd9966242734984b53e72f87afc21f7171b118e6defa1f166
initial_ast: 17aa90c54b70219cdff169e7ed82be28aea9c01dc2067908326c1fa384b83a0a
unrolled_ast: 17aa90c54b70219cdff169e7ed82be28aea9c01dc2067908326c1fa384b83a0a
ssa_ast: 2249e2ea876fdfe5eb154d93d65ebae837f21ed1b5ff6bd4f2bf6e58df4ad0cd

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 5e0a61d909d2e94dfbc95775e4c5c356adb61375ceef2d583a5ab927b3b6342e
initial_ast: f5bdcda30d93281577f0a33daf097e79dd014e0e4bd4cbddfa48ff5f03d2bfd6
unrolled_ast: f5bdcda30d93281577f0a33daf097e79dd014e0e4bd4cbddfa48ff5f03d2bfd6
ssa_ast: 95d8ac9a0a6031987c83dfb86e233ca9ef50cbf4c9a5973baafee8cc9845d3d4

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 4e3882d83c8044e40258f8414966b09c715b00e08bc3383030cecf2c4a825c60
initial_ast: c725f6542c21f1324498289b9496365219fc1dd662af5767d36455751ceb7f95
unrolled_ast: c725f6542c21f1324498289b9496365219fc1dd662af5767d36455751ceb7f95
ssa_ast: c4a54b742dba9420a390ec9385d38e1886e95002caa1a332de3552e7dac7f555

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: e19dcac0064fed4ec8293b9b40ec70cb94b5fdb05f1081fc29f46a023bf79b09
initial_ast: 227822368d014c2dc16897037981786cba91cade824469e49792f12638e8ce6c
unrolled_ast: 227822368d014c2dc16897037981786cba91cade824469e49792f12638e8ce6c
ssa_ast: af6bbdc8053ab7f8d8f5334339b158e6bdbaef57fcaa8a42a95af5400c42e243

View File

@ -4,5 +4,6 @@ expectation: Pass
outputs:
- output:
- initial_input_ast: ae0703890dbea144e675f85228e958d6903df0d1ebd88f16a531624270205cc2
initial_ast: 85d116272d15c32e476cdebedd70fb8d5ce65b5d0046603759760fd6ca8e11fc
unrolled_ast: 85d116272d15c32e476cdebedd70fb8d5ce65b5d0046603759760fd6ca8e11fc
initial_ast: 829a0b024b5b379ba3234ad3d6cdca35209888d3525b5296df72359449081832
unrolled_ast: 829a0b024b5b379ba3234ad3d6cdca35209888d3525b5296df72359449081832
ssa_ast: 102e7525a4abfa9e2b24f324f6e2e8801b408402e84172c619686dfce92eb076

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 20c78fdc1aa61aab7198367aa6618923dca3acf5d2a4abaf1724786fb156234b
initial_ast: 407db22dfe3948bf184ac69c411a5c36013de6c26a57249eb86976d5ff67e711
unrolled_ast: 7d7fa7866e30ad3239889b312fddf3605ad108c982b88f57eb0735df678d56df
ssa_ast: eb633397825e2b57f7c9cab7d06586db2feed0d66d3838d2ed6e1e1d67834f67

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: ab4d5c5684c7eac16c0eef269b050be81f9b5af82e541d9bc43b1da5f2dfce33
initial_ast: 1d5f753676d2482f2b90736dee972528b5e48c8e1455f60cc56fd99f2ba4bef4
unrolled_ast: d8777c1e1c37dac0af3672f7fe4e52278c9544b26812131bfa18634b07fde06e
ssa_ast: 1c4e1db69e76fb066bf3376753f822919662db4cdd813ff4d87f70c7e27f9ff8

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 9af772ba2f200de34252f1d0306b9dd5d028f5019ff940be9c3624c4c78e1518
initial_ast: 1a150cbe9adfca148c28fe63a7c9cb069765df53cd76c09c1ac9a9fb62f106c3
unrolled_ast: 1a150cbe9adfca148c28fe63a7c9cb069765df53cd76c09c1ac9a9fb62f106c3
ssa_ast: 1c7fb3babeb1dd663cb1dcac6fcd7f2018780d5e63d19f8aeb9a2554e2e5519d

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: b36400e27028efb4b1b490603f9f3d4f1ab5b6c0d2556e951977e5cc6637221e
initial_ast: d133b0af76fb581e058760fc979962d122e3bc43cf156684449335c405536d33
unrolled_ast: d133b0af76fb581e058760fc979962d122e3bc43cf156684449335c405536d33
ssa_ast: acd6b4245214ea65af88a83ccc099cd03f500fa3207545c184ff5f92fbe93c6f

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: a28b88dc36ace78ed0de93b56db7274f16521597ccf1e820a17fdb60a1e3d92a
initial_ast: 1fd23b66e7ae5a9ec8168c559fa30f913c853d2a8bd64e8e5a56d39a3c18e33f
unrolled_ast: 1fd23b66e7ae5a9ec8168c559fa30f913c853d2a8bd64e8e5a56d39a3c18e33f
ssa_ast: 89eac5c5de25f1c670dbe77bd9c422550cabf27415f297a91586898ed02b6df4

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 00c3cc87ce3c30894ad6b6874ce6dacfa02c9f2bc171615ff627f06f2e201997
initial_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
unrolled_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
ssa_ast: f2ae99852d42ea765c8cbe2292155fc5ca577b451f0a59dc055b3a3e0d86dd64

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 6a4eb74c21f558af5535a98ec6ce7ef851896476f7dfb6ae8385b9a0cb865973
unrolled_ast: 6a4eb74c21f558af5535a98ec6ce7ef851896476f7dfb6ae8385b9a0cb865973
ssa_ast: 13325fec05940394c3513afdd63ad4cc7d7087ea80908b63a9c765c17472b242

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 88c6a9385764de1df4430489aec43d7eb2563886fa65239ca068793d611a737d
unrolled_ast: 88c6a9385764de1df4430489aec43d7eb2563886fa65239ca068793d611a737d
ssa_ast: 5f59e8f16ed79a1950744c549c53fff88855c46f49b7c6a2fefce8b58f9a4f00

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 8e7932b4bd7ef496bdec3763e81a8e09908c212db21c89e1e0c966e82574fac9
unrolled_ast: 8e7932b4bd7ef496bdec3763e81a8e09908c212db21c89e1e0c966e82574fac9
ssa_ast: 40f7365611890658ccf286a8035a4af5ece8eb5161b39757d070c973ae1649d5

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 00c3cc87ce3c30894ad6b6874ce6dacfa02c9f2bc171615ff627f06f2e201997
initial_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
unrolled_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
ssa_ast: f2ae99852d42ea765c8cbe2292155fc5ca577b451f0a59dc055b3a3e0d86dd64

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: c93f9fd667509aa0aa3896c261cb48c7d579d9856d0a14b96e9b2c7e04566a0a
initial_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
unrolled_ast: 05854d5dee2e23cb551eb131f9d3bec4d7f33fbbed3c29dbecf055f7d8995d9e
ssa_ast: f2ae99852d42ea765c8cbe2292155fc5ca577b451f0a59dc055b3a3e0d86dd64

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 7b0236b04ad9caa4039a989b91e7f49021a9daf09a495a9cdad7c371ee196761
initial_ast: 02a6c3318bd4ccf328425cd890e16d2933b67e5b9b14ac6244c743275f03a1a6
unrolled_ast: 02a6c3318bd4ccf328425cd890e16d2933b67e5b9b14ac6244c743275f03a1a6
ssa_ast: cc5334998cd5dc90e14e882a9704e16627698eb31d9fc3f10a5d3df9a5e9006c

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 5e1e23855cb6841ee210c8a24e11cc819e91ce3b087a8c961035c574baa1784b
initial_ast: d362fcd34317c6af293d435e907f137d6aa71f95683629e0cb9c68f92393a677
unrolled_ast: d362fcd34317c6af293d435e907f137d6aa71f95683629e0cb9c68f92393a677
ssa_ast: efdda1a1d2b3842c5d29b923fc1c12520df322cf95dca68b9fa5481885c91412

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: c834a4d1f5aecdef55a47337945ef657e297f539f5db59f1d40eef343306e7fd
unrolled_ast: c834a4d1f5aecdef55a47337945ef657e297f539f5db59f1d40eef343306e7fd
ssa_ast: 374a2c1089cebb08c8c30f01863b169d10e1fb72113fbffaa68f69b2b31b8bad

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: a60503e3f83fbee658d02fb3806b3a3326fc6d4f4e43ac05bce7b16ac0552edb
initial_ast: 7ffab494280736953e34589589b16b26cead38323b223ae95ef7199fa7a48e04
unrolled_ast: 7ffab494280736953e34589589b16b26cead38323b223ae95ef7199fa7a48e04
ssa_ast: 1438c89357d3ec3ba2b5def56660c60b5385610798048e77202ca4134127490f

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: c77551885b3f06f2ec5db73868feb8fd7521c4c2c57bbf95c367e56cd93ab991
unrolled_ast: c77551885b3f06f2ec5db73868feb8fd7521c4c2c57bbf95c367e56cd93ab991
ssa_ast: bdf64a26de6aeecffcc6052bfc8c515cac27c6d5f6cddf63e82072dfdf2d3106

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 1b5330a3356c437ddc09afc027d1365eedb24c56777772fd83b9167cfebb4435
initial_ast: 89a9088b787d344c1d55e7fba58c94773a58dde967d2499e897ce19ba5587653
unrolled_ast: 89a9088b787d344c1d55e7fba58c94773a58dde967d2499e897ce19ba5587653
ssa_ast: 71b7378f39431df859e595efc138ed174818cd59877fd06a682f6c6cf4bb048e

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 72dd086a27dd223b377616799a8b5f47a71e017ca28747ae671e17caf14f8e69
unrolled_ast: 72dd086a27dd223b377616799a8b5f47a71e017ca28747ae671e17caf14f8e69
ssa_ast: b27f28826cd6866f4753445859903b10a67a3ccee4c1ada10ba4395d3c2ac1ec

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: a28b88dc36ace78ed0de93b56db7274f16521597ccf1e820a17fdb60a1e3d92a
initial_ast: 9e499c18873510b9af07db030a4d5a14064e9353fd68ab0462c3ebd9a24fe066
unrolled_ast: 9e499c18873510b9af07db030a4d5a14064e9353fd68ab0462c3ebd9a24fe066
ssa_ast: 4b4138b73f53b31ac6de117f3d6f92e347e410b7479ab8db5826e50ac7e6b136

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: d12e492b73a208051457ad2ce9ed8dbbb5a8096f32f52d697c41972ba8b88d35
initial_ast: fd9ea0e771a68f031e301d5b106341c3292dfa37572b83b917fb75676c55ab98
unrolled_ast: fd9ea0e771a68f031e301d5b106341c3292dfa37572b83b917fb75676c55ab98
ssa_ast: be281e2ccbeba28b40d08efbf515a8eccea2b0c4cd80325b9b52ed795c73ec11

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 8bbb4a56d86d6180a6d335bfcc622ebba88f413907c2eb39e93ef5b6597609d4
unrolled_ast: 8bbb4a56d86d6180a6d335bfcc622ebba88f413907c2eb39e93ef5b6597609d4
ssa_ast: 608c790d7e7e4057a44b5c7d5e754c4145de342685fef381ac085b99d1828b43

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: d81b26d3c0174fb20bd498c956af3c4f20efca37c817bd06e8ce6bd02f7ad5c3
unrolled_ast: d81b26d3c0174fb20bd498c956af3c4f20efca37c817bd06e8ce6bd02f7ad5c3
ssa_ast: 3f5dd352321065c15eee0b57d2849107211ea687353690e9ac6cd5e220a2faa4

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 68f706452fc9304afe6345986e427b65d6df41157750241d2693a8c0edd24add
unrolled_ast: 68f706452fc9304afe6345986e427b65d6df41157750241d2693a8c0edd24add
ssa_ast: dc7f78f28a62935e1011f7974c99174ead1e0e2cfda8e337a71e5a1ee2c44795

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 43d64ea2b989308c0fb65af569584f70dde0ffba12966885c7dc3025f9fbc5d2
unrolled_ast: 43d64ea2b989308c0fb65af569584f70dde0ffba12966885c7dc3025f9fbc5d2
ssa_ast: 9cc98d653b965ea3613f855e0039f61d7862b73b4f093d240d92b74bae5d64dc

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 7fc9bd0e110f629f957d5deeca1a1586e197c883aac5211c741944c1229943f6
unrolled_ast: 7fc9bd0e110f629f957d5deeca1a1586e197c883aac5211c741944c1229943f6
ssa_ast: a893fb34a04397aa236acc13ae96ac17601bd5ccc82488606d1162af0cbdfd61

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 042a092947df5736285b242eac235d558237d3ead82a163a11db9c078f6cefaa
unrolled_ast: 042a092947df5736285b242eac235d558237d3ead82a163a11db9c078f6cefaa
ssa_ast: 34317aaa85d069972bc549612ccddad791aa3cdbc9f401c5550d6cdf5d276afc

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: e9f0202a1f7177791abab18c4b675b377c8b6c937447e264921ded60f9f3bdd8
unrolled_ast: e9f0202a1f7177791abab18c4b675b377c8b6c937447e264921ded60f9f3bdd8
ssa_ast: fbeb47642f8100044e8b1505fa774327db108d59605204b620aff4f43a063a35

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: no input
initial_ast: 1563a60bff892a9e6e94548ae3f805baaa8eab13b8e3fb45d9609c3d8cd32f6b
unrolled_ast: 1563a60bff892a9e6e94548ae3f805baaa8eab13b8e3fb45d9609c3d8cd32f6b
ssa_ast: d09d7ccf2a8d42276dabed140708d2a13ba129e61813db5bc81098c0e1ede294

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 23e62412d2a9377334d90aaeb6629b73c77e045ce87f23bd6ae2e2cd242e70f0
initial_ast: 2b0cfb24941fb39064b691681ecccae455fcc86feb65ea08bcce1e4ba994839f
unrolled_ast: 2b0cfb24941fb39064b691681ecccae455fcc86feb65ea08bcce1e4ba994839f
ssa_ast: e8a850d37dfb2810edcebd0ff2d077210cff9f8fa517884d790845b80a372a63

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 2b6bc4ade2305a65746066befacf6a0a18382f754d4d7911d0c6e0abef682114
initial_ast: f7431d8437cd9f6da8f21260e823610ccb88cb7588e429d5a9b56546af48538b
unrolled_ast: f7431d8437cd9f6da8f21260e823610ccb88cb7588e429d5a9b56546af48538b
ssa_ast: f4d96c3e1dffbc527c29919914ff40327342d534c8fa90c0b7b50f7b626af705

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 4001f721e97052bdea8fafe39846356011e335e40281e0082c0b406cd6a85947
initial_ast: 29be2867fe0418b46a1db5936f21636cf5b694a48112f3e7f85628f972f631e5
unrolled_ast: 29be2867fe0418b46a1db5936f21636cf5b694a48112f3e7f85628f972f631e5
ssa_ast: 82fee3950cbf846fa44125856ac4f279129a48c377fbb373ac234883918824d8

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: e626f055978f5125bc292065d74aab5b679229a5364f150ccbe1f07d0c167c3d
initial_ast: 2ce20f8feffa0ac800dae56918308b9715a742f9c135dc1368bfacbb1dd4509c
unrolled_ast: 2ce20f8feffa0ac800dae56918308b9715a742f9c135dc1368bfacbb1dd4509c
ssa_ast: 4430cc07bb5c3217ce66d2430278d910d0efc38949dd2c2e61dd4b3c93ca594e

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 3eaa98274698edacf455de40418ea012234a1355d5b50b9063ee0e06d3d26709
initial_ast: be527618fdb6232d3b482186db6323321682604f3b9ab8870700ef24be8fe60f
unrolled_ast: be527618fdb6232d3b482186db6323321682604f3b9ab8870700ef24be8fe60f
ssa_ast: d476d945d54caf24d511a241f3fbeb2a2958fa750ec069beffe99261eee6d594

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 4efe3ae5f2d6a95663ca302b60d4e430b63bcb7c5a19d638ec7613b7dc099825
initial_ast: 29ae1d1c58b9b58e41b74a1fa52dadd4d373652d0af51e614ef88c356b8c0143
unrolled_ast: 29ae1d1c58b9b58e41b74a1fa52dadd4d373652d0af51e614ef88c356b8c0143
ssa_ast: 6d24c34d9c47445dc2a893d840a56033d109388233680cc4f580622ed289a4f3

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: e7173f6b8aa8aa40bcb167fa4de0b5d5a7f1b6d245a78dcb5ad70a73b53ef7de
initial_ast: 9f4dab18c4804c4b7229c563f6bd7cb2d5cb5d53df2dbfdb70b3f1f05420546b
unrolled_ast: 9f4dab18c4804c4b7229c563f6bd7cb2d5cb5d53df2dbfdb70b3f1f05420546b
ssa_ast: fc613d21d8034b1a9a9ced7ef45a965ffe0a743430704a6fc6ecca9d550ed336

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: b7a1796fa4abcee565ee7dea475310c15b5881c60a2acc7a2b65cea9a84acf56
initial_ast: 111c4da5ce13398915927b769845f270d4d76001cde3fca5ef2f6827a3f55cc6
unrolled_ast: 111c4da5ce13398915927b769845f270d4d76001cde3fca5ef2f6827a3f55cc6
ssa_ast: bb394ca1e51666ed8fd63d3820fc933a4a6467bff71843186bcf9132ca368dfa

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: b7a1796fa4abcee565ee7dea475310c15b5881c60a2acc7a2b65cea9a84acf56
initial_ast: d15a50234ff5afb09f488cdac1e83588c1449bc2f1b8c042692ba9e073462ecc
unrolled_ast: d15a50234ff5afb09f488cdac1e83588c1449bc2f1b8c042692ba9e073462ecc
ssa_ast: 62c72973afc16bdc53a92f074e29d3dcf3d882d39319aba118899b4b506a4ebb

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 809c4e4298aa9ee1320cb7b491bc3dc81deb71a691cdc7add970e2c2bf5f47b5
initial_ast: e3402b36581c371ef4eb1faad35ffd2fcafae3cb1e1e4b3b0802b5e8d7a241c4
unrolled_ast: e3402b36581c371ef4eb1faad35ffd2fcafae3cb1e1e4b3b0802b5e8d7a241c4
ssa_ast: 841632483e0c454cdb0f64c9ac50db6359ddc4c8fbd8307904d87cd4ba605eaf

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 7034fae8c2db1f78f9f42400f5a6b28d498a7d31f7e35923584622420bfa0ef6
initial_ast: 1e81162d25e94d5b228c5a038e5fe1e0de683e287a43a37d462ee269ebdfc085
unrolled_ast: 1e81162d25e94d5b228c5a038e5fe1e0de683e287a43a37d462ee269ebdfc085
ssa_ast: 0e75ba788c3699798619fce4340a9838306c3d2acdc867c296f8cca75bcbb571

View File

@ -6,3 +6,4 @@ outputs:
- initial_input_ast: 91219f5a1516834f9c60220a65cece763ae40c916f636fed729b1fd91e25310a
initial_ast: ca133cfbaa8f8d3878fdb49176b664e02838ac1f2772f2882a8246cd4912defa
unrolled_ast: ca133cfbaa8f8d3878fdb49176b664e02838ac1f2772f2882a8246cd4912defa
ssa_ast: 4476933f34fc773f90a8d6119534583bf4b078e10ccd9891aad1cddc07906798

View File

@ -7,3 +7,4 @@ outputs:
- initial_input_ast: 9036921d0594f2bc8402c7364492ca47d57d34e8588b0bef6491ae6978454e31
initial_ast: 25767879a281db7754873fcee20c4ab34d0749a34cbc462802712b2de644eb80
unrolled_ast: 25767879a281db7754873fcee20c4ab34d0749a34cbc462802712b2de644eb80
ssa_ast: 10467f41b268e304d30e1d891b44af7c6f2733c14347e349b1eb86f5c54af38d

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