-
Notifications
You must be signed in to change notification settings - Fork 109
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
chore: upstream classical tactic (#242)
- Loading branch information
Showing
3 changed files
with
84 additions
and
0 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
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,51 @@ | ||
/- | ||
Copyright (c) 2021 Mario Carneiro. All rights reserved. | ||
Released under Apache 2.0 license as described in the file LICENSE. | ||
Authors: Mario Carneiro | ||
-/ | ||
import Lean.Elab.ElabRules | ||
|
||
/-! # `classical` and `classical!` tactics -/ | ||
|
||
namespace Std.Tactic | ||
open Lean Meta | ||
|
||
/-- | ||
`classical!` adds a proof of `Classical.propDecidable` as a local variable, which makes it | ||
available for instance search and effectively makes all propositions decidable. | ||
``` | ||
noncomputable def foo : Bool := by | ||
classical! | ||
have := ∀ p, decide p -- uses the classical instance | ||
exact decide (0 < 1) -- uses the classical instance even though `0 < 1` is decidable | ||
``` | ||
Consider using `classical` instead if you want to use the decidable instance when available. | ||
-/ | ||
macro (name := classical!) "classical!" : tactic => | ||
`(tactic| have em := Classical.propDecidable) | ||
|
||
/-- | ||
`classical tacs` runs `tacs` in a scope where `Classical.propDecidable` is a low priority | ||
local instance. It differs from `classical!` in that `classical!` uses a local variable, | ||
which has high priority: | ||
``` | ||
noncomputable def foo : Bool := by | ||
classical! | ||
have := ∀ p, decide p -- uses the classical instance | ||
exact decide (0 < 1) -- uses the classical instance even though `0 < 1` is decidable | ||
def bar : Bool := by | ||
classical | ||
have := ∀ p, decide p -- uses the classical instance | ||
exact decide (0 < 1) -- uses the decidable instance | ||
``` | ||
Note that (unlike lean 3) `classical` is a scoping tactic - it adds the instance only within the | ||
scope of the tactic. | ||
-/ | ||
-- FIXME: using ppDedent looks good in the common case, but produces the incorrect result when | ||
-- the `classical` does not scope over the rest of the block. | ||
elab "classical" tacs:ppDedent(tacticSeq) : tactic => do | ||
modifyEnv Meta.instanceExtension.pushScope | ||
Meta.addInstance ``Classical.propDecidable .local 10 | ||
try Elab.Tactic.evalTactic tacs | ||
finally modifyEnv Meta.instanceExtension.popScope |
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,32 @@ | ||
import Std.Tactic.Classical | ||
import Std.Tactic.PermuteGoals | ||
import Std.Tactic.GuardExpr | ||
|
||
noncomputable example : Bool := by | ||
fail_if_success have := ∀ p, decide p -- no classical in scope | ||
classical! | ||
have := ∀ p, decide p -- uses the classical instance | ||
-- uses the classical instance even though `0 < 1` is decidable | ||
guard_expr decide (0 < 1) = @decide (0 < 1) (‹(a : Prop) → Decidable a› _) | ||
exact decide (0 < 1) | ||
|
||
example : Bool := by | ||
fail_if_success have := ∀ p, decide p -- no classical in scope | ||
classical | ||
have := ∀ p, decide p -- uses the classical instance | ||
guard_expr decide (0 < 1) = @decide (0 < 1) (Nat.decLt 0 1) | ||
exact decide (0 < 1) -- uses the decidable instance | ||
|
||
-- double check no leakage | ||
example : Bool := by | ||
fail_if_success have := ∀ p, decide p -- no classical in scope | ||
exact decide (0 < 1) -- uses the decidable instance | ||
|
||
-- check that classical respects tactic blocks | ||
example : Bool := by | ||
fail_if_success have := ∀ p, decide p -- no classical in scope | ||
on_goal 1 => | ||
classical | ||
have := ∀ p, decide p -- uses the classical instance | ||
fail_if_success have := ∀ p, decide p -- no classical in scope again | ||
exact decide (0 < 1) -- uses the decidable instance |