{-# OPTIONS --with-K --safe #-}
module Data.Star.Environment {ℓ} (Ty : Set ℓ) where
open import Level
open import Data.Star.List
open import Data.Star.Decoration
open import Data.Star.Pointer as Pointer hiding (lookup)
open import Data.Unit
open import Function.Base hiding (_∋_)
open import Relation.Binary.PropositionalEquality
open import Relation.Binary.Construct.Closure.ReflexiveTransitive
Ctxt : Set ℓ
Ctxt = List Ty
infix 4 _∋_
_∋_ : Ctxt → Ty → Set ℓ
Γ ∋ σ = Any (const (Lift ℓ ⊤)) (σ ≡_) Γ
vz : ∀ {Γ σ} → Γ ▻ σ ∋ σ
vz = this refl
vs : ∀ {Γ σ τ} → Γ ∋ τ → Γ ▻ σ ∋ τ
vs = that _
Env : ∀ {e} → (Ty → Set e) → (Ctxt → Set (ℓ ⊔ e))
Env T Γ = All T Γ
lookup : ∀ {Γ σ} {T : Ty → Set} → Env T Γ → Γ ∋ σ → T σ
lookup ρ i with Pointer.lookup ρ i
... | result refl x = x