diff --git a/Batteries/Lean/Name.lean b/Batteries/Lean/Name.lean index b3705e6a2e..b5530e49d1 100644 --- a/Batteries/Lean/Name.lean +++ b/Batteries/Lean/Name.lean @@ -16,8 +16,13 @@ Generally, user code should not explicitly use internal names. def isInternalDetail : Name → Bool | .str p s => s.startsWith "_" - || s.startsWith "match_" - || s.startsWith "proof_" + || matchPrefix s "eq_" + || matchPrefix s "match_" + || matchPrefix s "proof_" || p.isInternalOrNum | .num _ _ => true | p => p.isInternalOrNum +where + /-- Check that a string begins with the given prefix, and then is only digit characters. -/ + matchPrefix (s : String) (pre : String) := + s.startsWith pre && (s |>.drop pre.length |>.all Char.isDigit) diff --git a/Batteries/Util/ExtendedBinder.lean b/Batteries/Util/ExtendedBinder.lean index 33629b7cc0..d577e1d785 100644 --- a/Batteries/Util/ExtendedBinder.lean +++ b/Batteries/Util/ExtendedBinder.lean @@ -52,22 +52,3 @@ macro_rules -- TODO: merging the two macro_rules breaks expansion | `(∀ᵉ _ : $ty:term, $b) => `(∀ _ : $ty:term, $b) | `(∀ᵉ $x:ident : $ty:term, $b) => `(∀ $x:ident : $ty:term, $b) | `(∀ᵉ $x:binderIdent $p:binderPred, $b) => `(∀ $x:binderIdent $p:binderPred, $b) - -open Parser.Command in -/-- -Declares a binder predicate. For example: -``` -binder_predicate x " > " y:term => `($x > $y) -``` --/ -syntax (name := binderPredicate) (docComment)? (Parser.Term.attributes)? (attrKind)? - "binder_predicate" optNamedName optNamedPrio ppSpace ident (ppSpace macroArg)* " => " - term : command - -open Linter.MissingDocs Parser Term in -/-- Missing docs handler for `binder_predicate` -/ -@[missing_docs_handler binderPredicate] -def checkBinderPredicate : SimpleHandler := fun stx => do - if stx[0].isNone && stx[2][0][0].getKind != ``«local» then - if stx[4].isNone then lint stx[3] "binder predicate" - else lintNamed stx[4][0][3] "binder predicate"