Skip to content

Commit 971a2d7

Browse files
committed
message assertion
1 parent 924b104 commit 971a2d7

File tree

3 files changed

+76
-46
lines changed

3 files changed

+76
-46
lines changed

lean/main/05_syntax.lean

Lines changed: 39 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -18,18 +18,18 @@ import Lean
1818
-- XOR, denoted \oplus
1919
infixl:60 " ⊕ " => fun l r => (!l && r) || (l && !r)
2020

21-
#eval truetrue -- false
22-
#eval truefalse -- true
23-
#eval false true -- true
24-
#eval falsefalse -- false
21+
#guard !(truetrue)
22+
#guard truefalse
23+
#guard falsetrue
24+
#guard !(falsefalse)
2525

2626
-- with `notation`, "left XOR"
2727
notation:10 l:10 " LXOR " r:11 => (!l && r)
2828

29-
#eval true LXOR true -- false
30-
#eval true LXOR false -- false
31-
#eval false LXOR true -- true
32-
#eval false LXOR false -- false
29+
#guard !(true LXOR true)
30+
#guard !(true LXOR false)
31+
#guard false LXOR true
32+
#guard !(false LXOR false)
3333

3434
/- As we can see the `infixl` command allows us to declare a notation for
3535
a binary operation that is infix, meaning that the operator is in between
@@ -52,9 +52,9 @@ The two unintuitive parts about these two are:
5252
precedence, meaning how strong they bind to their arguments, let's see this in action:
5353
-/
5454

55-
#eval truefalse LXOR false -- false
56-
#eval (truefalse) LXOR false -- false
57-
#eval true ⊕ (false LXOR false) -- true
55+
#guard !(truefalse LXOR false)
56+
#guard !((truefalse) LXOR false)
57+
#guard true ⊕ (false LXOR false)
5858

5959
/-!
6060
As we can see, the Lean interpreter analyzed the first term without parentheses
@@ -154,9 +154,12 @@ We can now write `MyTerm` in place of things like `1 + 1` and it will be
154154
it just means that the Lean parser can understand it:
155155
-/
156156

157+
/-⋆-//--
158+
info: elaboration function for 'termMyTerm' has not been implemented
159+
MyTerm
160+
-/
161+
#guard_msgs in --#
157162
#check_failure MyTerm
158-
-- elaboration function for 'termMyTerm' has not been implemented
159-
-- MyTerm
160163

161164
/-! Note: `#check_failure` command allows incorrectly typed terms to be indicated without error. -/
162165

@@ -276,13 +279,20 @@ syntax binNumber := binDigit,+
276279
/-!
277280
Since we can just use named parsers in place of syntax categories, we can now easily
278281
add this to the `term` category:
282+
-/
283+
284+
open Lean Parser in
285+
/-- function to check if parsing passes -/
286+
def parse (cat : Name) (s : String) : MetaM Syntax := do
287+
ofExcept <| runParserCategory (← getEnv) cat s
279288

280-
```lean
281289
syntax "bin(" binNumber ")" : term
282-
#check bin(Z, O, Z, Z, O) -- elaboration function hasn't been implemented but parsing passes
283-
#check bin() -- fails to parse because `binNumber` is "one or many": expected 'O' or 'Z'
284-
```
285-
-/
290+
#check_failure bin(Z, O, Z, Z, O) -- elaboration function hasn't been implemented but parsing passes
291+
292+
-- fails to parse `bin()` because `binNumber` is "one or many": expected 'O' or 'Z'
293+
/-⋆-//-- error: <input>:1:4: expected 'O' or 'Z' -/
294+
#guard_msgs in --#
295+
#eval parse `term "bin()"
286296

287297
syntax binNumber' := binDigit,* -- note the *
288298
syntax "emptyBin(" binNumber' ")" : term
@@ -390,8 +400,8 @@ def isAdd11 : Syntax → Bool
390400
| `(Nat.add 1 1) => true
391401
| _ => false
392402

393-
#eval isAdd11 (Syntax.mkApp (mkIdent `Nat.add) #[Syntax.mkNumLit "1", Syntax.mkNumLit "1"]) -- true
394-
#eval isAdd11 (Syntax.mkApp (mkIdent `Nat.add) #[mkIdent `foo, Syntax.mkNumLit "1"]) -- false
403+
#guard isAdd11 (Syntax.mkApp (mkIdent `Nat.add) #[Syntax.mkNumLit "1", Syntax.mkNumLit "1"]) -- true
404+
#guard ! isAdd11 (Syntax.mkApp (mkIdent `Nat.add) #[mkIdent `foo, Syntax.mkNumLit "1"]) -- false
395405

396406
/-!
397407
The next level with matches is to capture variables from the input instead
@@ -426,8 +436,8 @@ def isLitAdd : TSyntax `term → Option Nat
426436
| `(Nat.add $x:num $y:num) => some (x.getNat + y.getNat)
427437
| _ => none
428438

429-
#eval isLitAdd (Syntax.mkApp (mkIdent `Nat.add) #[Syntax.mkNumLit "1", Syntax.mkNumLit "1"]) -- some 2
430-
#eval isLitAdd (Syntax.mkApp (mkIdent `Nat.add) #[mkIdent `foo, Syntax.mkNumLit "1"]) -- none
439+
#guard isLitAdd (Syntax.mkApp (mkIdent `Nat.add) #[Syntax.mkNumLit "1", Syntax.mkNumLit "1"]) = some 2
440+
#guard isLitAdd (Syntax.mkApp (mkIdent `Nat.add) #[mkIdent `foo, Syntax.mkNumLit "1"]) = none
431441

432442
/-!
433443
If you want to access the `Syntax` behind a `TSyntax` you can do this using
@@ -448,6 +458,7 @@ chapters.
448458
declare_syntax_cat arith
449459

450460
syntax num : arith
461+
451462
syntax arith "-" arith : arith
452463
syntax arith "+" arith : arith
453464
syntax "(" arith ")" : arith
@@ -466,7 +477,9 @@ def test : Elab.TermElabM Nat := do
466477
let stx ← `(arith| (12 + 3) - 4)
467478
pure (denoteArith stx)
468479

469-
#eval test -- 11
480+
/-⋆-//-- info: 11 -/
481+
#guard_msgs in --#
482+
#eval test
470483

471484
/-!
472485
Feel free to play around with this example and extend it in whatever way
@@ -547,7 +560,9 @@ basic version of our notation:
547560
-/
548561
notation "{ " x " | " p " }" => setOf (fun x => p)
549562

550-
#check { (x : Nat) | x ≤ 1 } -- { x | x ≤ 1 } : Set Nat
563+
/-⋆-//-- info: { x | x ≤ 1 } : Set Nat -/
564+
#guard_msgs in --#
565+
#check { (x : Nat) | x ≤ 1 }
551566

552567
example : 1 ∈ { (y : Nat) | y ≤ 1 } := by simp[Membership.mem, Set.mem, setOf]
553568
example : 2 ∈ { (y : Nat) | y ≤ 31 ≤ y } := by simp[Membership.mem, Set.mem, setOf]

lean/main/06_macros.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -154,7 +154,9 @@ compiler, you could also declare functions that are specific to certain
154154
`TSyntax` variants. For example as we have seen in the syntax chapter
155155
there exists the function:
156156
-/
157-
#check TSyntax.getNat -- TSyntax.getNat : TSyntax numLitKind → Nat
157+
/-⋆-//-- info: Lean.TSyntax.getNat (s : NumLit) : Nat -/
158+
#guard_msgs in
159+
#check TSyntax.getNat
158160
/-!
159161
Which is guaranteed to not panic because we know that the `Syntax` that
160162
the function is receiving is a numeric literal and can thus naturally

lean/main/09_tactics.lean

Lines changed: 34 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -39,11 +39,10 @@ We start by simply declaring the tactic with no implementation:
3939

4040
syntax "custom_tactic" : tactic
4141

42-
/-- error: tactic 'tacticCustom_tactic' has not been implemented -/
42+
/-⋆-//-- error: tactic 'tacticCustom_tactic' has not been implemented -/
4343
#guard_msgs in --#
4444
example : 42 = 42 := by
4545
custom_tactic
46-
sorry
4746

4847
/-
4948
We will now add the `rfl` tactic into `custom_tactic`, which will allow us to
@@ -54,20 +53,23 @@ macro_rules
5453
| `(tactic| custom_tactic) => `(tactic| rfl)
5554

5655
example : 42 = 42 := by
57-
custom_tactic
58-
-- Goals accomplished 🎉
56+
custom_tactic
57+
-- Goals accomplished 🎉
5958

6059
/-
6160
We can now try a harder problem, that cannot be immediately dispatched by `rfl`:
6261
-/
6362

64-
#check_failure (by custom_tactic : 43 = 4342 = 42)
65-
-- type mismatch
66-
-- Iff.rfl
67-
-- has type
68-
-- ?m.1437 ↔ ?m.1437 : Prop
69-
-- but is expected to have type
70-
-- 43 = 43 ∧ 42 = 42 : Prop
63+
/-⋆-//--
64+
error: tactic 'rfl' failed, the left-hand side
65+
43 = 43
66+
is not definitionally equal to the right-hand side
67+
42 = 42
68+
⊢ 43 = 43 ∧ 42 = 42
69+
-/
70+
#guard_msgs in --#
71+
example : 43 = 4342 = 42 := by
72+
custom_tactic
7173

7274
/-
7375
We extend the `custom_tactic` tactic with a tactic that tries to break `And`
@@ -89,7 +91,7 @@ that we dispatch the theorem.
8991

9092
example : 43 = 4342 = 42 := by
9193
custom_tactic
92-
-- Goals accomplished 🎉
94+
-- Goals accomplished 🎉
9395

9496
/-
9597
In summary, we declared an extensible tactic called `custom_tactic`. It
@@ -120,9 +122,12 @@ macro_rules
120122
theorem test_and_then: 1 = 12 = 2 := by
121123
apply And.intro and_then rfl
122124

125+
/-⋆-//--
126+
info: theorem test_and_then : 1 = 1 ∧ 2 = 2 :=
127+
⟨Eq.refl 1, Eq.refl 2⟩
128+
-/
129+
#guard_msgs in --#
123130
#print test_and_then
124-
-- theorem test_and_then : 1 = 1 ∧ 2 = 2 :=
125-
-- { left := Eq.refl 1, right := Eq.refl 2 }
126131

127132
/-
128133
## Exploring `TacticM`
@@ -394,10 +399,14 @@ elab "custom_assump_2" : tactic =>
394399
example (H1 : 1 = 1) (H2 : 2 = 2) : 2 = 2 := by
395400
custom_assump_2
396401

397-
#check_failure (by custom_assump_2 : (H1 : 1 = 1) → 2 = 2)
398-
-- tactic 'custom_assump_2' failed, unable to find matching hypothesis of type (2 = 2)
399-
-- H1 : 1 = 1
400-
-- ⊢ 2 = 2
402+
/-⋆-//--
403+
error: tactic 'custom_assump_2' failed, unable to find matching hypothesis of type (2 = 2)
404+
H1 : 1 = 1
405+
⊢ 2 = 2
406+
-/
407+
#guard_msgs in --#
408+
example (H1 : 1 = 1) : 2 = 2 := by
409+
custom_assump_2
401410

402411
/-
403412
### Tweaking the context
@@ -578,9 +587,13 @@ elab "faq_throw_error" : tactic =>
578587
let goal ← Lean.Elab.Tactic.getMainGoal
579588
Lean.Meta.throwTacticEx `faq_throw_error goal "throwing an error at the current goal"
580589

581-
#check_failure (by faq_throw_error : (b : Bool) → b = true)
582-
-- tactic 'faq_throw_error' failed, throwing an error at the current goal
583-
-- ⊢ ∀ (b : Bool), b = true
590+
/-⋆-//--
591+
error: tactic 'faq_throw_error' failed, throwing an error at the current goal
592+
⊢ ∀ (b : Bool), b = true
593+
-/
594+
#guard_msgs in --#
595+
example : (b : Bool) → b = true := by
596+
faq_throw_error
584597

585598
/-!
586599
**Q: What is the difference between `Lean.Elab.Tactic.*` and `Lean.Meta.Tactic.*`?**

0 commit comments

Comments
 (0)