-
Notifications
You must be signed in to change notification settings - Fork 446
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
- Loading branch information
1 parent
a1c8e0d
commit 5d364cf
Showing
2 changed files
with
55 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,52 @@ | ||
import Lean | ||
|
||
/-! | ||
# Validation of `Json Empty` instance | ||
This test serves a similar purpose to the `Repr Empty` test in `repr_empty.lean`. | ||
-/ | ||
|
||
structure Foo (α : Type) where | ||
y : Option α | ||
deriving Lean.ToJson, Lean.FromJson, Repr | ||
|
||
def fromStrHelper (res : Type) [Lean.FromJson res] (s : String) : Except String res := | ||
(Lean.Json.parse s) >>= Lean.fromJson? | ||
|
||
/-! First, two basic examples with α not-Empty -/ | ||
/-- info: {"y": 1} -/ | ||
#guard_msgs in | ||
#eval Lean.toJson (⟨some 1⟩ : Foo Nat) | ||
|
||
/-! Ensure we can parse the type back -/ | ||
/-- info: Except.ok { y := some 1 } -/ | ||
#guard_msgs in | ||
#eval fromStrHelper (Foo Nat) "{\"y\": 1}" | ||
|
||
/-- info: {"y": null} -/ | ||
#guard_msgs in | ||
#eval Lean.toJson (⟨none⟩ : Foo Nat) | ||
|
||
/-- info: Except.ok { y := none } -/ | ||
#guard_msgs in | ||
#eval fromStrHelper (Foo Nat) "{\"y\": null}" | ||
|
||
/-! Examples with the `Empty` type -/ | ||
/-- info: {"y": null} -/ | ||
#guard_msgs in | ||
#eval Lean.toJson (⟨none⟩ : Foo Empty) | ||
|
||
/-! Show that we can round-trip from string -/ | ||
/-- info: Except.ok { y := none } -/ | ||
#guard_msgs in | ||
#eval fromStrHelper (Foo Empty) "{\"y\": null}" | ||
|
||
/-! Show that parsing fails -/ | ||
/-- info: Except.error "no constructor matched JSON value '\"Yo!\"'" -/ | ||
#guard_msgs in | ||
#eval fromStrHelper (Empty) "\"Yo!\"" | ||
|
||
/-! Show that parsing fails if we supply anything else but `null` -/ | ||
/-- info: Except.error "Foo.y: no constructor matched JSON value '1'" -/ | ||
#guard_msgs in | ||
#eval fromStrHelper (Foo Empty) "{\"y\": 1}" |