Skip to content

Commit

Permalink
feat: (↑) notation for coercions (#199)
Browse files Browse the repository at this point in the history
See leanprover-community/mathport#116 (comment)

The notation is a bit problematic for the pretty-printer, as it doesn't automatically insert parentheses to disambiguate it from the built-in coercion notation.  Alternatives:
 1. Use a different token.
 2. Magically define `(↑ ∙)` to be the eta-reduction of `fun x => ↑x`.

PS: the old macro for `↑ x` was redundant since it's been moved to core some time ago.
  • Loading branch information
gebner committed Feb 21, 2022
1 parent 792d625 commit 5544193
Show file tree
Hide file tree
Showing 2 changed files with 16 additions and 12 deletions.
2 changes: 1 addition & 1 deletion Mathlib/Data/Subtype.lean
Original file line number Diff line number Diff line change
Expand Up @@ -84,7 +84,7 @@ f x
lemma restrict_apply {α} {β : α → Type _} (f : ∀x, β x) (p : α → Prop) (x : Subtype p) :
restrict f p x = f x.1 := rfl

lemma restrict_def {α β} (f : α → β) (p : α → Prop) : restrict f p = f ∘ CoeHead.coe := rfl
lemma restrict_def {α β} (f : α → β) (p : α → Prop) : restrict f p = f ∘ (↑) := rfl

lemma restrict_injective {α β} {f : α → β} (p : α → Prop) (h : injective f) :
injective (restrict f p) :=
Expand Down
26 changes: 15 additions & 11 deletions Mathlib/Tactic/Coe.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,20 +5,24 @@ Author: Gabriel Ebner
-/
import Lean

open Lean Elab Term
open Lean Elab Term Meta

/-!
Redefine the ↑-notation to elaborate in the same way as type annotations
(i.e., unfolding the coercion instance).
Define a `(↑)` notation for coercions equivalent to the eta-reduction of `(↑ ·)`.
-/

namespace Lean.Elab.Term.CoeImpl

scoped elab "coe%" x:term : term <= expectedType => do
tryPostponeIfMVar expectedType
let x ← elabTerm x none
synthesizeSyntheticMVarsUsingDefault
ensureHasType expectedType x

macro_rules
| `(↑ $x) => `(coe% $x)
elab "(" "↑" ")" : term <= expectedType => do
let expectedType ← instantiateMVars expectedType
let Expr.forallE _ a b .. := expectedType | do
tryPostpone
throwError "(↑) must have a function type, not{indentExpr expectedType}"
if b.hasLooseBVars then
tryPostpone
throwError "(↑) must have a non-dependent function type, not{indentExpr expectedType}"
if a.hasExprMVar then tryPostpone
if b.hasExprMVar then tryPostpone
let f ← withLocalDeclD `x a fun x => do
mkLambdaFVars #[x] (← mkCoe b a x)
return f.etaExpanded?.getD f

0 comments on commit 5544193

Please sign in to comment.