-
Notifications
You must be signed in to change notification settings - Fork 463
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
feat: use reserved name infrastructure for functional induction (#3776)
no need to enter `derive_functional_induction` anymore. (Will remove the support for `derive_functional_induction` after the next stage0 update, since we are already using it in Init.)
- Loading branch information
Showing
11 changed files
with
127 additions
and
204 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,45 @@ | ||
-- Some of these tests made more sense when we had a | ||
-- derive_functional_induction command. | ||
|
||
#check doesNotExist.induct | ||
|
||
def takeWhile (p : α → Bool) (as : Array α) : Array α := | ||
foo 0 #[] | ||
where | ||
foo (i : Nat) (r : Array α) : Array α := | ||
if h : i < as.size then | ||
let a := as.get ⟨i, h⟩ | ||
if p a then | ||
foo (i+1) (r.push a) | ||
else | ||
r | ||
else | ||
r | ||
termination_by as.size - i | ||
|
||
-- Checks the error message when the users tries to access the induct rule for the wrong function | ||
-- (before we used reserved names for this feature we did give a more helpful error message here) | ||
#check takeWhile.induct | ||
|
||
#check takeWhile.foo.induct | ||
|
||
|
||
-- this tests the error we get when trying to access the induct rules for | ||
-- a function that recurses over an inductive *predicate* (not yet supported) | ||
|
||
inductive Even : Nat → Prop where | ||
| zero : Even 0 | ||
| plus2 : Even n → Even (n + 2) | ||
|
||
def idEven : Even n → Even n | ||
| .zero => .zero | ||
| .plus2 p => .plus2 (idEven p) | ||
|
||
#check idEven.induct | ||
|
||
-- this tests the error we get when trying to access the induct rules for | ||
-- a function that recurses over `Acc` | ||
|
||
def idAcc : Acc p x → Acc p x | ||
| Acc.intro x f => Acc.intro x (fun y h => idAcc (f y h)) | ||
#check idAcc.induct |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,23 @@ | ||
funind_errors.lean:4:7-4:26: error: unknown identifier 'doesNotExist.induct' | ||
funind_errors.lean:22:7-22:23: error: invalid field notation, type is not of the form (C ...) where C is a constant | ||
takeWhile | ||
has type | ||
(?m → Bool) → Array ?m → Array ?m | ||
takeWhile.foo.induct.{u_1} {α : Type u_1} (p : α → Bool) (as : Array α) (motive : Nat → Array α → Prop) | ||
(case1 : | ||
∀ (i : Nat) (r : Array α) (h : i < as.size), | ||
let a := as.get ⟨i, h⟩; | ||
p a = true → motive (i + 1) (r.push a) → motive i r) | ||
(case2 : | ||
∀ (i : Nat) (r : Array α) (h : i < as.size), | ||
let a := as.get ⟨i, h⟩; | ||
¬p a = true → motive i r) | ||
(case3 : ∀ (i : Nat) (r : Array α), ¬i < as.size → motive i r) (i : Nat) (r : Array α) : motive i r | ||
funind_errors.lean:38:7-38:20: error: invalid field notation, type is not of the form (C ...) where C is a constant | ||
idEven | ||
has type | ||
Even ?m → Even ?m | ||
funind_errors.lean:45:7-45:19: error: invalid field notation, type is not of the form (C ...) where C is a constant | ||
idAcc | ||
has type | ||
Acc ?m ?m → Acc ?m ?m |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Oops, something went wrong.