Skip to content

Commit

Permalink
feat: generalize panic to Sort
Browse files Browse the repository at this point in the history
  • Loading branch information
tydeu committed Dec 10, 2024
1 parent 3f79193 commit 43f5568
Show file tree
Hide file tree
Showing 2 changed files with 4 additions and 4 deletions.
4 changes: 2 additions & 2 deletions src/Init/Prelude.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2550,7 +2550,7 @@ When we reimplement the specializer, we may consider copying `inst` if it also
occurs outside binders or if it is an instance.
-/
@[never_extract, extern "lean_panic_fn"]
def panicCore {α : Type u} [Inhabited α] (msg : String) : α := default
def panicCore {α : Sort u} [Inhabited α] (msg : String) : α := default

/--
`(panic "msg" : α)` has a built-in implementation which prints `msg` to
Expand All @@ -2564,7 +2564,7 @@ Because this is a pure function with side effects, it is marked as
elimination and other optimizations that assume that the expression is pure.
-/
@[noinline, never_extract]
def panic {α : Type u} [Inhabited α] (msg : String) : α :=
def panic {α : Sort u} [Inhabited α] (msg : String) : α :=
panicCore msg

-- TODO: this be applied directly to `Inhabited`'s definition when we remove the above workaround
Expand Down
4 changes: 2 additions & 2 deletions src/Init/Util.lean
Original file line number Diff line number Diff line change
Expand Up @@ -33,13 +33,13 @@ def dbgSleep {α : Type u} (ms : UInt32) (f : Unit → α) : α := f ()
@[noinline] private def mkPanicMessage (modName : String) (line col : Nat) (msg : String) : String :=
"PANIC at " ++ modName ++ ":" ++ toString line ++ ":" ++ toString col ++ ": " ++ msg

@[never_extract, inline] def panicWithPos {α : Type u} [Inhabited α] (modName : String) (line col : Nat) (msg : String) : α :=
@[never_extract, inline] def panicWithPos {α : Sort u} [Inhabited α] (modName : String) (line col : Nat) (msg : String) : α :=
panic (mkPanicMessage modName line col msg)

@[noinline] private def mkPanicMessageWithDecl (modName : String) (declName : String) (line col : Nat) (msg : String) : String :=
"PANIC at " ++ declName ++ " " ++ modName ++ ":" ++ toString line ++ ":" ++ toString col ++ ": " ++ msg

@[never_extract, inline] def panicWithPosWithDecl {α : Type u} [Inhabited α] (modName : String) (declName : String) (line col : Nat) (msg : String) : α :=
@[never_extract, inline] def panicWithPosWithDecl {α : Sort u} [Inhabited α] (modName : String) (declName : String) (line col : Nat) (msg : String) : α :=
panic (mkPanicMessageWithDecl modName declName line col msg)

@[extern "lean_ptr_addr"]
Expand Down

0 comments on commit 43f5568

Please sign in to comment.