isla_lib/ir/
ssa.rs

1// BSD 2-Clause License
2//
3// Copyright (c) 2020 Alasdair Armstrong
4//
5// All rights reserved.
6//
7// Redistribution and use in source and binary forms, with or without
8// modification, are permitted provided that the following conditions are
9// met:
10//
11// 1. Redistributions of source code must retain the above copyright
12// notice, this list of conditions and the following disclaimer.
13//
14// 2. Redistributions in binary form must reproduce the above copyright
15// notice, this list of conditions and the following disclaimer in the
16// documentation and/or other materials provided with the distribution.
17//
18// THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
19// "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
20// LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
21// A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
22// HOLDER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
23// SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
24// LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
25// DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
26// THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
27// (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
28// OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
29
30//! This module provides a control flow graph type [CFG] for IR
31//! analysis and transforms, and additionally supports conversion of
32//! that CFG into single static assignment (SSA) form via
33//! [CFG::ssa()].
34
35use petgraph::algo::dominators::{self, Dominators};
36use petgraph::graph::{Graph, NodeIndex};
37use petgraph::{Directed, Direction};
38use std::collections::HashMap;
39use std::fmt;
40use std::io::Write;
41use std::usize;
42
43use super::*;
44use crate::primop::{Binary, Unary, Variadic};
45
46/// A [SSAName] is a [Name] augmented with an additional number. The
47/// number is a signed integer, with the value `-1` representing a
48/// name that does not have an SSA number (either because the CFG is
49/// not in SSA form, or because it is a type name, function identifer,
50/// or similar that does not need one in SSA form).
51#[derive(Copy, Clone, Debug, PartialEq, Eq, Hash, Serialize, Deserialize)]
52pub struct SSAName {
53    name: Name,
54    number: i32,
55}
56
57impl SSAName {
58    pub fn new(name: Name) -> Self {
59        SSAName { name, number: -1 }
60    }
61
62    pub fn new_ssa(name: Name, number: i32) -> Self {
63        SSAName { name, number }
64    }
65
66    pub fn base_name(self) -> Name {
67        self.name
68    }
69
70    fn ssa_number_mut(&mut self) -> &mut i32 {
71        &mut self.number
72    }
73
74    pub fn ssa_number(self) -> i32 {
75        self.number
76    }
77
78    pub(crate) fn unssa(self, symtab: &mut Symtab, generated: &mut HashMap<SSAName, Name>) -> Name {
79        if self.number < 0 {
80            self.name
81        } else if let Some(name) = generated.get(&self) {
82            *name
83        } else {
84            let gs = symtab.gensym();
85            generated.insert(self, gs);
86            gs
87        }
88    }
89
90    fn write(self, output: &mut dyn Write, symtab: &Symtab) -> std::io::Result<()> {
91        if self.number >= 0 {
92            write!(output, "{}/{}", zencode::decode(symtab.to_str(self.name)), self.number)
93        } else {
94            write!(output, "{}", zencode::decode(symtab.to_str(self.name)))
95        }
96    }
97}
98
99pub(crate) fn unssa_ty(ty: &Ty<SSAName>) -> Ty<Name> {
100    use Ty::*;
101    match ty {
102        I64 => I64,
103        I128 => I128,
104        AnyBits => AnyBits,
105        Bits(n) => Bits(*n),
106        Unit => Unit,
107        Bool => Bool,
108        Bit => Bit,
109        String => String,
110        Real => Real,
111        Enum(id) => {
112            assert!(id.number < 0);
113            Enum(id.name)
114        }
115        Struct(id) => {
116            assert!(id.number < 0);
117            Struct(id.name)
118        }
119        Union(id) => {
120            assert!(id.number < 0);
121            Union(id.name)
122        }
123        Vector(ty) => Vector(Box::new(unssa_ty(ty))),
124        FixedVector(n, ty) => FixedVector(*n, Box::new(unssa_ty(ty))),
125        List(ty) => List(Box::new(unssa_ty(ty))),
126        Ref(ty) => Ref(Box::new(unssa_ty(ty))),
127    }
128}
129
130#[derive(Clone, Debug)]
131pub enum BlockLoc {
132    Id(SSAName),
133    // Field locations contain the previous name so that we can update one field at a time
134    Field(Box<BlockLoc>, SSAName, SSAName),
135    Addr(Box<BlockLoc>),
136}
137
138impl BlockLoc {
139    fn id(&self) -> SSAName {
140        match self {
141            BlockLoc::Id(id) => *id,
142            BlockLoc::Field(loc, _, _) | BlockLoc::Addr(loc) => loc.id(),
143        }
144    }
145
146    fn ids(&self) -> (SSAName, Option<SSAName>) {
147        match self {
148            BlockLoc::Id(id) => (*id, None),
149            BlockLoc::Field(loc, base_id, _) => (loc.id(), Some(*base_id)),
150            BlockLoc::Addr(loc) => (loc.id(), None),
151        }
152    }
153
154    fn collect_variables<'a, 'b>(&'a mut self, vars: &'b mut Vec<Variable<'a, SSAName>>) {
155        match self {
156            BlockLoc::Id(id) => vars.push(Variable::Declaration(id)),
157            BlockLoc::Field(loc, id, _) => {
158                loc.collect_variables(vars);
159                vars.push(Variable::Usage(id))
160            }
161            BlockLoc::Addr(loc) => loc.collect_variables(vars),
162        }
163    }
164}
165
166impl From<&Loc<Name>> for BlockLoc {
167    fn from(loc: &Loc<Name>) -> Self {
168        match loc {
169            Loc::Id(id) => BlockLoc::Id(SSAName::new(*id)),
170            Loc::Field(loc, field) => {
171                let base_name = SSAName::new(loc.id());
172                BlockLoc::Field(Box::new(Self::from(loc.as_ref())), base_name, SSAName::new(*field))
173            }
174            Loc::Addr(loc) => BlockLoc::Addr(Box::new(Self::from(loc.as_ref()))),
175        }
176    }
177}
178
179/// [BlockInstr] is the same as [Instr] but restricted to just
180/// instructions that can appear in basic blocks, and with all names
181/// replaced by [SSAName].
182pub enum BlockInstr<B> {
183    Decl(SSAName, Ty<SSAName>),
184    Init(SSAName, Ty<SSAName>, Exp<SSAName>),
185    Copy(BlockLoc, Exp<SSAName>),
186    Monomorphize(SSAName),
187    Call(BlockLoc, bool, Name, Vec<Exp<SSAName>>),
188    PrimopUnary(BlockLoc, Unary<B>, Exp<SSAName>),
189    PrimopBinary(BlockLoc, Binary<B>, Exp<SSAName>, Exp<SSAName>),
190    PrimopVariadic(BlockLoc, Variadic<B>, Vec<Exp<SSAName>>),
191}
192
193impl<B: BV> BlockInstr<B> {
194    // Returns the written-to variable, and its previous name if it's a field access
195    pub fn write_ssa(&self) -> Option<(SSAName, Option<SSAName>)> {
196        use BlockInstr::*;
197        match self {
198            Decl(id, _) | Init(id, _, _) => Some((*id, None)),
199            Copy(loc, _)
200            | Call(loc, _, _, _)
201            | PrimopUnary(loc, _, _)
202            | PrimopBinary(loc, _, _, _)
203            | PrimopVariadic(loc, _, _) => Some(loc.ids()),
204            _ => None,
205        }
206    }
207
208    pub fn write(&self) -> Option<Name> {
209        self.write_ssa().map(|(id, _)| id.name)
210    }
211
212    pub fn declares(&self) -> Option<Name> {
213        use BlockInstr::*;
214        match self {
215            Decl(id, _) | Init(id, _, _) => Some(id.name),
216            _ => None,
217        }
218    }
219
220    fn declares_typed(&self) -> Option<(Name, Ty<Name>)> {
221        use BlockInstr::*;
222        match self {
223            Decl(id, ty) | Init(id, ty, _) => Some((id.name, unssa_ty(ty))),
224            _ => None,
225        }
226    }
227
228    fn collect_variables<'a, 'b>(&'a mut self, vars: &'b mut Vec<Variable<'a, SSAName>>) {
229        use BlockInstr::*;
230        match self {
231            Decl(id, _) => vars.push(Variable::Declaration(id)),
232            Init(id, _, exp) => {
233                vars.push(Variable::Declaration(id));
234                exp.collect_variables(vars)
235            }
236            Copy(loc, exp) => {
237                loc.collect_variables(vars);
238                exp.collect_variables(vars)
239            }
240            Monomorphize(id) => vars.push(Variable::Usage(id)),
241            Call(loc, _, _, args) => {
242                loc.collect_variables(vars);
243                args.iter_mut().for_each(|exp| exp.collect_variables(vars))
244            }
245            PrimopUnary(loc, _, exp) => {
246                loc.collect_variables(vars);
247                exp.collect_variables(vars)
248            }
249            PrimopBinary(loc, _, lhs, rhs) => {
250                loc.collect_variables(vars);
251                lhs.collect_variables(vars);
252                rhs.collect_variables(vars)
253            }
254            PrimopVariadic(loc, _, args) => {
255                loc.collect_variables(vars);
256                args.iter_mut().for_each(|exp| exp.collect_variables(vars))
257            }
258        }
259    }
260
261    fn variables(&mut self) -> Variables<'_, SSAName> {
262        let mut vec = Vec::new();
263        self.collect_variables(&mut vec);
264        Variables::from_vec(vec)
265    }
266}
267
268impl<B: fmt::Debug> fmt::Debug for BlockInstr<B> {
269    fn fmt(&self, f: &mut fmt::Formatter<'_>) -> fmt::Result {
270        use BlockInstr::*;
271        match self {
272            Decl(id, ty) => write!(f, "{:?} : {:?}", id, ty),
273            Init(id, ty, exp) => write!(f, "{:?} : {:?} = {:?}", id, ty, exp),
274            Copy(loc, exp) => write!(f, "{:?} = {:?}", loc, exp),
275            Monomorphize(id) => write!(f, "mono {:?}", id),
276            Call(loc, ext, id, args) => write!(f, "{:?} = {:?}<{:?}>({:?})", loc, id, ext, args),
277            _ => write!(f, "primop"),
278        }
279    }
280}
281
282/// An instruction that can only occur at the end of a basic block.
283#[derive(Debug)]
284pub enum Terminator {
285    Continue,
286    Goto(usize),
287    Jump(Exp<SSAName>, usize, String),
288    Failure,
289    Arbitrary,
290    End,
291}
292
293impl Terminator {
294    fn collect_variables<'a, 'b>(&'a mut self, vars: &'b mut Vec<Variable<'a, SSAName>>) {
295        if let Terminator::Jump(exp, _, _) = self {
296            exp.collect_variables(vars)
297        }
298    }
299
300    fn variables(&mut self) -> Variables<'_, SSAName> {
301        let mut vec = Vec::new();
302        self.collect_variables(&mut vec);
303        Variables::from_vec(vec)
304    }
305}
306
307#[derive(Debug)]
308pub struct Block<B> {
309    pub label: Option<usize>,
310    pub phis: Vec<(SSAName, Vec<SSAName>)>,
311    pub instrs: Vec<BlockInstr<B>>,
312    pub terminator: Terminator,
313}
314
315impl<B: BV> Block<B> {
316    fn insert_phi(&mut self, id: Name, num_preds: usize) {
317        self.phis.push((SSAName::new(id), vec![SSAName::new(id); num_preds]))
318    }
319
320    fn rename(&mut self, counts: &mut HashMap<Name, i32>, stacks: &mut HashMap<Name, Vec<i32>>) {
321        for (id, _) in self.phis.iter_mut() {
322            let i = counts.entry(id.name).or_default();
323            *i += 1;
324            stacks.entry(id.name).or_default().push(*i);
325            *id.ssa_number_mut() = *i;
326        }
327
328        for instr in self.instrs.iter_mut() {
329            for variable_use in instr.variables() {
330                match variable_use {
331                    Variable::Declaration(id) => {
332                        let i = counts.entry(id.name).or_default();
333                        *i += 1;
334                        stacks.entry(id.name).or_default().push(*i);
335                        *id.ssa_number_mut() = *i
336                    }
337                    Variable::Usage(id) => {
338                        if let Some(stack) = stacks.get(&id.name) {
339                            let i = stack.last().expect("Empty stack when renaming variables");
340                            *id.ssa_number_mut() = *i
341                        }
342                    }
343                }
344            }
345        }
346
347        for variable_use in self.terminator.variables() {
348            if let Variable::Usage(id) = variable_use {
349                if let Some(stack) = stacks.get(&id.name) {
350                    let i = stack.last().expect("Empty stack when renaming variables");
351                    *id.ssa_number_mut() = *i
352                }
353            }
354        }
355    }
356
357    fn rename_phi_arg(&mut self, j: usize, stacks: &mut HashMap<Name, Vec<i32>>) {
358        for (id, args) in self.phis.iter_mut() {
359            let i = stacks.get(&id.name).and_then(|v| v.last()).expect("Empty stack when renaming phi arg");
360            if *i != 0 {
361                if !args.is_empty() {
362                    *args[j].ssa_number_mut() = *i
363                }
364            } else {
365                // A phi function that has variable x/0 is pointing to
366                // an undeclared variable x, which implies that x has
367                // gone out of scope, and so this phi function can be
368                // pruned.
369                *args = vec![]
370            }
371        }
372    }
373}
374
375/// An edge between basic blocks in the control flow graph. The edge
376/// corresponds to the [Terminator] of the node the edge comes from.
377#[derive(Debug)]
378pub enum Edge {
379    /// True if the Jump expression was true, and the jump was taken.
380    Jump(bool),
381    Goto,
382    Continue,
383}
384
385pub struct CFG<B> {
386    pub root: NodeIndex,
387    pub graph: Graph<Block<B>, Edge, Directed>,
388}
389
390/// Takes the final instruction in a basic block, and returns the
391/// appropriate terminator for that block.
392fn to_terminator<B: BV>(instr: &Instr<Name, B>) -> Terminator {
393    match instr {
394        Instr::Goto(target) => Terminator::Goto(*target),
395        Instr::Jump(cond, target, loc) => Terminator::Jump(block_exp(cond), *target, loc.clone()),
396        Instr::Failure => Terminator::Failure,
397        Instr::Arbitrary => Terminator::Arbitrary,
398        Instr::End => Terminator::End,
399        _ => Terminator::Continue,
400    }
401}
402
403type TerminatorSplit<'a, B> = (&'a [LabeledInstr<B>], Option<usize>, Terminator, &'a [LabeledInstr<B>]);
404
405fn split_terminator<B: BV>(instrs: &[LabeledInstr<B>]) -> Option<TerminatorSplit<'_, B>> {
406    for (i, instr) in instrs.iter().enumerate() {
407        match instr.strip_ref() {
408            // Any labeled instruction after the first becomes the start of a new block
409            _ if i > 0 && instr.is_labeled() => return Some((&instrs[0..i], None, Terminator::Continue, &instrs[i..])),
410            Instr::Goto(_) | Instr::Jump(_, _, _) | Instr::Failure | Instr::Arbitrary | Instr::End => {
411                return Some((&instrs[0..i], instr.label(), to_terminator(instr.strip_ref()), &instrs[(i + 1)..]))
412            }
413            _ => (),
414        }
415    }
416    None
417}
418
419fn block_label<B: BV>(instrs: &[LabeledInstr<B>], terminator_label: Option<usize>) -> Option<usize> {
420    match instrs.first() {
421        None => terminator_label,
422        Some(instr) => instr.label(),
423    }
424}
425
426fn block_ty(ty: &Ty<Name>) -> Ty<SSAName> {
427    use Ty::*;
428    match ty {
429        I64 => I64,
430        I128 => I128,
431        AnyBits => AnyBits,
432        Bits(n) => Bits(*n),
433        Unit => Unit,
434        Bool => Bool,
435        Bit => Bit,
436        String => String,
437        Real => Real,
438        Enum(id) => Enum(SSAName::new(*id)),
439        Struct(id) => Struct(SSAName::new(*id)),
440        Union(id) => Union(SSAName::new(*id)),
441        Vector(ty) => Vector(Box::new(block_ty(ty))),
442        FixedVector(n, ty) => FixedVector(*n, Box::new(block_ty(ty))),
443        List(ty) => List(Box::new(block_ty(ty))),
444        Ref(ty) => Ref(Box::new(block_ty(ty))),
445    }
446}
447
448fn block_exp(exp: &Exp<Name>) -> Exp<SSAName> {
449    use Exp::*;
450    match exp {
451        Id(id) => Id(SSAName::new(*id)),
452        Ref(reg) => Ref(SSAName::new(*reg)),
453        Bool(b) => Bool(*b),
454        Bits(bv) => Bits(*bv),
455        String(s) => String(s.clone()),
456        Unit => Unit,
457        I64(n) => I64(*n),
458        I128(n) => I128(*n),
459        Undefined(ty) => Undefined(block_ty(ty)),
460        Struct(s, fields) => {
461            Struct(SSAName::new(*s), fields.iter().map(|(field, exp)| (SSAName::new(*field), block_exp(exp))).collect())
462        }
463        Kind(ctor, exp) => Kind(SSAName::new(*ctor), Box::new(block_exp(exp))),
464        Unwrap(ctor, exp) => Unwrap(SSAName::new(*ctor), Box::new(block_exp(exp))),
465        Field(exp, field) => Field(Box::new(block_exp(exp)), SSAName::new(*field)),
466        Call(op, args) => Call(*op, args.iter().map(block_exp).collect()),
467    }
468}
469
470fn block_instrs<B: BV>(instrs: &[LabeledInstr<B>]) -> Vec<BlockInstr<B>> {
471    use BlockInstr::*;
472
473    instrs
474        .iter()
475        .enumerate()
476        .map(|(i, instr)| {
477            // We should never have jump targets into the middle of a
478            // basic block, so all instructions other than the first
479            // should be unlabelled.
480            assert!(i == 0 || instr.is_unlabeled());
481
482            match instr.strip_ref() {
483                Instr::Decl(v, ty) => Decl(SSAName::new(*v), block_ty(ty)),
484                Instr::Init(v, ty, exp) => Init(SSAName::new(*v), block_ty(ty), block_exp(exp)),
485                Instr::Copy(loc, exp) => Copy(BlockLoc::from(loc), block_exp(exp)),
486                Instr::Monomorphize(v) => Monomorphize(SSAName::new(*v)),
487                Instr::Call(loc, ext, f, args) => {
488                    Call(BlockLoc::from(loc), *ext, *f, args.iter().map(block_exp).collect())
489                }
490                Instr::PrimopUnary(loc, fptr, exp) => PrimopUnary(BlockLoc::from(loc), *fptr, block_exp(exp)),
491                Instr::PrimopBinary(loc, fptr, exp1, exp2) => {
492                    PrimopBinary(BlockLoc::from(loc), *fptr, block_exp(exp1), block_exp(exp2))
493                }
494                Instr::PrimopVariadic(loc, fptr, args) => {
495                    PrimopVariadic(BlockLoc::from(loc), *fptr, args.iter().map(block_exp).collect())
496                }
497                // All other cases should be terminators
498                _ => panic!("Invalid block instruction {:?}", instr),
499            }
500        })
501        .collect()
502}
503
504struct DominanceFrontiers {
505    frontiers: Vec<HashSet<NodeIndex>>,
506}
507
508impl DominanceFrontiers {
509    /// Dominance frontier algorithm from 'A Simple, Fast Dominance
510    /// Algorithm' by Cooper et al.
511    fn from_graph<N, E>(graph: &Graph<N, E>, doms: &Dominators<NodeIndex>) -> Self {
512        let mut frontiers = vec![HashSet::new(); graph.node_count()];
513
514        for b in graph.node_indices() {
515            let mut predecessors = graph.neighbors_directed(b, Direction::Incoming);
516            if let Some(first) = predecessors.next() {
517                if let Some(second) = predecessors.next() {
518                    for p in [first, second].iter().copied().chain(predecessors) {
519                        let mut runner = p;
520                        while runner != doms.immediate_dominator(b).unwrap() {
521                            frontiers[runner.index()].insert(b);
522                            runner = doms.immediate_dominator(runner).unwrap()
523                        }
524                    }
525                }
526            }
527        }
528
529        DominanceFrontiers { frontiers }
530    }
531
532    fn get(&self, n: NodeIndex) -> &HashSet<NodeIndex> {
533        &self.frontiers[n.index()]
534    }
535}
536
537struct DominatorTree {
538    tree: HashMap<NodeIndex, Vec<NodeIndex>>,
539}
540
541impl DominatorTree {
542    fn children(&self, n: NodeIndex) -> Option<&[NodeIndex]> {
543        self.tree.get(&n).map(|v| &**v)
544    }
545}
546
547impl<B: BV> CFG<B> {
548    /// Construct a control flow graph from a slice of labeled
549    /// instructions. Note that the set of labels should be pruned
550    /// with [super::prune_labels], otherwise the control flow graph
551    /// will end up containing redundant blocks.
552    pub fn new(instrs: &[LabeledInstr<B>]) -> Self {
553        let mut remaining = instrs;
554        let mut graph = Graph::new();
555        let mut block_indices: Vec<NodeIndex> = Vec::new();
556
557        while let Some((before, terminator_label, terminator, after)) = split_terminator(remaining) {
558            remaining = after;
559            block_indices.push(graph.add_node(Block {
560                label: block_label(before, terminator_label),
561                phis: Vec::new(),
562                instrs: block_instrs(before),
563                terminator,
564            }))
565        }
566
567        let mut targets: HashMap<usize, NodeIndex> = HashMap::new();
568
569        for ix in graph.node_indices() {
570            if let Some(label) = graph.node_weight(ix).unwrap().label {
571                targets.insert(label, ix);
572            }
573        }
574
575        // Consecutive blocks get fallthrough edges either for regular
576        // control flow into a label (continue), or for jumps not
577        // taken.
578        for consecutive in block_indices.windows(2) {
579            match *consecutive {
580                [ix1, ix2] => match graph.node_weight(ix1).unwrap().terminator {
581                    Terminator::Continue => {
582                        graph.add_edge(ix1, ix2, Edge::Continue);
583                    }
584                    Terminator::Jump(_, _, _) => {
585                        graph.add_edge(ix1, ix2, Edge::Jump(false));
586                    }
587                    _ => (),
588                },
589                _ => unreachable!(),
590            }
591        }
592
593        for ix in &block_indices {
594            match graph[*ix].terminator {
595                Terminator::Jump(_, target, _) => {
596                    let destination =
597                        targets.get(&target).unwrap_or_else(|| panic!("No block found for jump target {}!", target));
598                    graph.add_edge(*ix, *destination, Edge::Jump(true));
599                }
600                Terminator::Goto(target) => {
601                    let destination =
602                        targets.get(&target).unwrap_or_else(|| panic!("No block found for goto target {}!", target));
603                    graph.add_edge(*ix, *destination, Edge::Goto);
604                }
605                _ => (),
606            }
607        }
608
609        CFG { root: *block_indices.first().expect("Control flow graph has zero blocks!"), graph }
610    }
611
612    fn dominator_tree(&self, doms: &Dominators<NodeIndex>) -> DominatorTree {
613        let mut tree: HashMap<_, Vec<_>> = HashMap::new();
614
615        for ix in self.graph.node_indices() {
616            if let Some(idom) = doms.immediate_dominator(ix) {
617                tree.entry(idom).or_default().push(ix)
618            }
619        }
620
621        DominatorTree { tree }
622    }
623
624    fn node_writes(&self) -> Vec<HashSet<Name>> {
625        let mut writes = vec![HashSet::new(); self.graph.node_count()];
626
627        for ix in self.graph.node_indices() {
628            for instr in &self.graph.node_weight(ix).unwrap().instrs {
629                if let Some(id) = instr.write() {
630                    writes[ix.index()].insert(id);
631                }
632            }
633        }
634
635        writes
636    }
637
638    fn all_vars(&self) -> HashSet<Name> {
639        let mut vars = HashSet::new();
640
641        vars.insert(RETURN);
642
643        for ix in self.graph.node_indices() {
644            for instr in &self.graph[ix].instrs {
645                if let Some(id) = instr.declares() {
646                    vars.insert(id);
647                }
648            }
649        }
650
651        vars
652    }
653
654    /// Returns a HashMap of all variables declared in a CFG. Also
655    /// includes the special RETURN variable which is used to signal
656    /// the return value of a function, hence why the return type of
657    /// the function is also passed as an argument.
658    pub fn all_vars_typed(&self, ret_ty: &Ty<Name>) -> HashMap<Name, Ty<Name>> {
659        let mut vars = HashMap::new();
660
661        vars.insert(RETURN, ret_ty.clone());
662
663        for ix in self.graph.node_indices() {
664            for instr in &self.graph[ix].instrs {
665                if let Some((id, ty)) = instr.declares_typed() {
666                    vars.insert(id, ty);
667                }
668            }
669        }
670
671        vars
672    }
673
674    fn place_phi_functions(&mut self, frontiers: &DominanceFrontiers, all_vars: &HashSet<Name>) {
675        let node_writes: Vec<HashSet<Name>> = self.node_writes();
676        let mut needs_phi: HashMap<Name, HashSet<NodeIndex>> = HashMap::new();
677        let mut defsites: HashMap<Name, HashSet<NodeIndex>> = HashMap::new();
678
679        for ix in self.graph.node_indices() {
680            for a in &node_writes[ix.index()] {
681                defsites.entry(*a).or_default().insert(ix);
682            }
683        }
684
685        for a in all_vars {
686            let mut worklist: Vec<NodeIndex> = defsites.get_mut(a).unwrap().drain().collect();
687
688            while !worklist.is_empty() {
689                let n = worklist.pop().unwrap();
690
691                for y in frontiers.get(n) {
692                    if !needs_phi.entry(*a).or_default().contains(y) {
693                        let num_preds = self.graph.edges_directed(*y, Direction::Incoming).count();
694                        self.graph.node_weight_mut(*y).unwrap().insert_phi(*a, num_preds);
695                        needs_phi.entry(*a).or_default().insert(*y);
696                        if node_writes[y.index()].contains(a) {
697                            worklist.push(*y)
698                        }
699                    }
700                }
701            }
702        }
703    }
704
705    fn rename_node(
706        &mut self,
707        n: NodeIndex,
708        dominator_tree: &DominatorTree,
709        counts: &mut HashMap<Name, i32>,
710        stacks: &mut HashMap<Name, Vec<i32>>,
711    ) {
712        self.graph[n].rename(counts, stacks);
713
714        let succs: Vec<NodeIndex> = self.graph.neighbors_directed(n, Direction::Outgoing).collect();
715        for succ in succs {
716            let mut j = usize::MAX;
717            for (k, pred) in self.graph.neighbors_directed(succ, Direction::Incoming).enumerate() {
718                if pred == n {
719                    j = k;
720                    break;
721                }
722            }
723            assert!(j != usize::MAX);
724
725            self.graph[succ].rename_phi_arg(j, stacks)
726        }
727
728        if let Some(children) = dominator_tree.children(n) {
729            for child in children {
730                self.rename_node(*child, dominator_tree, counts, stacks)
731            }
732        }
733
734        for instr in &self.graph.node_weight(n).unwrap().instrs {
735            if let Some(name) = instr.write() {
736                stacks.get_mut(&name).and_then(Vec::pop);
737            }
738        }
739    }
740
741    fn rename(&mut self, dominator_tree: &DominatorTree, all_vars: &HashSet<Name>) {
742        let mut counts = HashMap::new();
743        let mut stacks: HashMap<Name, Vec<i32>> = HashMap::new();
744
745        for a in all_vars {
746            stacks.insert(*a, vec![0]);
747        }
748
749        self.rename_node(self.root, dominator_tree, &mut counts, &mut stacks)
750    }
751
752    /// Put the CFG into single static assignment (SSA) form.
753    pub fn ssa(&mut self) {
754        let dominators = dominators::simple_fast(&self.graph, self.root);
755        let dominator_tree = self.dominator_tree(&dominators);
756        let frontiers = DominanceFrontiers::from_graph(&self.graph, &dominators);
757        let all_vars: HashSet<Name> = self.all_vars();
758        self.place_phi_functions(&frontiers, &all_vars);
759        self.rename(&dominator_tree, &all_vars);
760    }
761
762    /// Generate a dot file of the CFG. For debugging.
763    pub fn dot(&self, output: &mut dyn Write, symtab: &Symtab) -> std::io::Result<()> {
764        writeln!(output, "digraph CFG {{")?;
765
766        for ix in self.graph.node_indices() {
767            let node = self.graph.node_weight(ix).unwrap();
768            write!(output, "  n{} [shape=box;style=filled;label=\"", ix.index())?;
769            for (id, args) in &node.phis {
770                id.write(output, symtab)?;
771                write!(output, " = Φ")?;
772                for (i, arg) in args.iter().enumerate() {
773                    if i == 0 {
774                        write!(output, "(")?
775                    } else {
776                        write!(output, ", ")?
777                    }
778                    arg.write(output, symtab)?
779                }
780                write!(output, ")\\n")?;
781            }
782            writeln!(output, "\"]")?
783        }
784
785        for edge in self.graph.edge_indices() {
786            if let Some((ix1, ix2)) = self.graph.edge_endpoints(edge) {
787                writeln!(output, "  n{} -> n{}", ix1.index(), ix2.index())?
788            }
789        }
790
791        writeln!(output, "}}")
792    }
793}