Equations
Instances For
Equations
- Lean.IR.Borrow.OwnedSet.beq x✝ x = match x✝, x with | (f₁, x₁), (f₂, x₂) => f₁ == f₂ && x₁ == x₂
Instances For
Equations
Equations
- Lean.IR.Borrow.OwnedSet.getHash x = match x with | (f, x) => mixHash (hash f) (hash x)
Instances For
Instances For
Equations
- s.insert k = Lean.HashMap.insert s k ()
Instances For
Equations
- s.contains k = Lean.HashMap.contains s k
Instances For
We perform borrow inference in a block of mutually recursive functions.
Join points are viewed as local functions, and are identified using
their local id + the name of the surrounding function.
We keep a mapping from function and joint points to parameters (Array Param).
Recall that Param contains the field borrow.
- decl: Lean.IR.FunId → Lean.IR.Borrow.ParamMap.Key
- jp: Lean.IR.FunId → Lean.IR.JoinPointId → Lean.IR.Borrow.ParamMap.Key
Instances For
Equations
Equations
- Lean.IR.Borrow.ParamMap.getHash x = match x with | Lean.IR.Borrow.ParamMap.Key.decl n => hash n | Lean.IR.Borrow.ParamMap.Key.jp n id => mixHash (hash n) (hash id)
Instances For
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.IR.Borrow.instToFormatParamMap = { format := Lean.IR.Borrow.ParamMap.fmt }
Equations
- Lean.IR.Borrow.instToStringParamMap = { toString := fun (m : Lean.IR.Borrow.ParamMap) => (Std.format m).pretty Std.Format.defWidth 0 }
Mark parameters that take a reference as borrow
Equations
- Lean.IR.Borrow.InitParamMap.initBorrow ps = Array.map (fun (p : Lean.IR.Param) => { x := p.x, borrow := p.ty.isObj, ty := p.ty }) ps
Instances For
We do perform borrow inference for constants marked as export.
Reason: we current write wrappers in C++ for using exported functions.
These wrappers use smart pointers such as object_ref.
When writing a new wrapper we need to know whether an argument is a borrow
inference or not.
We can revise this decision when we implement code for generating
the wrappers automatically.
Equations
- Lean.IR.Borrow.InitParamMap.initBorrowIfNotExported exported ps = if exported = true then ps else Lean.IR.Borrow.InitParamMap.initBorrow ps
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.IR.Borrow.mkInitParamMap env decls = StateT.run' (SeqRight.seqRight (Lean.IR.Borrow.InitParamMap.visitDecls env decls) fun (x : Unit) => get) ∅
Instances For
Apply the inferred borrow annotations stored at ParamMap to a block of mutually
recursive functions.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.IR.Borrow.applyParamMap decls map = Lean.IR.Borrow.ApplyParamMap.visitDecls decls map
Instances For
- env : Lean.Environment
- decls : Array Lean.IR.Decl
- currFn : Lean.IR.FunId
- paramSet : Lean.IR.IndexSet
Instances For
- owned : Lean.IR.Borrow.OwnedSetSet of variables that must be owned.
- modified : Bool
- paramMap : Lean.IR.Borrow.ParamMap
Instances For
Equations
Instances For
Equations
- Lean.IR.Borrow.getCurrFn = do let ctx ← read pure ctx.currFn
Instances For
Equations
- Lean.IR.Borrow.markModified = modify fun (s : Lean.IR.Borrow.BorrowInfState) => { owned := s.owned, modified := true, paramMap := s.paramMap }
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.IR.Borrow.ownArg x = match x with | Lean.IR.Arg.var x => Lean.IR.Borrow.ownVar x | x => pure ()
Instances For
Equations
Instances For
Equations
- Lean.IR.Borrow.isOwned x = do let currFn ← Lean.IR.Borrow.getCurrFn let s ← get pure (s.owned.contains (currFn, x.idx))
Instances For
Updates map[k] using the current set of owned variables.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
For each ps[i], if ps[i] is owned, then mark xs[i] as owned.
Equations
- Lean.IR.Borrow.ownArgsUsingParams xs ps = xs.size.forM fun (i : Nat) => let x := xs[i]!; let p := ps[i]!; if p.borrow = true then pure PUnit.unit else Lean.IR.Borrow.ownArg x
Instances For
For each xs[i], if xs[i] is owned, then mark ps[i] as owned.
We use this action to preserve tail calls. That is, if we have
a tail call f xs, if the i-th parameter is borrowed, but xs[i] is owned
we would have to insert a dec xs[i] after f xs and consequently
"break" the tail call.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Mark xs[i] as owned if it is one of the parameters ps.
We use this action to mark function parameters that are being "packed" inside constructors.
This is a heuristic, and is not related with the effectiveness of the reset/reuse optimization.
It is useful for code such as
def f (x y : obj) :=
let z := ctor_1 x y;
ret z
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Keep executing x until it reaches a fixpoint
Equations
- Lean.IR.Borrow.collectDecls = do let __do_lift ← read Lean.IR.Borrow.whileModifing (Array.forM Lean.IR.Borrow.collectDecl __do_lift.decls 0) let s ← get pure s.paramMap
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.IR.inferBorrow decls = do let env ← Lean.IR.getEnv let paramMap : Lean.IR.Borrow.ParamMap := Lean.IR.Borrow.infer env decls pure (Lean.IR.Borrow.applyParamMap decls paramMap)