Skip to content

Commit d3cb87d

Browse files
committed
feat: add ` (command| ..) `` quote (#116)
1 parent d9f8224 commit d3cb87d

File tree

2 files changed

+15
-0
lines changed

2 files changed

+15
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -39,6 +39,7 @@ import Mathlib.Mathport.Syntax
3939
import Mathlib.Tactic.Basic
4040
import Mathlib.Tactic.Cache
4141
import Mathlib.Tactic.Coe
42+
import Mathlib.Tactic.CommandQuote
4243
import Mathlib.Tactic.Core
4344
import Mathlib.Tactic.Ext
4445
import Mathlib.Tactic.Find

Mathlib/Tactic/CommandQuote.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,14 @@
1+
/-
2+
Copyright (c) 2021 Microsoft Corporation. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Mario Carneiro
5+
-/
6+
import Lean
7+
8+
namespace Lean
9+
10+
@[termParser default+1] def Parser.Term.Command.quot : Parser :=
11+
leading_parser "`(command|" >> incQuotDepth commandParser >> ")"
12+
13+
@[termElab Command.quot] def Elab.Term.elabCommandQuot : TermElab :=
14+
adaptExpander Quotation.stxQuot.expand

0 commit comments

Comments
 (0)