Skip to content

Commit

Permalink
Add more tests (#57)
Browse files Browse the repository at this point in the history
  • Loading branch information
eyihluyc authored Dec 7, 2024
1 parent 1e67968 commit 1e861f0
Showing 1 changed file with 227 additions and 31 deletions.
258 changes: 227 additions & 31 deletions Minimal/Examples.lean
Original file line number Diff line number Diff line change
Expand Up @@ -9,14 +9,37 @@ open Reduce
This file contains examples of minimal φ-calculus terms and their reductions.
Summary:
1. R_dot, locator substitution:
⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.x ⇝* ⟦ ⟧
2. R_app, R_dot, cong_DOT:
⟦ x ↦ ∅ ⟧(x ↦ ⟦ ⟧).x ⇝* ⟦ ⟧
3. R_dot, locator substitution, cong_DOT:
⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝* ⟦ ⟧
4. R_phi:
⟦ φ ↦ ⟦ x ↦ ⟦ ⟧ ⟧ ⟧.x ⇝* ⟦ ⟧
5. cong_APP, R_app, locator increment
⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ρ⁰.y) ⟧ ⇝* ⟦ x ↦ ⟦ a ↦ ρ¹.y ⟧ ⟧
6. Infinite reduction sequence presented after Definition 3.1 [KS22]
⟦ x ↦ ρ⁰.y, y ↦ ρ⁰.x ⟧.x ⇝ ⟦ x ↦ ρ⁰.y, y ↦ ρ⁰.x ⟧.y
⟦ x ↦ ρ⁰.y, y ↦ ρ⁰.x ⟧.y ⇝ ⟦ x ↦ ρ⁰.y, y ↦ ρ⁰.x ⟧.x
7. Examples/tests of different reductions of the same term presented after Definition 3.1 [KS22]
⟦ x ↦ ⟦ y ↦ ∅ ⟧ ⟧.x(y ↦ ⟦ z ↦ w ⟧.z)
8. Locator decrement in substitution, cong_OBJ, R_dot
⟦ b ↦ ⟦ c ↦ ρ¹ ⟧.c ⟧ ⇝* ⟦ b ↦ ρ⁰ ⟧
9. R_dot, R_app, cong_OBJ, congAPPₗ
⟦ a ↦ ⟦ b ↦ ∅ ⟧ ⟧.a (b ↦ ⟦ ⟧) ⇝* ⟦ b ↦ ⟦ ⟧ ⟧
10. congAPPᵣ, R_dot
⟦ ⟧ (a ↦ ⟦ b ↦ ⟦ ⟧ ⟧.b) ⇝* ⟦ ⟧ (a ↦ ⟦ ⟧)
## References
* [Nikolai Kudasov and Violetta Sim. 2023. Formalizing 𝜑-Calculus: A Purely Object-Oriented Calculus of Decorated Objects.][KS22]
-/

-- TEST 1: ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.x ⇝* ⟦ ⟧
-- TEST 1: ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.x ⇝* ⟦ ⟧

-- bindings: x ↦ ξ.y, y ↦ ⟦ ⟧
-- bindings: x ↦ ρ⁰.y, y ↦ ⟦ ⟧
def bindings1 : Bindings ["x", "y"] :=
Bindings.cons
"x"
Expand All @@ -25,19 +48,19 @@ def bindings1 : Bindings ["x", "y"] :=
(Bindings.cons
"y"
(by simp)
(attached (obj (Bindings.nil)))
(attached (obj Bindings.nil))
Bindings.nil)

-- ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.x
-- ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.x
def example1 : Term := dot (obj bindings1) "x"

-- ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.y
-- ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.y
def example1_a : Term := dot (obj bindings1) "y"

-- ⟦ ⟧
def example1_res : Term := obj Bindings.nil

-- ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.x ⇝ ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.y
-- ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.x ⇝ ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.y
def test1_red1 : example1 ⇝ example1_a := by
rw [example1, example1_a]
have := (@dot_c
Expand All @@ -50,7 +73,7 @@ def test1_red1 : example1 ⇝ example1_a := by
(by simp [lookup]))
simp at this; exact this

-- ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.y ⇝ ⟦ ⟧
-- ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.y ⇝ ⟦ ⟧
def test1_red2 : example1_a ⇝ example1_res := by
rw [example1_a, example1_res]
exact (dot_c
Expand All @@ -60,7 +83,7 @@ def test1_red2 : example1_a ⇝ example1_res := by
rfl
(by simp [lookup]))

-- ⟦ x ↦ ξ.y, y ↦ ⟦ ⟧ ⟧.x ⇝* ⟦ ⟧
-- ⟦ x ↦ ρ⁰.y, y ↦ ⟦ ⟧ ⟧.x ⇝* ⟦ ⟧
def test1 : example1 ⇝* example1_res :=
ReflTransGen.head
test1_red1
Expand Down Expand Up @@ -127,7 +150,7 @@ def test2_red2 : example2_a ⇝ example2_res := by
(Bindings.cons
"x"
(by simp)
(attached (obj (Bindings.nil)))
(attached (obj Bindings.nil))
Bindings.nil)
(by simp)
(by simp [lookup]))
Expand All @@ -141,9 +164,9 @@ def test2 : example2 ⇝* example2_res :=
ReflTransGen.refl)

------------------------------------------
-- TEST 3: ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝* ⟦ ⟧
-- TEST 3: ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝* ⟦ ⟧

-- ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z
-- ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z
def example3 : Term :=
dot
(dot
Expand All @@ -159,7 +182,7 @@ def example3 : Term :=
"x")
"z"

-- ⟦ z ↦ ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.y ⟧.z
-- ⟦ z ↦ ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.y ⟧.z
def example3_a : Term :=
dot
(obj
Expand All @@ -180,7 +203,7 @@ def example3_a : Term :=
Bindings.nil))
"z"

-- ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.y
-- ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.y
def example3_b : Term :=
dot
(obj
Expand All @@ -198,7 +221,7 @@ def example3_b : Term :=
-- ⟦ ⟧
def example3_res : Term := obj Bindings.nil

-- ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝ ⟦ z ↦ ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.y ⟧.z
-- ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝ ⟦ z ↦ ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.y ⟧.z
def test3_red1 : example3 ⇝ example3_a := congDOT
(dot
(obj
Expand Down Expand Up @@ -326,7 +349,7 @@ def test3_red3 : example3_b ⇝ example3_res := by
(by simp)
(by simp [lookup]))

-- ⟦ x ↦ ⟦ z ↦ ρ.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝* ⟦ ⟧
-- ⟦ x ↦ ⟦ z ↦ ρ¹.y ⟧, y ↦ ⟦ ⟧ ⟧.x.z ⇝* ⟦ ⟧
def test3 : example3 ⇝* example3_res :=
ReflTransGen.head
test3_red1
Expand Down Expand Up @@ -465,9 +488,9 @@ def test4 : example4 ⇝* example4_res :=
ReflTransGen.refl)))

------------------------------------------
-- TEST 5: ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ξ.y) ⟧ ⇝* ⟦ x ↦ ⟦ a ↦ ρ.y ⟧ ⟧
-- TEST 5: ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ρ⁰.y) ⟧ ⇝* ⟦ x ↦ ⟦ a ↦ ρ¹.y ⟧ ⟧

-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ξ.y) ⟧
-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ρ⁰.y) ⟧
def example5 : Term :=
(obj
(Bindings.cons "x" (by simp)
Expand All @@ -478,7 +501,7 @@ def example5 : Term :=
(dot (loc 0) "y")))
Bindings.nil))

-- ⟦ x ↦ ⟦ a ↦ ρ.y ⟧ ⟧
-- ⟦ x ↦ ⟦ a ↦ ρ¹.y ⟧ ⟧
def example5_res : Term :=
(obj
(Bindings.cons "x" (by simp)
Expand All @@ -488,19 +511,11 @@ def example5_res : Term :=
Bindings.nil)))
Bindings.nil))

-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ξ.y) ⟧ ⇝ ⟦ x ↦ ⟦ a ↦ ρ.y ⟧ ⟧
-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ρ⁰.y) ⟧ ⇝ ⟦ x ↦ ⟦ a ↦ ρ¹.y ⟧ ⟧
def test5_red : example5 ⇝ example5_res := by
simp [example5, example5_res]
exact (@congOBJ
(app
(obj (Bindings.cons "a" (by simp) void Bindings.nil))
"a"
(dot (loc 0) "y"))
(obj (Bindings.cons "a" (by simp)
(attached (dot (loc 1) "y"))
Bindings.nil))
exact (congOBJ
"x"
["x"]
(Bindings.cons "x" (by simp)
(attached
(app
Expand All @@ -509,18 +524,17 @@ def test5_red : example5 ⇝ example5_res := by
(dot (loc 0) "y")))
Bindings.nil)
(by
have := @app_c
have := app_c
(obj (Bindings.cons "a" (by simp) void Bindings.nil))
(dot (loc 0) "y")
"a"
["a"]
(Bindings.cons "a" (by simp) void Bindings.nil)
(by simp)
(by simp [lookup])
simp at this; exact this)
(by apply IsAttached.zeroth_attached))

-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ξ.y) ⟧ ⇝* ⟦ x ↦ ⟦ a ↦ ρ.y ⟧ ⟧
-- ⟦ x ↦ ⟦ a ↦ ∅ ⟧ (a ↦ ρ⁰.y) ⟧ ⇝* ⟦ x ↦ ⟦ a ↦ ρ¹.y ⟧ ⟧
def test5 : example5 ⇝* example5_res :=
ReflTransGen.head
test5_red
Expand Down Expand Up @@ -694,3 +708,185 @@ def test_graph_7 : example_graph_r2 ⇝ example_graph_last
(by simp [lookup])
simp [insert_φ] at reduction
exact reduction

------------------------------------------
-- TEST 8: ⟦ b ↦ ⟦ c ↦ ρ¹ ⟧.c ⟧ ⇝* ⟦ b ↦ ρ⁰ ⟧

-- ⟦ b ↦ ⟦ c ↦ ρ¹ ⟧.c ⟧
def example8 : Term :=
(obj
(Bindings.cons "b" (by simp)
(attached
(dot
(obj (Bindings.cons "c" (by simp)
(attached (loc 1))
Bindings.nil))
"c"))
Bindings.nil))

-- ⟦ b ↦ ρ⁰ ⟧
def example8_res : Term :=
(obj
(Bindings.cons "b" (by simp)
(attached (loc 0))
Bindings.nil))

-- ⟦ b ↦ ⟦ c ↦ ρ¹ ⟧.c ⟧ ⇝ ⟦ b ↦ ρ⁰ ⟧
def test8_red : example8 ⇝ example8_res := by
simp [example8, example8_res]
have := @congOBJ
(dot
(obj (Bindings.cons "c" (by simp)
(attached (loc 1))
Bindings.nil))
"c")
(loc 0)
"b"
["b"]
(Bindings.cons "b" (by simp)
(attached
(dot
(obj (Bindings.cons "c" (by simp)
(attached (loc 1))
Bindings.nil))
"c"))
Bindings.nil)
(by
have := @dot_c
(obj (Bindings.cons "c" (by simp)
(attached (loc 1))
Bindings.nil))
(loc 1)
"c"
["c"]
(Bindings.cons "c" (by simp)
(attached (loc 1))
Bindings.nil)
(by simp)
(by simp [lookup])
simp at this; exact this)
(by apply IsAttached.zeroth_attached)
simp [insert_φ] at this; exact this

-- ⟦ b ↦ ⟦ c ↦ ρ¹ ⟧.c ⟧ ⇝* ⟦ b ↦ ρ⁰ ⟧
def test8 : example8 ⇝* example8_res :=
ReflTransGen.head
test8_red
ReflTransGen.refl

------------------------------------------
-- TEST 9: ⟦ a ↦ ⟦ b ↦ ∅ ⟧ ⟧.a (b ↦ ⟦ ⟧) ⇝* ⟦ b ↦ ⟦ ⟧ ⟧

-- ⟦ a ↦ ⟦ b ↦ ∅ ⟧ ⟧.a (b ↦ ⟦ ⟧)
def example9 : Term :=
app
(dot
(obj
(Bindings.cons "a" (by simp)
(attached (obj
(Bindings.cons "b" (by simp) void Bindings.nil)))
Bindings.nil))
"a")
"b"
(obj Bindings.nil)

-- ⟦ b ↦ ∅ ⟧ (b ↦ ⟦ ⟧)
def example9_a : Term :=
app
(obj (Bindings.cons "b" (by simp) void Bindings.nil))
"b"
(obj Bindings.nil)

-- ⟦ b ↦ ⟦ ⟧ ⟧
def example9_res : Term :=
(obj
(Bindings.cons "b" (by simp)
(attached (obj Bindings.nil))
Bindings.nil))

-- ⟦ a ↦ ⟦ b ↦ ∅ ⟧ ⟧.a (b ↦ ⟦ ⟧) ⇝ ⟦ b ↦ ∅ ⟧ (b ↦ ⟦ ⟧)
def test9_red1 : example9 ⇝ example9_a := by
simp [example9, example9_a]
exact (congAPPₗ _ _ _ _
(by
have := @dot_c
(obj
(Bindings.cons "a" (by simp)
(attached (obj
(Bindings.cons "b" (by simp) void Bindings.nil)))
Bindings.nil))
(obj (Bindings.cons "b" (by simp) void Bindings.nil))
"a"
["a"]
(Bindings.cons "a" (by simp)
(attached (obj
(Bindings.cons "b" (by simp) void Bindings.nil)))
Bindings.nil)
(by simp)
(by simp [lookup])
simp at this; exact this))

-- ⟦ b ↦ ∅ ⟧ (b ↦ ⟦ ⟧) ⇝ ⟦ b ↦ ⟦ ⟧ ⟧
def test9_red2 : example9_a ⇝ example9_res := by
simp [example9_a, example9_res]
exact (by
have := app_c
(obj (Bindings.cons "b" (by simp) void Bindings.nil))
(obj Bindings.nil)
"b"
(Bindings.cons "b" (by simp) void Bindings.nil)
(by simp)
(by simp [lookup])
simp at this; exact this)

-- ⟦ b ↦ ∅ ⟧ (b ↦ ⟦ ⟧) ⇝* ⟦ b ↦ ⟦ ⟧ ⟧
def test9 : example9 ⇝* example9_res :=
ReflTransGen.head test9_red1
(ReflTransGen.head test9_red2
ReflTransGen.refl)

------------------------------------------
-- TEST 10: ⟦ ⟧ (a ↦ ⟦ b ↦ ⟦ ⟧ ⟧.b) ⇝* ⟦ ⟧ (a ↦ ⟦ ⟧)

-- ⟦ ⟧ (a ↦ ⟦ b ↦ ⟦ ⟧ ⟧.b)
def example10 : Term :=
app
(obj Bindings.nil)
"a"
(dot
(obj (Bindings.cons "b" (by simp)
(attached (obj Bindings.nil))
Bindings.nil))
"b")

-- ⟦ ⟧ (a ↦ ⟦ ⟧)
def example10_res : Term :=
app
(obj Bindings.nil)
"a"
(obj Bindings.nil)

-- ⟦ ⟧ (a ↦ ⟦ b ↦ ⟦ ⟧ ⟧.b) ⇝ ⟦ ⟧ (a ↦ ⟦ ⟧)
def test10_red : example10 ⇝ example10_res := by
simp [example10, example10_res]
exact (congAPPᵣ _ _ _ _
(by
have := @dot_c
(obj
(Bindings.cons "b" (by simp)
(attached (obj Bindings.nil))
Bindings.nil))
(obj Bindings.nil)
"b"
["b"]
(Bindings.cons "b" (by simp)
(attached (obj Bindings.nil))
Bindings.nil)
(by simp)
(by simp [lookup])
simp at this; exact this))

-- ⟦ ⟧ (a ↦ ⟦ b ↦ ⟦ ⟧ ⟧.b) ⇝* ⟦ ⟧ (a ↦ ⟦ ⟧)
def test10 : example10 ⇝* example10_res :=
ReflTransGen.head test10_red
ReflTransGen.refl

1 comment on commit 1e861f0

@0pdd
Copy link

@0pdd 0pdd commented on 1e861f0 Dec 7, 2024

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I wasn't able to retrieve PDD puzzles from the code base and submit them to github. If you think that it's a bug on our side, please submit it to yegor256/0pdd:

set -x && set -e && set -o pipefail && cd /tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA && pdd -v -f /tmp/20241207-4475-y308k1 [1]: + set -e + set -o pipefail + cd /tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA + pdd -v -f...

Please, copy and paste this stack trace to GitHub:

UserError
set -x && set -e && set -o pipefail && cd /tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA && pdd -v -f /tmp/20241207-4475-y308k1 [1]:
+ set -e
+ set -o pipefail
+ cd /tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA
+ pdd -v -f /tmp/20241207-4475-y308k1

My version is 0.24.0
Ruby version is 3.1.4 at x86_64-linux
Reading from root dir /tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA
/tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA/lake-manifest.json is a binary file (4329 bytes)
/tmp/0pdd20241207-2-3w2sdt/Z2l0QGdpdGh1Yi5jb206b2JqZWN0aW9uYXJ5L3Byb29mLmdpdA/renovate.json is a binary file (114 bytes)
Reading .github/workflows/lean.yml ...
Reading .gitignore ...
Reading Extended/ARS.lean ...
Reading Extended/Confluence/CompleteDevelopment.lean ...
Reading Extended/Confluence/Equivalence.lean ...
Reading Extended/Record.lean ...
Reading Extended/Reduction/Parallel/Definition.lean ...
Reading Extended/Reduction/Parallel/Properties.lean ...
Reading Extended/Reduction/Regular/Definition.lean ...
ERROR: ERROR: Extended/Reduction/Regular/Definition.lean; PDD::Error at Extended/Reduction/Regular/Definition.lean:102: TODO found, but puzzle can't be parsed, most probably because TODO is not followed by a puzzle marker, as this page explains: https://github.com/cqfn/pdd#how-to-format
If you can't understand the cause of this issue or you don't know how to fix it, please submit a GitHub issue, we will try to help you: https://github.com/cqfn/pdd/issues. This tool is still in its beta version and we will appreciate your feedback. Here is where you can find more documentation: https://github.com/cqfn/pdd/blob/master/README.md.
Exit code is 1

/app/objects/git_repo.rb:74:in `rescue in block in xml'
/app/objects/git_repo.rb:71:in `block in xml'
/app/vendor/ruby-3.1.4/lib/ruby/3.1.0/tempfile.rb:317:in `open'
/app/objects/git_repo.rb:70:in `xml'
/app/objects/puzzles.rb:46:in `deploy'
/app/objects/jobs/job.rb:38:in `proceed'
/app/objects/jobs/job_starred.rb:32:in `proceed'
/app/objects/jobs/job_recorded.rb:31:in `proceed'
/app/objects/jobs/job_emailed.rb:33:in `proceed'
/app/objects/jobs/job_commiterrors.rb:33:in `proceed'
/app/objects/jobs/job_detached.rb:48:in `exclusive'
/app/objects/jobs/job_detached.rb:36:in `block in proceed'
/app/objects/jobs/job_detached.rb:36:in `fork'
/app/objects/jobs/job_detached.rb:36:in `proceed'
/app/0pdd.rb:549:in `process_request'
/app/0pdd.rb:380:in `block in <top (required)>'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1804:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1804:in `block in compile!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1071:in `block (3 levels) in route!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1089:in `route_eval'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1071:in `block (2 levels) in route!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1120:in `block in process_route'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1118:in `catch'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1118:in `process_route'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1069:in `block in route!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1066:in `each'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1066:in `route!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1190:in `block in dispatch!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1161:in `catch'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1161:in `invoke'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1185:in `dispatch!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1001:in `block in call!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1161:in `catch'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1161:in `invoke'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1001:in `call!'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:990:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/rewindable_input.rb:25:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/deflater.rb:47:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/xss_header.rb:20:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/path_traversal.rb:18:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/json_csrf.rb:28:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/base.rb:53:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/base.rb:53:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-protection-4.0.0/lib/rack/protection/frame_options.rb:33:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/logger.rb:19:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/common_logger.rb:43:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:266:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:259:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/head.rb:15:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rack-3.0.9/lib/rack/method_override.rb:28:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:224:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:2115:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1674:in `block in call'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1890:in `synchronize'
/app/vendor/bundle/ruby/3.1.0/gems/sinatra-4.0.0/lib/sinatra/base.rb:1674:in `call'
/app/vendor/bundle/ruby/3.1.0/gems/rackup-2.1.0/lib/rackup/handler/webrick.rb:111:in `service'
/app/vendor/bundle/ruby/3.1.0/gems/webrick-1.8.1/lib/webrick/httpserver.rb:140:in `service'
/app/vendor/bundle/ruby/3.1.0/gems/webrick-1.8.1/lib/webrick/httpserver.rb:96:in `run'
/app/vendor/bundle/ruby/3.1.0/gems/webrick-1.8.1/lib/webrick/server.rb:310:in `block in start_thread'

Please sign in to comment.