Skip to content

Commit

Permalink
fix: accidental ownership with specialization
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha committed Jun 7, 2024
1 parent 745d77b commit d85d3d5
Show file tree
Hide file tree
Showing 5 changed files with 57 additions and 15 deletions.
8 changes: 4 additions & 4 deletions src/Lean/Compiler/IR/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -24,12 +24,12 @@ abbrev Index := Nat
/-- Variable identifier -/
structure VarId where
idx : Index
deriving Inhabited
deriving Inhabited, Repr

/-- Join point identifier -/
structure JoinPointId where
idx : Index
deriving Inhabited
deriving Inhabited, Repr

abbrev Index.lt (a b : Index) : Bool := a < b

Expand Down Expand Up @@ -83,7 +83,7 @@ inductive IRType where
| irrelevant | object | tobject
| struct (leanTypeName : Option Name) (types : Array IRType) : IRType
| union (leanTypeName : Name) (types : Array IRType) : IRType
deriving Inhabited
deriving Inhabited, Repr

namespace IRType

Expand Down Expand Up @@ -236,7 +236,7 @@ structure Param where
x : VarId
borrow : Bool
ty : IRType
deriving Inhabited
deriving Inhabited, Repr

@[export lean_ir_mk_param]
def mkParam (x : VarId) (borrow : Bool) (ty : IRType) : Param := ⟨x, borrow, ty⟩
Expand Down
3 changes: 2 additions & 1 deletion src/Lean/Compiler/IR/Borrow.lean
Original file line number Diff line number Diff line change
Expand Up @@ -258,7 +258,8 @@ def preserveTailCall (x : VarId) (v : Expr) (b : FnBody) : M Unit := do
let ctx ← read
match v, b with
| (Expr.fap g ys), (FnBody.ret (Arg.var z)) =>
if ctx.decls.any (·.name == g) && x == z then
-- NOTE: we currently support TCO for self-calls only
if ctx.currFn == g && x == z then
let ps ← getParamInfo (ParamMap.Key.decl g)
ownParamsUsingArgs ys ps
| _, _ => pure ()
Expand Down
25 changes: 25 additions & 0 deletions tests/lean/4240.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,25 @@
/-! Check for bug accidentally marking `m` as owned. -/

class MyClass (α : Type u) where

instance : MyClass Nat where

inductive MyOption (α : Type u) where
| none
| some (key : α)

namespace MyOption

def isSomeWithInstance [MyClass α] : MyOption α → Bool
| none => false
| some _ => true

def isSome : MyOption α → Bool
| none => false
| some _ => true

end MyOption

set_option trace.compiler.ir.result true in
def isSomeWithInstanceNat (m : { m : Array (MyOption Nat) // 0 < m.size }) : Bool :=
(m.1.uget 0 m.2).isSomeWithInstance
26 changes: 26 additions & 0 deletions tests/lean/4240.lean.expected.out
Original file line number Diff line number Diff line change
@@ -0,0 +1,26 @@

[result]
def MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 (x_1 : @& obj) : u8 :=
case x_1 : obj of
MyOption.none →
let x_2 : u8 := 0;
ret x_2
MyOption.some →
let x_3 : u8 := 1;
ret x_3
def isSomeWithInstanceNat (x_1 : @& obj) : u8 :=
let x_2 : usize := 0;
let x_3 : obj := Array.uget ◾ x_1 x_2 ◾;
let x_4 : u8 := MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 x_3;
dec x_3;
ret x_4
def MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1._boxed (x_1 : obj) : obj :=
let x_2 : u8 := MyOption.isSomeWithInstance._at.isSomeWithInstanceNat._spec_1 x_1;
dec x_1;
let x_3 : obj := box x_2;
ret x_3
def isSomeWithInstanceNat._boxed (x_1 : obj) : obj :=
let x_2 : u8 := isSomeWithInstanceNat x_1;
dec x_1;
let x_3 : obj := box x_2;
ret x_3
10 changes: 0 additions & 10 deletions tests/lean/603.lean

This file was deleted.

0 comments on commit d85d3d5

Please sign in to comment.