Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit b39040f

Browse files
fpvandoornmergify[bot]
authored andcommitted
feat(sanity_check): improvements (#1487)
* checkpoint * checkpoint * checkpoint * feat(sanity_check): improvements * Now check for classical.[some|choice] and [gt|ge] in theorem statements * Add [sanity_skip] attribute to skip checks * Add #sanity_check_all to check all declarations * Add `-` after command to only execute fast tests * Reduce code duplication * Make the performed checks configurable. * some cleanup * fix tests * clarification * reviewer comments * fix typo in docstring * reviewer comments * update docstring * Update tactics.md
1 parent 3e5dc88 commit b39040f

File tree

4 files changed

+187
-99
lines changed

4 files changed

+187
-99
lines changed

docs/tactics.md

Lines changed: 21 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1138,10 +1138,29 @@ Here too, the `reassoc` attribute can be used instead. It works well when combin
11381138
attribute [simp, reassoc] some_class.bar
11391139
```
11401140
### sanity_check
1141+
User commands to spot common mistakes in the code
11411142

1142-
The `#sanity_check` command checks for common mistakes in the current file or in all of mathlib, respectively.
1143+
* `#sanity_check`: check all declarations in the current file
1144+
* `#sanity_check_mathlib`: check all declarations in mathlib (so excluding core or other projects,
1145+
and also excluding the current file)
1146+
* `#sanity_check_all`: check all declarations in the environment (the current file and all
1147+
imported files)
11431148

1144-
Currently this will check for unused arguments in declarations and whether a declaration is incorrectly marked as a def/lemma.
1149+
Currently this will check for
1150+
1151+
1. unused arguments in declarations,
1152+
2. whether a declaration is incorrectly marked as a def/lemma,
1153+
3. whether a namespace is duplicated in the name of a declaration
1154+
4. whether ≥/> is used in the declaration
1155+
1156+
You can append a `-` to any command (e.g. `#sanity_check_mathlib-`) to omit the slow tests (4).
1157+
1158+
You can customize the performed checks like this:
1159+
```lean
1160+
meta def my_check (d : declaration) : tactic (option string) :=
1161+
return $ if d.to_name = `foo then some "gotcha!" else none
1162+
run_cmd sanity_check tt [(my_check, "found nothing", "found something")] >>= trace
1163+
```
11451164

11461165
### lift
11471166

src/meta/expr.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -186,6 +186,11 @@ e.fold mk_name_set $ λ e' _ l,
186186
| _ := l
187187
end
188188

189+
/-- Returns true if `e` contains a name `n` where `p n` is true.
190+
Returns `true` if `p name.anonymous` is true -/
191+
meta def contains_constant (e : expr) (p : name → Prop) [decidable_pred p] : bool :=
192+
e.fold ff (λ e' _ b, if p (e'.const_name) then tt else b)
193+
189194
/--
190195
is_num_eq n1 n2 returns true if n1 and n2 are both numerals with the same numeral structure,
191196
ignoring differences in type and type class arguments.
@@ -291,7 +296,7 @@ e.fold (return x) (λ d t, t >>= fn d)
291296
meta def mfilter (e : environment) (test : declaration → tactic bool) : tactic (list declaration) :=
292297
e.mfold [] $ λ d ds, do b ← test d, return $ if b then d::ds else ds
293298

294-
/-- Checks whether `ml` is a prefix of the file where `n` is declared.
299+
/-- Checks whether `s` is a prefix of the file where `n` is declared.
295300
This is used to check whether `n` is declared in mathlib, where `s` is the mathlib directory. -/
296301
meta def is_prefix_of_file (e : environment) (s : string) (n : name) : bool :=
297302
s.is_prefix_of $ (e.decl_olean n).get_or_else ""

src/tactic/sanity_check.lean

Lines changed: 132 additions & 92 deletions
Original file line numberDiff line numberDiff line change
@@ -3,24 +3,46 @@ Copyright (c) 2019 Floris van Doorn. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Floris van Doorn
55
-/
6-
76
import tactic.core
8-
97
/-!
10-
# sanity_check command
11-
This file defines the `#sanity_check` and `#sanity_check_mathlib` commands, to spot common mistakes in the current file or in all of mathlib, respectively.
8+
# sanity_check command
9+
This file defines the following user commands to spot common mistakes in the code.
10+
* `#sanity_check`: check all declarations in the current file
11+
* `#sanity_check_mathlib`: check all declarations in mathlib (so excluding core or other projects,
12+
and also excluding the current file)
13+
* `#sanity_check_all`: check all declarations in the environment (the current file and all
14+
imported files)
15+
16+
Currently this will check for
17+
1. unused arguments in declarations,
18+
2. whether a declaration is incorrectly marked as a def/lemma,
19+
3. whether a namespace is duplicated in the name of a declaration
20+
4. whether ≥/> is used in the declaration
21+
22+
You can append a `-` to any command (e.g. `#sanity_check_mathlib-`) to omit the slow tests (4).
1223
13-
Currently this will check for unused arguments in declarations and whether a declaration is incorrectly marked as a def/lemma.
24+
You can customize the performed checks like this:
25+
```
26+
meta def my_check (d : declaration) : tactic (option string) :=
27+
return $ if d.to_name = `foo then some "gotcha!" else none
28+
run_cmd sanity_check tt [(my_check, "found nothing", "found something")] >>= trace
29+
```
1430
15-
## Tags
16-
sanity check, sanity_check, cleanup, command, tactic
31+
## Tags
32+
sanity check, sanity_check, cleanup, command, tactic
1733
-/
1834

1935
universe variable u
2036
open expr tactic native
2137

2238
reserve notation `#sanity_check`
2339
reserve notation `#sanity_check_mathlib`
40+
reserve notation `#sanity_check_all`
41+
42+
@[user_attribute]
43+
meta def sanity_skip_attr : user_attribute :=
44+
{ name := "sanity_skip",
45+
descr := "Do not report this declaration in any of the tests of `#sanity_check`" }
2446

2547
setup_tactic_parser
2648
universe variable v
@@ -32,11 +54,11 @@ l.mmap_filter $ λ d, option.map (λ x, (d, x)) <$> tac d
3254

3355
/-- Find all declarations in `l` where tac returns `some x` and sort the resulting list by file name. -/
3456
meta def fold_over_with_cond_sorted {α} (l : list declaration)
35-
(tac : declaration → tactic (option α)) : tactic (list (string × list (declaration × α))) :=
36-
do e ← get_env,
37-
ds ← fold_over_with_cond l tac,
38-
let ds₂ := rb_lmap.of_list (ds.map (λ x, ((e.decl_olean x.1.to_name).iget, x))),
39-
return $ ds₂.to_list
57+
(tac : declaration → tactic (option α)) : tactic (list (string × list (declaration × α))) := do
58+
e ← get_env,
59+
ds ← fold_over_with_cond l tac,
60+
let ds₂ := rb_lmap.of_list (ds.map (λ x, ((e.decl_olean x.1.to_name).iget, x))),
61+
return $ ds₂.to_list
4062

4163
/-- Make the output of `fold_over_with_cond` printable, in the following form:
4264
#print <name> <open multiline comment> <elt of α> <close multiline comment> -/
@@ -49,44 +71,17 @@ ds.foldl
4971
meta def print_decls_sorted {α} [has_to_format α] (ds : list (string × list (declaration × α))) :
5072
format :=
5173
ds.foldl
52-
(λ f x, f ++ "\n" ++ "\n" ++ to_fmt "-- " ++ to_fmt x.1 ++ print_decls x.2)
74+
(λ f x, f ++ "\n\n" ++ to_fmt "-- " ++ to_fmt x.1 ++ print_decls x.2)
5375
format.nil
5476

5577
/-- Same as `print_decls_sorted`, but removing the first `n` characters from the string.
5678
Useful for omitting the mathlib directory from the output. -/
5779
meta def print_decls_sorted_mathlib {α} [has_to_format α] (n : ℕ)
5880
(ds : list (string × list (declaration × α))) : format :=
5981
ds.foldl
60-
(λ f x, f ++ "\n" ++ "\n" ++ to_fmt "-- " ++ to_fmt (x.1.popn n) ++ print_decls x.2)
82+
(λ f x, f ++ "\n\n" ++ to_fmt "-- " ++ to_fmt (x.1.popn n) ++ print_decls x.2)
6183
format.nil
6284

63-
/- Print all (non-internal) declarations where tac return `some x`-/
64-
meta def print_all_decls {α} [has_to_format α] (tac : declaration → tactic (option α)) :
65-
tactic format :=
66-
do
67-
e ← get_env,
68-
l ← e.mfilter (λ d, return $ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e),
69-
print_decls_sorted <$> fold_over_with_cond_sorted l tac
70-
71-
/- Print (non-internal) declarations in the current file where tac return `some x`-/
72-
meta def print_decls_current_file {α} [has_to_format α] (tac : declaration → tactic (option α)) :
73-
tactic format :=
74-
do
75-
e ← get_env,
76-
l ← e.mfilter (λ d, return $
77-
e.in_current_file' d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e),
78-
print_decls <$> fold_over_with_cond l tac
79-
80-
/- Print (non-internal) declarations in mathlib where tac return `some x` -/
81-
meta def print_decls_mathlib {α} [has_to_format α] (tac : declaration → tactic (option α)) :
82-
tactic format :=
83-
do
84-
e ← get_env,
85-
ml ← get_mathlib_dir,
86-
l ← e.mfilter (λ d, return $
87-
e.is_prefix_of_file ml d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e),
88-
print_decls_sorted_mathlib ml.length <$> fold_over_with_cond_sorted l tac
89-
9085
/-- Auxilliary definition for `check_unused_arguments` -/
9186
meta def check_unused_arguments_aux : list ℕ → ℕ → ℕ → expr → list ℕ | l n n_max e :=
9287
if n > n_max then l else
@@ -112,8 +107,7 @@ let l2 := check_unused_arguments_aux [] 1 d.type.pi_arity d.type in
112107
the argument is a duplicate.
113108
See also `check_unused_arguments`.
114109
This tactic additionally filters out all unused arguments of type `parse _` -/
115-
meta def unused_arguments (d : declaration) : tactic (option string) :=
116-
do
110+
meta def unused_arguments (d : declaration) : tactic (option string) := do
117111
let ns := check_unused_arguments d,
118112
if ¬ ns.is_some then return none else do
119113
let ns := ns.iget,
@@ -126,22 +120,9 @@ do
126120
(if ds.countp (λ b', b.type = b'.type) ≥ 2 then " (duplicate)" else "")) <$> pp b),
127121
return $ some $ ns.to_string_aux tt
128122

129-
/-- Print all declarations with unused arguments -/
130-
meta def get_all_unused_args : tactic unit :=
131-
print_all_decls unused_arguments >>= trace
132-
133-
/-- Print all declarations in mathlib with unused arguments -/
134-
meta def get_all_unused_args_mathlib : tactic unit :=
135-
print_decls_mathlib unused_arguments >>= trace
136-
137-
/-- Print all declarations in current file with unused arguments. -/
138-
meta def get_all_unused_args_current_file : tactic unit :=
139-
print_decls_current_file unused_arguments >>= trace
140-
141123
/-- Checks whether the correct declaration constructor (definition of theorem) by comparing it
142124
to its sort. Instances will not be printed -/
143-
meta def incorrect_def_lemma (d : declaration) : tactic (option string) :=
144-
do
125+
meta def incorrect_def_lemma (d : declaration) : tactic (option string) := do
145126
e ← get_env,
146127
expr.sort n ← infer_type d.type,
147128
let is_def : Prop := d.is_definition,
@@ -152,70 +133,129 @@ do
152133
else if (d.is_definition : bool) then "is a def, should be a lemma/theorem"
153134
else "is a lemma/theorem, should be a def"
154135

155-
/-- Print all declarations in mathlib incorrectly marked as def/lemma -/
156-
meta def incorrect_def_lemma_mathlib : tactic unit :=
157-
print_decls_mathlib unused_arguments >>= trace
158-
159136
/-- Checks whether a declaration has a namespace twice consecutively in its name -/
160137
meta def dup_namespace (d : declaration) : tactic (option string) :=
161138
return $ let nm := d.to_name.components in if nm.chain' (≠) then none
162139
else let s := (nm.find $ λ n, nm.count n ≥ 2).iget.to_string in
163140
some $ "The namespace `" ++ s ++ "` is duplicated in the name"
164141

142+
/-- Checks whether a `>`/`≥` is used in the statement of `d`. -/
143+
-- TODO: the commented out code also checks for classicality in statements, but needs fixing
144+
-- TODO: this probably needs to also check whether the argument is a variable or @eq <var> _ _
145+
-- meta def illegal_constants_in_statement (d : declaration) : tactic (option string) :=
146+
-- return $ if d.type.contains_constant (λ n, (n.get_prefix = `classical ∧
147+
-- n.last ∈ ["prop_decidable", "dec", "dec_rel", "dec_eq"]) ∨ n ∈ [`gt, `ge])
148+
-- then
149+
-- let illegal1 := [`classical.prop_decidable, `classical.dec, `classical.dec_rel, `classical.dec_eq],
150+
-- illegal2 := [`gt, `ge],
151+
-- occur1 := illegal1.filter (λ n, d.type.contains_constant (eq n)),
152+
-- occur2 := illegal2.filter (λ n, d.type.contains_constant (eq n)) in
153+
-- some $ sformat!"the type contains the following declarations: {occur1 ++ occur2}." ++
154+
-- (if occur1 = [] then "" else " Add decidability type-class arguments instead.") ++
155+
-- (if occur2 = [] then "" else " Use ≤/< instead.")
156+
-- else none
157+
meta def illegal_constants_in_statement (d : declaration) : tactic (option string) :=
158+
return $ let illegal := [`gt, `ge] in if d.type.contains_constant (λ n, n ∈ illegal)
159+
then some "the type contains ≥/>. Use ≤/< instead."
160+
else none
161+
162+
/- The quick checks of `#sanity_check` -/
163+
meta def quick_checks : list ((declaration → tactic (option string)) × string × string) :=
164+
[ (unused_arguments, "No unused arguments", "UNUSED ARGUMENTS"),
165+
(incorrect_def_lemma, "All declarations correctly marked as def/lemma", "INCORRECT DEF/LEMMA"),
166+
(dup_namespace, "No declarations have a duplicate namespace", "DUPLICATED NAMESPACES IN NAME")]
167+
168+
/- The slow checks of `#sanity_check`. They are not executed with `#sanity_check-` -/
169+
meta def slow_checks : list ((declaration → tactic (option string)) × string × string) :=
170+
[ (illegal_constants_in_statement, "No illegal constants in declarations",
171+
"ILLEGAL CONSTANTS IN DECLARATIONS")]
172+
173+
/- The default checks of `#sanity_check`. Depends on whether `slow` is true -/
174+
meta def default_checks (slow : bool := tt) :
175+
list ((declaration → tactic (option string)) × string × string) :=
176+
quick_checks ++ if slow then slow_checks else []
177+
178+
/-- The common denominator of `#sanity_check[|mathlib|all]`.
179+
The different commands have different configurations for `l`, `printer` and `where_desc`.
180+
If `slow` is false, doesn't do the checks that take a lot of time.
181+
By setting `checks` you can customize which checks are performed. -/
182+
meta def sanity_check_aux (l : list declaration)
183+
(printer : (declaration → tactic (option string)) → tactic format)
184+
(where_desc : string) (slow : bool)
185+
(checks : list ((declaration → tactic (option string)) × string × string)) : tactic format := do
186+
let s : format := "/- Note: This command is still in development. -/\n",
187+
let s := s ++ format!"/- Checking {l.length} declarations {where_desc} -/\n\n",
188+
s ← checks.mfoldl (λ s ⟨tac, ok_string, warning_string⟩, show tactic format, from do
189+
f ← printer tac,
190+
return $ s ++ if f.is_nil then format!"/- OK: {ok_string}. -/\n"
191+
else format!"/- {warning_string}: -/" ++ f ++ "\n\n") s,
192+
return $ if slow then s else s ++ "/- (slow tests skipped) -/\n"
193+
165194
/-- Return the message printed by `#sanity_check`. -/
166-
meta def sanity_check : tactic format :=
167-
do
195+
meta def sanity_check (slow : bool := tt)
196+
(checks : list ((declaration → tactic (option string)) × string × string) :=
197+
default_checks slow) : tactic format := do
168198
e ← get_env,
169199
l ← e.mfilter (λ d,
170-
return $ e.in_current_file' d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e),
171-
let s : format := "/- Note: This command is still in development. -/\n",
172-
let s := s ++ "/- Checking " ++ l.length ++ " declarations in the current file -/\n\n",
173-
f ← print_decls <$> fold_over_with_cond l unused_arguments,
174-
let s := s ++ if f.is_nil then "/- OK: No unused arguments in the current file. -/\n"
175-
else "/- Unused arguments in the current file: -/" ++ f ++ "\n\n",
176-
f ← print_decls <$> fold_over_with_cond l incorrect_def_lemma,
177-
let s := s ++ if f.is_nil then "/- OK: All declarations correctly marked as def/lemma -/\n"
178-
else "/- Declarations incorrectly marked as def/lemma: -/" ++ f ++ "\n\n",
179-
f ← print_decls <$> fold_over_with_cond l dup_namespace,
180-
let s := s ++ if f.is_nil then "/- OK: No declarations have a duplicate namespace -/\n"
181-
else "/- Declarations with a namespace duplicated: -/" ++ f ++ "\n\n",
182-
return s
200+
if e.in_current_file' d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e
201+
then bnot <$> has_attribute' `sanity_skip d.to_name else return ff),
202+
sanity_check_aux l (λ t, print_decls <$> fold_over_with_cond l t)
203+
"in the current file" slow checks
183204

184205
/-- Return the message printed by `#sanity_check_mathlib`. -/
185-
meta def sanity_check_mathlib : tactic format :=
186-
do
206+
meta def sanity_check_mathlib (slow : bool := tt)
207+
(checks : list ((declaration → tactic (option string)) × string × string) :=
208+
default_checks slow) : tactic format := do
187209
e ← get_env,
188210
ml ← get_mathlib_dir,
189-
l ← e.mfilter (λ d, return $
190-
e.is_prefix_of_file ml d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e),
211+
/- note: we don't separate out some of these tests in `sanity_check_aux` because that causes a
212+
performance hit. That is also the reason for the current formulation using if then else. -/
213+
l ← e.mfilter (λ d,
214+
if e.is_prefix_of_file ml d.to_name ∧ ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e
215+
then bnot <$> has_attribute' `sanity_skip d.to_name else return ff),
191216
let ml' := ml.length,
192-
let s : format := "/- Note: This command is still in development. -/\n",
193-
let s := s ++ "/- Checking " ++ l.length ++ " declarations in mathlib (only in imported files) -/\n\n",
194-
f ← print_decls_sorted_mathlib ml' <$> fold_over_with_cond_sorted l unused_arguments,
195-
let s := s ++ "/- UNUSED ARGUMENTS: -/" ++ f ++ "\n\n",
196-
f ← print_decls_sorted_mathlib ml' <$> fold_over_with_cond_sorted l incorrect_def_lemma,
197-
let s := s ++ "/- INCORRECT DEF/LEMMA: -/" ++ f ++ "\n\n",
198-
f ← print_decls_sorted_mathlib ml' <$> fold_over_with_cond_sorted l dup_namespace,
199-
let s := s ++ "/- DUPLICATED NAMESPACES IN NAME: -/" ++ f ++ "\n\n",
200-
return s
217+
sanity_check_aux l (λ t, print_decls_sorted_mathlib ml' <$> fold_over_with_cond_sorted l t)
218+
"in mathlib (only in imported files)" slow checks
219+
220+
/-- Return the message printed by `#sanity_check_all`. -/
221+
meta def sanity_check_all (slow : bool := tt)
222+
(checks : list ((declaration → tactic (option string)) × string × string) :=
223+
default_checks slow) : tactic format := do
224+
e ← get_env,
225+
l ← e.mfilter (λ d, if ¬ d.to_name.is_internal ∧ ¬ d.is_auto_generated e
226+
then bnot <$> has_attribute' `sanity_skip d.to_name else return ff),
227+
sanity_check_aux l (λ t, print_decls_sorted <$> fold_over_with_cond_sorted l t)
228+
"in all imported files (including this one)" slow checks
201229

202230
/-- The command `#sanity_check` at the bottom of a file will warn you about some common mistakes
203231
in that file. -/
204232
@[user_command] meta def sanity_check_cmd (_ : parse $ tk "#sanity_check") : parser unit :=
205-
do s ← sanity_check, trace s
233+
do b ← optional (tk "-"), s ← sanity_check b.is_none, trace s
206234

207235
/-- The command `#sanity_check_mathlib` checks all of mathlib for certain mistakes. -/
208236
@[user_command] meta def sanity_check_mathlib_cmd (_ : parse $ tk "#sanity_check_mathlib") :
209237
parser unit :=
210-
do s ← sanity_check_mathlib, trace s
238+
do b ← optional (tk "-"), s ← sanity_check_mathlib b.is_none, trace s
239+
240+
/-- The command `#sanity_check_mathlib` checks all of mathlib for certain mistakes. -/
241+
@[user_command] meta def sanity_check_all_cmd (_ : parse $ tk "#sanity_check_all") :
242+
parser unit :=
243+
do b ← optional (tk "-"), s ← sanity_check_all b.is_none, trace s
211244

245+
/-- Use sanity check as a hole command. Note: In a large file, there might be some delay between
246+
choosing the option and the information appearing -/
212247
@[hole_command] meta def sanity_check_hole_cmd : hole_command :=
213248
{ name := "Sanity Check",
214-
descr := "Sanity check: Find mistakes in current file.",
249+
descr := "Sanity check: Find common mistakes in current file.",
215250
action := λ es, do s ← sanity_check, return [(s.to_string,"")] }
216251

217252
-- set_option profiler true
218253
-- run_cmd sanity_check
219254
-- run_cmd sanity_check_mathlib
255+
-- run_cmd sanity_check_all
220256
-- #sanity_check
221257
-- #sanity_check_mathlib
258+
-- #sanity_check_all
259+
-- #sanity_check-
260+
-- #sanity_check_mathlib-
261+
-- #sanity_check_all-

0 commit comments

Comments
 (0)