oplss2024/ahmed/day1.agda
2024-06-04 18:56:00 -04:00

34 lines
1.1 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module ahmed.day1 where
open import Agda.Primitive
open import Data.Nat
open import Data.Fin
open import Data.Product
data Type : Set where
Bool : Type
_-→_ : Type Type Type
data Term (n : ) : Set where
Var : (m : Fin n) Term n
True : Term n
False : Term n
If_Then_Else_ : Term n Term n Term n Term n
λ[_::_]_ : Term n Type Term n Term n
_∙_ : Term n Term n Term n
data isValue (n : ) : Term n Set where
TrueValue : isValue n True
FalseValue : isValue n False
LambdaValue : (x : Term n) (τ : Type) (e : Term n) isValue n (λ[ x :: τ ] e)
Value = (n : ) Σ (Term n) (isValue n)
data evaluationContext : (n : ) Set where
dot : evaluationContext 0
EIf_Then_Else_ : {n : } evaluationContext n Term n Term n evaluationContext n
EAppLeft : {n : } evaluationContext n Term n evaluationContext n
EAppRight : {n : } Value evaluationContext n evaluationContext n
data isValidValueForType (n : ) (t : Type) : (e : Term n) Set where
TrueV : isValidValueForType n Bool True