blob: 44e3b6c54c88f92e59df9ad481d20a2c2a5c89ba (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
|
{-# OPTIONS --without-K --safe #-}
open import Relation.Binary as B using (Setoid)
module Cfe.Language.Indexed.Homogeneous
{c ℓ} (over : Setoid c ℓ)
where
open import Cfe.Language over
open import Level
open import Data.List
open import Data.Product
open import Relation.Binary.Indexed.Heterogeneous
open _≈_
open Setoid over using () renaming (Carrier to C)
record IndexedLanguage i iℓ a : Set (ℓ ⊔ suc (c ⊔ i ⊔ iℓ ⊔ a)) where
field
Carrierᵢ : Set i
_≈ᵢ_ : B.Rel Carrierᵢ iℓ
isEquivalenceᵢ : B.IsEquivalence _≈ᵢ_
F : Carrierᵢ → Language a
cong : F B.Preserves _≈ᵢ_ ⟶ _≈_
open B.IsEquivalence isEquivalenceᵢ using () renaming (refl to reflᵢ; sym to symᵢ; trans to transᵢ) public
Tagged : List C → Set (i ⊔ a)
Tagged l = ∃[ i ] l ∈ F i
|