Skip to content

Commit a6fffcf

Browse files
authored
Merge pull request #3 from leanprover-community/moogle
Moogle syntax
2 parents c260ed9 + 05f5fe6 commit a6fffcf

File tree

3 files changed

+189
-58
lines changed

3 files changed

+189
-58
lines changed

LeanSearchClient/Basic.lean

+5
Original file line numberDiff line numberDiff line change
@@ -4,3 +4,8 @@ register_option leansearch.queries : Nat :=
44
{ defValue := 6
55
group := "leansearch"
66
descr := "Number of results requested from leansearch (default 6)" }
7+
8+
register_option moogle.queries : Nat :=
9+
{ defValue := 6
10+
group := "moogle"
11+
descr := "Number of results requested from leansearch (default 6)" }

LeanSearchClient/MoogleExamples.lean

+34
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,34 @@
1+
import LeanSearchClient.Syntax
2+
3+
/-!
4+
# Moogle Examples
5+
6+
Examples of using the Moogle API. The search is triggered when the sentence ends with a full stop (period) or a question mark.
7+
-/
8+
9+
/--
10+
warning: #moogle query should end with a `.` or `?`.
11+
Note this command sends your query to an external service at https://www.moogle.ai/api/search.
12+
-/
13+
#guard_msgs in
14+
#moogle "If a natural number n is less than m, then the successor of n is less than the successor of m"
15+
16+
/--
17+
warning: #moogle query should end with a `.` or `?`.
18+
Note this command sends your query to an external service at https://www.moogle.ai/api/search.
19+
-/
20+
#guard_msgs in
21+
example := #moogle "If a natural number n is less than m, then the successor of n is less than the successor of m"
22+
23+
set_option moogle.queries 1
24+
25+
/--
26+
info: From: Nat.lt (type: Nat → Nat → Prop)
27+
• have : Nat → Nat → Prop := Nat.lt
28+
---
29+
warning: declaration uses 'sorry'
30+
-/
31+
#guard_msgs in
32+
example : 35 := by
33+
#moogle "If a natural number n is less than m, then the successor of n is less than the successor of m."
34+
sorry

LeanSearchClient/Syntax.lean

+150-58
Original file line numberDiff line numberDiff line change
@@ -11,17 +11,24 @@ import LeanSearchClient.Basic
1111
# LeanSearchClient
1212
1313
In this file, we provide syntax for search using the [leansearch API](https://leansearch.net/)
14+
and the [Moogle API](https://www.moogle.ai/api/search).
1415
from within Lean. It allows you to search for Lean tactics and theorems using natural language.
1516
1617
We provide syntax to make a query and generate `TryThis` options to click or
1718
use a code action to use the results.
1819
19-
The queries are of three forms:
20+
The queries are of three forms. For leansearch these are:
2021
2122
* `Command` syntax: `#leansearch "search query"` as a command.
2223
* `Term` syntax: `#leansearch "search query"` as a term.
2324
* `Tactic` syntax: `#leansearch "search query"` as a tactic.
2425
26+
The corresponding syntax for Moogle is:
27+
28+
* `Command` syntax: `#moogle "search query"` as a command.
29+
* `Term` syntax: `#moogle "search query"` as a term.
30+
* `Tactic` syntax: `#moogle "search query"` as a tactic.
31+
2532
In all cases results are displayed in the Lean Infoview and clicking these replaces the query text.
2633
In the cases of a query for tactics only valid tactics are displayed.
2734
-/
@@ -30,27 +37,43 @@ namespace LeanSearchClient
3037

3138
open Lean Meta Elab Tactic Parser Term
3239

33-
def getQueryJson (s : String) (num_results : Nat := 6) : IO <| Array Json := do
40+
def getLeanSearchQueryJson (s : String) (num_results : Nat := 6) : IO <| Array Json := do
3441
let apiUrl := "https://leansearch.net/api/search"
3542
let s' := s.replace " " "%20"
3643
let q := apiUrl ++ s!"?query={s'}&num_results={num_results}"
3744
let s ← IO.Process.output {cmd := "curl", args := #["-X", "GET", q]}
3845
let js := Json.parse s.stdout |>.toOption |>.get!
3946
return js.getArr? |>.toOption |>.get!
4047

48+
def getMoogleQueryJson (s : String) (num_results : Nat := 6) : IO <| Array Json := do
49+
let apiUrl := "https://www.moogle.ai/api/search"
50+
let data := Json.arr
51+
#[Json.mkObj [("isFind", false), ("contents", s)]]
52+
let s ← IO.Process.output {cmd := "curl", args := #[apiUrl, "-H", "content-type: application/json", "--data", data.pretty]}
53+
match Json.parse s.stdout with
54+
| Except.error e =>
55+
IO.throwServerError s!"Could not parse JSON from {s.stdout}; error: {e}"
56+
| Except.ok js =>
57+
let result? := js.getObjValAs? Json "data"
58+
match result? with
59+
| Except.ok result =>
60+
match result.getArr? with
61+
| Except.ok arr => return arr[0:num_results]
62+
| Except.error e => IO.throwServerError s!"Could not obtain array from {js}; error: {e}"
63+
| _ => IO.throwServerError s!"Could not obtain data from {js}"
64+
65+
4166
structure SearchResult where
4267
name : String
4368
type? : Option String
4469
docString? : Option String
4570
doc_url? : Option String
4671
kind? : Option String
47-
48-
def queryNum : CoreM Nat := do
49-
return leansearch.queries.get (← getOptions)
72+
deriving Repr
5073

5174
namespace SearchResult
5275

53-
def ofJson? (js : Json) : Option SearchResult :=
76+
def ofLeanSearchJson? (js : Json) : Option SearchResult :=
5477
match js.getObjValAs? String "formal_name" with
5578
| Except.ok name =>
5679
let type? := js.getObjValAs? String "formal_type" |>.toOption
@@ -61,10 +84,22 @@ def ofJson? (js : Json) : Option SearchResult :=
6184
some {name := name, type? := type?, docString? := doc?, doc_url? := docurl?, kind? := kind?}
6285
| _ => none
6386

64-
def query (s : String) (num_results : Nat) :
65-
IO <| Array SearchResult := do
66-
let jsArr ← getQueryJson s num_results
67-
return jsArr.filterMap ofJson?
87+
def ofMoogleJson? (js : Json) : MetaM <| Option SearchResult :=
88+
match js.getObjValAs? String "declarationName" with
89+
| Except.ok name => do
90+
let type? ←
91+
try
92+
let info ← getConstInfo name.toName
93+
let fmt ← PrettyPrinter.ppExpr info.type
94+
pure <| some fmt.pretty
95+
catch _ =>
96+
pure none
97+
let doc? := js.getObjValAs? String "declarationDocString" |>.toOption
98+
let doc? := doc?.filter fun s => s != ""
99+
let docurl? := none
100+
let kind? := js.getObjValAs? String "declarationType" |>.toOption
101+
return some {name := name, type? := type?, docString? := doc?, doc_url? := docurl?, kind? := kind?}
102+
| _ => return none
68103

69104
def toCommandSuggestion (sr : SearchResult) : TryThis.Suggestion :=
70105
let data := match sr.docString? with
@@ -87,24 +122,16 @@ def toTacticSuggestions (sr : SearchResult) : Array TryThis.Suggestion :=
87122

88123
end SearchResult
89124

90-
def getQueryCommandSuggestions (s : String) (num_results : Nat) :
91-
IO <| Array TryThis.Suggestion := do
92-
let searchResults ← SearchResult.query s num_results
93-
return searchResults.map SearchResult.toCommandSuggestion
94125

95-
def getQueryTermSuggestions (s : String) (num_results : Nat) :
96-
IO <| Array TryThis.Suggestion := do
97-
let searchResults SearchResult.query s num_results
98-
return searchResults.map SearchResult.toTermSuggestion
126+
def queryLeanSearch (s : String) (num_results : Nat) :
127+
MetaM <| Array SearchResult := do
128+
let jsArrgetLeanSearchQueryJson s num_results
129+
return jsArr.filterMap SearchResult.ofLeanSearchJson?
99130

100-
def getQueryTacticSuggestionGroups (s : String) (num_results : Nat) :
101-
IO <| Array (String × Array TryThis.Suggestion) := do
102-
let searchResults ← SearchResult.query s num_results
103-
return searchResults.map fun sr =>
104-
let fullName := match sr.type? with
105-
| some type => s!"{sr.name} (type: {type})"
106-
| none => sr.name
107-
(fullName, sr.toTacticSuggestions)
131+
def queryMoogle (s : String) (num_results : Nat) :
132+
MetaM <| Array SearchResult := do
133+
let jsArr ← getMoogleQueryJson s num_results
134+
jsArr.filterMapM SearchResult.ofMoogleJson?
108135

109136
def defaultTerm (expectedType? : Option Expr) : MetaM Expr := do
110137
match expectedType? with
@@ -128,52 +155,72 @@ def checkTactic (target : Expr) (tac : Syntax) :
128155
catch _ =>
129156
return none
130157

131-
def incompleteQuery : String :=
132-
"#leansearch query should end with a `.` or `?`.\n\
133-
Note this command sends your query to an external service at https://leansearch.net/."
158+
structure SearchServer where
159+
name : String
160+
url : String
161+
cmd: String
162+
query : String → Nat → MetaM (Array SearchResult)
163+
queryNum : CoreM Nat
134164

135-
open Command
165+
def leanSearchServer : SearchServer :=
166+
{name := "LeanSearch", cmd := "#leansearch", url := "https://leansearch.net/",
167+
query := queryLeanSearch, queryNum := return leansearch.queries.get (← getOptions)}
136168

137-
syntax (name := leansearch_cmd) "#leansearch" str : command
169+
def moogleServer : SearchServer :=
170+
{name := "Moogle", cmd := "#moogle", url := "https://www.moogle.ai/api/search",
171+
query := queryMoogle, queryNum := return moogle.queries.get (← getOptions)}
138172

139-
@[command_elab leansearch_cmd] def leanSearchCommandImpl : CommandElab :=
140-
fun stx => Command.liftTermElabM do
141-
match stx with
142-
| `(command| #leansearch $s) =>
173+
instance : Inhabited SearchServer := ⟨leanSearchServer⟩
174+
175+
namespace SearchServer
176+
177+
def getCommandSuggestions (ss : SearchServer) (s : String) (num_results : Nat) :
178+
MetaM (Array TryThis.Suggestion) := do
179+
let suggestions ← ss.query s num_results
180+
return suggestions.map SearchResult.toCommandSuggestion
181+
182+
def getTermSuggestions (ss : SearchServer) (s : String) (num_results : Nat) :
183+
MetaM (Array TryThis.Suggestion) := do
184+
let suggestions ← ss.query s num_results
185+
return suggestions.map SearchResult.toTermSuggestion
186+
187+
def getTacticSuggestionGroups (ss : SearchServer) (s : String) (num_results : Nat) :
188+
MetaM (Array (String × Array TryThis.Suggestion)) := do
189+
let suggestions ← ss.query s num_results
190+
return suggestions.map fun sr =>
191+
let fullName := match sr.type? with
192+
| some type => s!"{sr.name} (type: {type})"
193+
| none => sr.name
194+
(fullName, sr.toTacticSuggestions)
195+
196+
def incompleteSearchQuery (ss : SearchServer) : String :=
197+
s!"{ss.cmd} query should end with a `.` or `?`.\n\
198+
Note this command sends your query to an external service at {ss.url}."
199+
200+
open Command
201+
def searchCommandSuggestions (ss: SearchServer) (stx: Syntax) (s: TSyntax `str) : CommandElabM Unit := Command.liftTermElabM do
143202
let s := s.getString
144203
if s.endsWith "." || s.endsWith "?" then
145-
let suggestions ← getQueryCommandSuggestions s (← queryNum)
146-
TryThis.addSuggestions stx suggestions (header := "Lean Search Results")
204+
let suggestions ← ss.getCommandSuggestions s (← ss.queryNum)
205+
TryThis.addSuggestions stx suggestions (header := s!"{ss.name} Search Results")
147206
else
148-
logWarning incompleteQuery
149-
| _ => throwUnsupportedSyntax
150-
151-
syntax (name := leansearch_term) "#leansearch" str : term
207+
logWarning <| ss.incompleteSearchQuery
152208

153-
@[term_elab leansearch_term] def leanSearchTermImpl : TermElab :=
154-
fun stx expectedType? => do
155-
match stx with
156-
| `(#leansearch $s) =>
209+
def searchTermSuggestions (ss: SearchServer) (stx: Syntax)
210+
(s: TSyntax `str) : TermElabM Unit := do
157211
let s := s.getString
158212
if s.endsWith "." || s.endsWith "?" then
159-
let suggestions ← getQueryTermSuggestions s (← queryNum)
160-
TryThis.addSuggestions stx suggestions (header := "Lean Search Results")
213+
let suggestions ← ss.getTermSuggestions s (← ss.queryNum)
214+
TryThis.addSuggestions stx suggestions (header := s!"{ss.name} Search Results")
161215
else
162-
logWarning incompleteQuery
163-
defaultTerm expectedType?
164-
| _ => throwUnsupportedSyntax
165-
166-
syntax (name := leansearch_tactic) "#leansearch" str : tactic
216+
logWarning <| ss.incompleteSearchQuery
167217

168-
@[tactic leansearch_tactic] def leanSearchTacticImpl : Tactic :=
169-
fun stx => withMainContext do
170-
match stx with
171-
| `(tactic|#leansearch $s) =>
218+
def searchTacticSuggestions (ss: SearchServer) (stx: Syntax) (s: TSyntax `str) : TacticM Unit := do
172219
let s := s.getString
173220
if s.endsWith "." || s.endsWith "?" then
174221
let target ← getMainTarget
175222
let suggestionGroups ←
176-
getQueryTacticSuggestionGroups s (← queryNum)
223+
ss.getTacticSuggestionGroups s (← ss.queryNum)
177224
for (name, sg) in suggestionGroups do
178225
let sg ← sg.filterM fun s =>
179226
let sugTxt := s.suggestion
@@ -190,5 +237,50 @@ syntax (name := leansearch_tactic) "#leansearch" str : tactic
190237
unless sg.isEmpty do
191238
TryThis.addSuggestions stx sg (header := s!"From: {name}")
192239
else
193-
logWarning incompleteQuery
240+
logWarning <| ss.incompleteSearchQuery
241+
242+
end SearchServer
243+
244+
open Command
245+
246+
declare_syntax_cat search_cmd
247+
248+
syntax "#leansearch" : search_cmd
249+
syntax "#moogle" : search_cmd
250+
251+
syntax (name := search_cmd) search_cmd str : command
252+
253+
def getSearchServer (sc : Syntax) : SearchServer :=
254+
match sc with
255+
| `(search_cmd| #leansearch) => leanSearchServer
256+
| `(search_cmd| #moogle) => moogleServer
257+
| _ => panic! "unsupported search command"
258+
259+
@[command_elab search_cmd] def searchCommandImpl : CommandElab :=
260+
fun stx =>
261+
match stx with
262+
| `(command| $sc:search_cmd $s) => do
263+
let ss := getSearchServer sc
264+
ss.searchCommandSuggestions stx s
265+
| _ => throwUnsupportedSyntax
266+
267+
syntax (name := search_term) search_cmd str : term
268+
269+
@[term_elab search_term] def searchTermImpl : TermElab :=
270+
fun stx expectedType? => do
271+
match stx with
272+
| `($sc:search_cmd $s) =>
273+
let ss := getSearchServer sc
274+
ss.searchTermSuggestions stx s
275+
defaultTerm expectedType?
276+
| _ => throwUnsupportedSyntax
277+
278+
syntax (name := search_tactic) search_cmd str : tactic
279+
280+
@[tactic search_tactic] def searchTacticImpl : Tactic :=
281+
fun stx => withMainContext do
282+
match stx with
283+
| `(tactic|$sc:search_cmd $s) =>
284+
let ss := getSearchServer sc
285+
ss.searchTacticSuggestions stx s
194286
| _ => throwUnsupportedSyntax

0 commit comments

Comments
 (0)