116 lines
4.0 KiB
Agda
116 lines
4.0 KiB
Agda
open import Relation.Binary.PropositionalEquality
|
||
open import Data.Product
|
||
|
||
import Data.List as L
|
||
import Data.Vec as V
|
||
import Data.String as S
|
||
import Data.Char as C
|
||
open L using (List)
|
||
open V using (Vec)
|
||
open S using (String)
|
||
open C using (Char)
|
||
|
||
open import Base
|
||
open import Bits-and-Bytes
|
||
open import Parsing
|
||
import UTF-8
|
||
import ASCII
|
||
open import NonEmpty
|
||
|
||
open import HTTP
|
||
open import ParseHelpers
|
||
open import HTTP.URL
|
||
open import HTTP.Parse
|
||
|
||
module HTTP.Test where
|
||
|
||
enc : (s : String) → V.Vec Byte (L.length (UTF-8.encode-string s))
|
||
enc x = V.fromList (UTF-8.encode-string x)
|
||
|
||
module Test-HTTP where
|
||
test-GET : (parse-http-method .parse (enc "GET /mjau.html"))
|
||
≡ real ⟨ GET , _ , enc " /mjau.html" ⟩
|
||
test-GET = refl
|
||
|
||
test-DELETE : parse-http-method .parse (enc "DELETE haters.list")
|
||
≡ real ⟨ DELETE , _ , enc " haters.list" ⟩
|
||
test-DELETE = refl
|
||
|
||
test-PUT : parse-http-method .parse (enc "PUT /")
|
||
≡ real ⟨ PUT , _ , enc " /" ⟩
|
||
test-PUT = refl
|
||
|
||
test-POST : parse-http-method .parse (enc "POST")
|
||
≡ real ⟨ POST , _ , enc "" ⟩
|
||
test-POST = refl
|
||
|
||
test-CONNECT : parse-http-method .parse (enc "CONNECT")
|
||
≡ real ⟨ CONNECT , _ , enc "" ⟩
|
||
test-CONNECT = refl
|
||
|
||
test-OPTIONS : parse-http-method .parse (enc "OPTIONS")
|
||
≡ real ⟨ OPTIONS , _ , enc "" ⟩
|
||
test-OPTIONS = refl
|
||
|
||
test-TRACE : parse-http-method .parse (enc "TRACE")
|
||
≡ real ⟨ TRACE , _ , enc "" ⟩
|
||
test-TRACE = refl
|
||
|
||
test-PATCH : parse-http-method .parse (enc "PATCH")
|
||
≡ real ⟨ PATCH , _ , enc "" ⟩
|
||
test-PATCH = refl
|
||
|
||
test-scheme : parse-scheme .parse (enc "http://")
|
||
≡ real ⟨ "http" , _ , enc "://" ⟩
|
||
test-scheme = refl
|
||
|
||
test-auth-portless : parse-authority .parse (enc "coral.shoes/mjau.html")
|
||
≡ real ⟨ "coral.shoes" ꞉ 80 , _ , enc "/mjau.html" ⟩
|
||
test-auth-portless = refl
|
||
|
||
test-auth-portful : parse-authority .parse (enc "coral.shoes:80/mjau.html")
|
||
≡ real ⟨ "coral.shoes" ꞉ 80 , _ , enc "/mjau.html" ⟩
|
||
test-auth-portful = refl
|
||
|
||
test-auth-port-out-of-range : parse-authority .parse (enc "coral.shoes:13372137/mjau.html")
|
||
≡ cake
|
||
test-auth-port-out-of-range = refl
|
||
|
||
test-path : parse-path .parse (enc "/mjau/cat.html?meowing=true")
|
||
≡ real ⟨ "mjau" / "cat.html" / $ , _ , enc "?meowing=true" ⟩
|
||
test-path = refl
|
||
|
||
test-path-encoded : parse-path .parse (enc "/%68%65%77%77%6f/%3f%5e%2f%2f%5e%3f?mjau")
|
||
≡ real ⟨ "hewwo" / "?^//^?" / $ , _ , enc "?mjau" ⟩
|
||
test-path-encoded = refl
|
||
|
||
test-path-bare : parse-path .parse (enc "/")
|
||
≡ real ⟨ $ , _ , enc "" ⟩
|
||
test-path-bare = refl
|
||
|
||
test-url : parse-url .parse (enc "http://www.rfc-editor.org/rfc/rfc3986.html")
|
||
≡ real ⟨ http:// "www.rfc-editor.org" ꞉ 80 / ("rfc" / "rfc3986.html" / $) ¿ cake , _ , enc "" ⟩
|
||
test-url = refl
|
||
|
||
test-header : parse-header .parse (enc "Content-Length: 6\r\n")
|
||
≡ real ⟨ record { name = ASCII.encode-list⁺ "Content-Length" ; value = ASCII.encode-list⁺ "6" } , _ , enc "\r\n" ⟩
|
||
test-header = refl
|
||
|
||
test-request : parse-request .parse (enc "POST /site/index.html?mjau HTTP/1.1\r\nContent-Length: 6\r\nHost: agda.nu\r\n\r\nmjau:)\r\n")
|
||
≡ real
|
||
⟨ record
|
||
{ method = POST
|
||
; target = http:// "agda.nu" ꞉ 80 / ("site" / "index.html" / $) ¿ real "mjau"
|
||
; version = HTTP11
|
||
; headers =
|
||
record { name = ASCII.encode-list⁺ "Content-Length" ; value = ASCII.encode-list⁺ "6" }
|
||
L.∷ record { name = ASCII.encode-list⁺ "Host" ; value = ASCII.encode-list⁺ "agda.nu" }
|
||
L.∷ L.[]
|
||
; content = real (ASCII.encode-list "mjau:)")
|
||
}
|
||
, _
|
||
, enc "\r\n"
|
||
⟩
|
||
test-request = refl
|
||
|