@@ -589,18 +589,24 @@ private def ppParamsAt (proof : Expr) (numParams : Nat) (paramPos : List Nat) :
589589 msg := msg ++ m!"{x} : {← inferType x}"
590590 addMessageContextFull msg
591591
592+ private def logPatternWhen (showInfo : Bool) (origin : Origin) (patterns : List Expr) : MetaM Unit := do
593+ if showInfo then
594+ logInfo m!"{← origin.pp}: {patterns.map ppPattern}"
595+
592596/--
593597Creates an E-matching theorem for a theorem with proof `proof`, `numParams` parameters, and the given set of patterns.
594598Pattern variables are represented using de Bruijn indices.
595599-/
596- def mkEMatchTheoremCore (origin : Origin) (levelParams : Array Name) (numParams : Nat) (proof : Expr) (patterns : List Expr) (kind : EMatchTheoremKind) : MetaM EMatchTheorem := do
600+ def mkEMatchTheoremCore (origin : Origin) (levelParams : Array Name) (numParams : Nat) (proof : Expr)
601+ (patterns : List Expr) (kind : EMatchTheoremKind) (showInfo := false ) : MetaM EMatchTheorem := do
597602 let (patterns, symbols, bvarFound) ← NormalizePattern.main patterns
598603 if symbols.isEmpty then
599604 throwError "invalid pattern for `{← origin.pp}`{indentD (patterns.map ppPattern)}\n the pattern does not contain constant symbols for indexing"
600- trace[grind.ematch.pattern] "{MessageData.ofConst proof }: {patterns.map ppPattern}"
605+ trace[grind.ematch.pattern] "{← origin.pp }: {patterns.map ppPattern}"
601606 if let .missing pos ← checkCoverage proof numParams bvarFound then
602607 let pats : MessageData := m!"{patterns.map ppPattern}"
603608 throwError "invalid pattern(s) for `{← origin.pp}`{indentD pats}\n the following theorem parameters cannot be instantiated:{indentD (← ppParamsAt proof numParams pos)}"
609+ logPatternWhen showInfo origin patterns
604610 return {
605611 proof, patterns, numParams, symbols
606612 levelParams, origin, kind
@@ -627,7 +633,7 @@ Given a theorem with proof `proof` and type of the form `∀ (a_1 ... a_n), lhs
627633creates an E-matching pattern for it using `addEMatchTheorem n [lhs]`
628634If `normalizePattern` is true, it applies the `grind` simplification theorems and simprocs to the pattern.
629635-/
630- def mkEMatchEqTheoremCore (origin : Origin) (levelParams : Array Name) (proof : Expr) (normalizePattern : Bool) (useLhs : Bool) : MetaM EMatchTheorem := do
636+ def mkEMatchEqTheoremCore (origin : Origin) (levelParams : Array Name) (proof : Expr) (normalizePattern : Bool) (useLhs : Bool) (showInfo := false ) : MetaM EMatchTheorem := do
631637 let (numParams, patterns) ← forallTelescopeReducing (← inferType proof) fun xs type => do
632638 let (lhs, rhs) ← match_expr type with
633639 | Eq _ lhs rhs => pure (lhs, rhs)
@@ -640,15 +646,15 @@ def mkEMatchEqTheoremCore (origin : Origin) (levelParams : Array Name) (proof :
640646 trace[grind.debug.ematch.pattern] "mkEMatchEqTheoremCore: after preprocessing: {pat}, {← normalize pat normConfig}"
641647 let pats := splitWhileForbidden (pat.abstract xs)
642648 return (xs.size, pats)
643- mkEMatchTheoremCore origin levelParams numParams proof patterns (if useLhs then .eqLhs else .eqRhs)
649+ mkEMatchTheoremCore origin levelParams numParams proof patterns (if useLhs then .eqLhs else .eqRhs) (showInfo := showInfo)
644650
645- def mkEMatchEqBwdTheoremCore (origin : Origin) (levelParams : Array Name) (proof : Expr) : MetaM EMatchTheorem := do
651+ def mkEMatchEqBwdTheoremCore (origin : Origin) (levelParams : Array Name) (proof : Expr) (showInfo := false ) : MetaM EMatchTheorem := do
646652 let (numParams, patterns) ← forallTelescopeReducing (← inferType proof) fun xs type => do
647653 let_expr f@Eq α lhs rhs := type
648654 | throwError "invalid E-matching `←=` theorem, conclusion must be an equality{indentExpr type}"
649655 let pat ← preprocessPattern (mkEqBwdPattern f.constLevels! α lhs rhs)
650656 return (xs.size, [pat.abstract xs])
651- mkEMatchTheoremCore origin levelParams numParams proof patterns .eqBwd
657+ mkEMatchTheoremCore origin levelParams numParams proof patterns .eqBwd (showInfo := showInfo)
652658
653659/--
654660Given theorem with name `declName` and type of the form `∀ (a_1 ... a_n), lhs = rhs`,
@@ -657,8 +663,8 @@ creates an E-matching pattern for it using `addEMatchTheorem n [lhs]`
657663If `normalizePattern` is true, it applies the `grind` simplification theorems and simprocs to the
658664pattern.
659665-/
660- def mkEMatchEqTheorem (declName : Name) (normalizePattern := true ) (useLhs : Bool := true ) : MetaM EMatchTheorem := do
661- mkEMatchEqTheoremCore (.decl declName) #[] (← getProofFor declName) normalizePattern useLhs
666+ def mkEMatchEqTheorem (declName : Name) (normalizePattern := true ) (useLhs : Bool := true ) (showInfo := false ) : MetaM EMatchTheorem := do
667+ mkEMatchEqTheoremCore (.decl declName) #[] (← getProofFor declName) normalizePattern useLhs (showInfo := showInfo)
662668
663669/--
664670Adds an E-matching theorem to the environment.
@@ -844,13 +850,13 @@ since the theorem is already in the `grind` state and there is nothing to be ins
844850-/
845851def mkEMatchTheoremWithKind?
846852 (origin : Origin) (levelParams : Array Name) (proof : Expr) (kind : EMatchTheoremKind)
847- (groundPatterns := true ) : MetaM (Option EMatchTheorem) := do
853+ (groundPatterns := true ) (showInfo := false ) : MetaM (Option EMatchTheorem) := do
848854 if kind == .eqLhs then
849- return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := true ) (useLhs := true ))
855+ return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := true ) (useLhs := true ) (showInfo := showInfo) )
850856 else if kind == .eqRhs then
851- return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := true ) (useLhs := false ))
857+ return (← mkEMatchEqTheoremCore origin levelParams proof (normalizePattern := true ) (useLhs := false ) (showInfo := showInfo) )
852858 else if kind == .eqBwd then
853- return (← mkEMatchEqBwdTheoremCore origin levelParams proof)
859+ return (← mkEMatchEqBwdTheoremCore origin levelParams proof (showInfo := showInfo) )
854860 let type ← inferType proof
855861 /-
856862 Remark: we should not use `forallTelescopeReducing` (with default reducibility) here
@@ -894,25 +900,26 @@ where
894900 return none
895901 let numParams := xs.size
896902 trace[grind.ematch.pattern] "{← origin.pp}: {patterns.map ppPattern}"
903+ logPatternWhen showInfo origin patterns
897904 return some {
898905 proof, patterns, numParams, symbols
899906 levelParams, origin, kind
900907 }
901908
902- def mkEMatchTheoremForDecl (declName : Name) (thmKind : EMatchTheoremKind) : MetaM EMatchTheorem := do
903- let some thm ← mkEMatchTheoremWithKind? (.decl declName) #[] (← getProofFor declName) thmKind
909+ def mkEMatchTheoremForDecl (declName : Name) (thmKind : EMatchTheoremKind) (showInfo := false ) : MetaM EMatchTheorem := do
910+ let some thm ← mkEMatchTheoremWithKind? (.decl declName) #[] (← getProofFor declName) thmKind (showInfo := showInfo)
904911 | throwError "`@{thmKind.toAttribute} theorem {declName}` {thmKind.explainFailure}, consider using different options or the `grind_pattern` command"
905912 return thm
906913
907- def mkEMatchEqTheoremsForDef? (declName : Name) : MetaM (Option (Array EMatchTheorem)) := do
914+ def mkEMatchEqTheoremsForDef? (declName : Name) (showInfo := false ) : MetaM (Option (Array EMatchTheorem)) := do
908915 let some eqns ← getEqnsFor? declName | return none
909916 eqns.mapM fun eqn => do
910- mkEMatchEqTheorem eqn (normalizePattern := true )
917+ mkEMatchEqTheorem eqn (normalizePattern := true ) (showInfo := showInfo)
911918
912- private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (thmKind : EMatchTheoremKind) (useLhs := true ) : MetaM Unit := do
919+ private def addGrindEqAttr (declName : Name) (attrKind : AttributeKind) (thmKind : EMatchTheoremKind) (useLhs := true ) (showInfo := false ) : MetaM Unit := do
913920 if wasOriginallyTheorem (← getEnv) declName then
914- ematchTheoremsExt.add (← mkEMatchEqTheorem declName (normalizePattern := true ) (useLhs := useLhs)) attrKind
915- else if let some thms ← mkEMatchEqTheoremsForDef? declName then
921+ ematchTheoremsExt.add (← mkEMatchEqTheorem declName (normalizePattern := true ) (useLhs := useLhs) (showInfo := showInfo) ) attrKind
922+ else if let some thms ← mkEMatchEqTheoremsForDef? declName (showInfo := showInfo) then
916923 unless useLhs do
917924 throwError "`{declName}` is a definition, you must only use the left-hand side for extracting patterns"
918925 thms.forM (ematchTheoremsExt.add · attrKind)
@@ -935,20 +942,20 @@ def EMatchTheorems.eraseDecl (s : EMatchTheorems) (declName : Name) : MetaM EMat
935942 throwErr
936943 return s.erase <| .decl declName
937944
938- def addEMatchAttr (declName : Name) (attrKind : AttributeKind) (thmKind : EMatchTheoremKind) : MetaM Unit := do
945+ def addEMatchAttr (declName : Name) (attrKind : AttributeKind) (thmKind : EMatchTheoremKind) (showInfo := false ) : MetaM Unit := do
939946 if thmKind == .eqLhs then
940- addGrindEqAttr declName attrKind thmKind (useLhs := true )
947+ addGrindEqAttr declName attrKind thmKind (useLhs := true ) (showInfo := showInfo)
941948 else if thmKind == .eqRhs then
942- addGrindEqAttr declName attrKind thmKind (useLhs := false )
949+ addGrindEqAttr declName attrKind thmKind (useLhs := false ) (showInfo := showInfo)
943950 else if thmKind == .eqBoth then
944- addGrindEqAttr declName attrKind thmKind (useLhs := true )
945- addGrindEqAttr declName attrKind thmKind (useLhs := false )
951+ addGrindEqAttr declName attrKind thmKind (useLhs := true ) (showInfo := showInfo)
952+ addGrindEqAttr declName attrKind thmKind (useLhs := false ) (showInfo := showInfo)
946953 else
947954 let info ← getConstInfo declName
948955 if !wasOriginallyTheorem (← getEnv) declName && !info.isCtor && !info.isAxiom then
949- addGrindEqAttr declName attrKind thmKind
956+ addGrindEqAttr declName attrKind thmKind (showInfo := showInfo)
950957 else
951- let thm ← mkEMatchTheoremForDecl declName thmKind
958+ let thm ← mkEMatchTheoremForDecl declName thmKind (showInfo := showInfo)
952959 ematchTheoremsExt.add thm attrKind
953960
954961def eraseEMatchAttr (declName : Name) : MetaM Unit := do
0 commit comments