├── LICENSE ├── generic.agda-lib ├── readme.md └── src └── Generic ├── Core.agda ├── Function ├── Elim.agda └── FoldMono.agda ├── Lib ├── Category.agda ├── Data │ ├── List.agda │ ├── Maybe.agda │ ├── Nat.agda │ ├── Pow.agda │ ├── Product.agda │ ├── Sets.agda │ ├── String.agda │ └── Sum.agda ├── Decidable.agda ├── Equality │ ├── Coerce.agda │ ├── Congn.agda │ ├── Heteroindexed.agda │ └── Propositional.agda ├── Intro.agda ├── Prelude.agda └── Reflection │ ├── Core.agda │ └── Fold.agda ├── Main.agda ├── Property ├── Eq.agda └── Reify.agda ├── Reflection ├── DeriveEq.agda └── ReadData.agda ├── Test.agda └── Test ├── Data.agda ├── Data ├── Fin.agda ├── Lift.agda ├── List.agda ├── Maybe.agda ├── Product.agda ├── Vec.agda └── W.agda ├── DeriveEq.agda ├── Elim.agda ├── Eq.agda ├── Experiment.agda ├── ReadData.agda └── Reify.agda /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2020 effectfully 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /generic.agda-lib: -------------------------------------------------------------------------------- 1 | name: generic-0.1.0.2 2 | depend: standard-library 3 | include: src 4 | -------------------------------------------------------------------------------- /readme.md: -------------------------------------------------------------------------------- 1 | # Generic 2 | 3 | It's a library for doing generic programming in Agda. 4 | 5 | The library is tested with Agda-2.6.1 and Agda-2.6.1.2 and likely does not work with other versions of Agda. 6 | 7 | # A quick taste 8 | 9 | Deriving decidable equality for vectors: 10 | 11 | ``` 12 | open import Data.Vec using (Vec) renaming ([] to []ᵥ; _∷_ to _∷ᵥ_) 13 | 14 | instance VecEq : ∀ {n α} {A : Set α} {{aEq : Eq A}} -> Eq (Vec A n) 15 | unquoteDef VecEq = deriveEqTo VecEq (quote Vec) 16 | 17 | xs : Vec ℕ 3 18 | xs = 2 ∷ᵥ 4 ∷ᵥ 1 ∷ᵥ []ᵥ 19 | 20 | test₁ : xs ≟ xs ≡ yes refl 21 | test₁ = refl 22 | 23 | test₂ : xs ≟ (2 ∷ᵥ 4 ∷ᵥ 2 ∷ᵥ []ᵥ) ≡ no _ 24 | test₂ = refl 25 | ``` 26 | 27 | Same for `Data.Star`: 28 | 29 | ``` 30 | open import Data.Star 31 | 32 | instance StarEq : ∀ {i t} {I : Set i} {T : Rel I t} {i j} 33 | {{iEq : Eq I}} {{tEq : ∀ {i j} -> Eq (T i j)}} -> Eq (Star T i j) 34 | unquoteDef StarEq = deriveEqTo StarEq (quote Star) 35 | ``` 36 | 37 | # Internally 38 | 39 | Descriptions of constructors are defined as follows: 40 | 41 | ``` 42 | mutual 43 | Binder : ∀ {ι} α β γ -> Arg-info -> ι ⊔ lsuc (α ⊔ β) ≡ γ -> Set ι -> Set γ 44 | Binder α β γ i q I = Coerce q (∃ λ (A : Set α) -> < relevance i > A -> Desc I β) 45 | 46 | data Desc {ι} (I : Set ι) β : Set (ι ⊔ lsuc β) where 47 | var : I -> Desc I β 48 | π : ∀ {α} i 49 | -> (q : α ≤ℓ β) 50 | -> Binder α β _ i (cong (λ αβ -> ι ⊔ lsuc αβ) q) I 51 | -> Desc I β 52 | _⊛_ : Desc I β -> Desc I β -> Desc I β 53 | ``` 54 | 55 | Constructors are interpreted in the way described in [Descriptions](http://effectfully.blogspot.com/2016/04/descriptions.html) (in the `CompProp` module). That `Coerce` stuff is elaborated in [Emulating cumulativity in Agda](http://effectfully.blogspot.com/2016/07/cumu.html). 56 | 57 | A description of a data type is a list of named constructors 58 | 59 | ``` 60 | record Data {α} (A : Set α) : Set α where 61 | no-eta-equality 62 | constructor packData 63 | field 64 | dataName : Name 65 | parsTele : Type 66 | indsTele : Type 67 | consTypes : List A 68 | consNames : All (const Name) consTypes 69 | ``` 70 | 71 | For regular data types `A` is instantiated to `Type`, for described data types `A` is instantiated to `Desc I β` for some `I` and `β`. Descriptions also store the name of an original data type and telescopes of types of parameters and indices. `Name` and `Type` come from the `Reflection` module. 72 | 73 | There is a reflection machinery that allows to parse regular Agda data types into their described counterparts. An example from the [`Examples/ReadData.agda`](Examples/ReadData.agda) module: 74 | 75 | ``` 76 | data D {α β} (A : Set α) (B : ℕ -> Set β) : ∀ {n} -> B n -> List ℕ -> Set (α ⊔ β) where 77 | c₁ : ∀ {n} (y : B n) xs -> A -> D A B y xs 78 | c₂ : ∀ {y : B 0} -> (∀ {n} (y : B n) {{xs}} -> D A B y xs) -> List A -> D A B y [] 79 | 80 | D′ : ∀ {α β} (A : Set α) (B : ℕ -> Set β) {n} -> B n -> List ℕ -> Set (α ⊔ β) 81 | D′ = readData D 82 | 83 | pattern c₁′ {n} y xs x = #₀ (relv n , relv y , relv xs , relv x , lrefl) 84 | pattern c₂′ {y} r ys = !#₁ (relv y , r , irrv ys , lrefl) 85 | 86 | inj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D A B y xs -> D′ A B y xs 87 | inj (c₁ y xs x) = c₁′ y xs x 88 | inj (c₂ r ys) = c₂′ (λ y -> inj (r y)) ys 89 | 90 | outj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D′ A B y xs -> D A B y xs 91 | outj (c₁′ y xs x) = c₁ y xs x 92 | outj (c₂′ r ys) = c₂ (λ y -> outj (r y)) ys 93 | ``` 94 | 95 | So universe polymorphism is fully supported, as well as implicit and instance arguments, multiple (including single or none) parameters and indices, irrelevance (partly), higher-order inductive occurrences and you can define functions over described data types just like over the actual ones (though, [pattern synonyms are not equal in power to proper constructors](https://github.com/agda/agda/issues/2069)). 96 | 97 | There is a generic procedure that allows to coerce elements of described data type to elements of the corresponding regular data types, e.g. `outj` can be defined as 98 | 99 | ``` 100 | outj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D′ A B y xs -> D A B y xs 101 | outj d = guncoerce d 102 | ``` 103 | 104 | Internally it's a bit of reflection sugar on top of a generic fold defined on described data types (the [`Function/FoldMono.agda`](Function/FoldMono.agda) module). 105 | 106 | It's possible to coerce the other way around: 107 | 108 | ``` 109 | unquoteDecl foldD = deriveFoldTo foldD (quote D) 110 | 111 | inj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D A B y xs -> D′ A B y xs 112 | inj = gcoerce foldD 113 | ``` 114 | 115 | `foldD` is a derived (via reflection) indexed fold (like `foldr` on `Vec`) on `D`. The procedure that derives indexed folds for regular data types is in the [`Lib/Reflection/Fold.agda`](Lib/Reflection/Fold.agda) module. 116 | 117 | `D′` computes to the following term: 118 | 119 | ``` 120 | λ {.α} {.β} A B {n} z z₁ → 121 | μ 122 | (packData 123 | -- dataName 124 | (quote D) 125 | -- parsTele 126 | (implRelPi (pureDef (quote Level)) "α" 127 | (implRelPi (pureDef (quote Level)) "β" 128 | (explRelPi (sort (set (pureVar 1))) "A" 129 | (explRelPi (pureDef (quote ℕ) ‵→ sort (set (pureVar 2))) "B" 130 | unknown)))) 131 | -- indsTele 132 | (implRelPi (pureDef (quote ℕ)) "n" 133 | (appVar 1 (explRelArg (pureVar 0) ∷ []) ‵→ 134 | appDef (quote List) 135 | (implRelArg (pureDef (quote lzero)) ∷ 136 | explRelArg (pureDef (quote ℕ)) ∷ []) 137 | ‵→ 138 | sort 139 | (set 140 | (appDef (quote _⊔_) 141 | (explRelArg (pureVar 5) ∷ explRelArg (pureVar 6) ∷ []))))) 142 | -- consTypes 143 | (implRelDPi ℕ 144 | (λ rx → 145 | explRelDPi (B (unrelv rx)) 146 | (λ rx₁ → 147 | explRelDPi (List ℕ) 148 | (λ rx₂ → 149 | explRelDPi A (λ rx₃ → var (unrelv rx , unrelv rx₁ , unrelv rx₂))))) 150 | ∷ 151 | implRelDPi (B 0) 152 | (λ rx → 153 | implRelDPi ℕ 154 | (λ rx₁ → 155 | explRelDPi (B (unrelv rx₁)) 156 | (λ rx₂ → 157 | instRelDPi (List ℕ) 158 | (λ rx₃ → var (unrelv rx₁ , unrelv rx₂ , unrelv rx₃)))) 159 | ⊛ explIrrDPi (List A) (λ rx₁ → var (0 , unrelv rx , []))) 160 | ∷ []) 161 | -- consNames 162 | (quote c₁ , quote c₂ , tt)) 163 | (n , z , z₁) 164 | ``` 165 | 166 | Actual generic programming happens in the [`Property`](Property) subfolder. There is generic decidable equality defined over described data types. It can be used like this: 167 | 168 | ``` 169 | xs : Vec (List (Fin 4)) 3 170 | xs = (fsuc fzero ∷ fzero ∷ []) 171 | ∷ᵥ (fsuc (fsuc fzero) ∷ []) 172 | ∷ᵥ (fzero ∷ fsuc (fsuc (fsuc fzero)) ∷ []) 173 | ∷ᵥ []ᵥ 174 | 175 | test : xs ≟ xs ≡ yes refl 176 | test = refl 177 | ``` 178 | 179 | Equality for desribed `Vec`s, `List`s and `Fin`s is derived automatically. 180 | 181 | The [`Property/Reify.agda`](Property/Reify.agda) module implements coercion from described data types to `Term`s. Since stored names of described constructors are taken from actual constructors, reified elements of described data types are actually quoted elements of regular data types and hence the former can be converted to the latter (like with `guncoerce`, but deeply and accepts only canonical forms): 182 | 183 | ``` 184 | record Reify {α} (A : Set α) : Set α where 185 | field reify : A -> Term 186 | 187 | macro 188 | reflect : A -> Term -> TC _ 189 | reflect = unify ∘ reify 190 | open Reify {{...}} public 191 | 192 | instance 193 | DescReify : ∀ {i β} {I : Set i} {D : Desc I β} {j} 194 | {{reD : All (ExtendReify ∘ proj₂) D}} -> Reify (μ D j) 195 | DescReify = ... 196 | 197 | open import Generic.Examples.Data.Fin 198 | open import Generic.Examples.Data.Vec 199 | 200 | open import Data.Fin renaming (Fin to StdFin) 201 | open import Data.Vec renaming (Vec to StdVec) 202 | 203 | xs : Vec (Fin 4) 3 204 | xs = fsuc (fsuc (fsuc fzero)) ∷ᵥ fzero ∷ᵥ fsuc fzero ∷ᵥ []ᵥ 205 | 206 | xs′ : StdVec (StdFin 4) 3 207 | xs′ = suc (suc (suc zero)) ∷ zero ∷ (suc zero) ∷ [] 208 | 209 | test : reflect xs ≡ xs′ 210 | test = refl 211 | ``` 212 | 213 | Having decidable equality on `B` we can derive decidable equality on `A` if there is an injection `A ↦ B`. To construct an injection we need two functions `to : A -> B`, `from : B -> A` and a proof `from-to : from ∘ to ≗ id`. `to` and `from` are `gcoerce` and `guncoerce` from the above and `from-to` is another generic function (defined via reflection again, placed in [`Reflection/DeriveEq.agda`](Reflection/DeriveEq.agda): `fromToClausesOf` generates clauses for it) which uses universe polymorphic n-ary [`cong`](https://github.com/effectfully/Generic/blob/master/Lib/Equality/Congn.agda) under the hood. 214 | 215 | There are also generic `elim` in [`Function/Elim.agda`](Function/Elim.agda) (the idea is described in [Deriving eliminators of described data types](http://effectfully.blogspot.com/2016/06/deriving-eliminators-of-described-data.html) and `lookup` in [`Function/Lookup.agda`](Function/Lookup.agda) (broken currently). 216 | 217 | # Limitations 218 | 219 | - No support for mutually recursive data types. They can be supported, I just haven't implemented that. 220 | 221 | - No support for inductive-inductive or inductive-recursive data types. The latter [can be done](https://github.com/effectfully/random-stuff/blob/master/Desc/IRDesc.agda) at the cost of complicating the encoding. 222 | 223 | - No coinduction. 224 | 225 | - You can't describe a non-strictly positive data type. Yes, I think it's a limitation. I have [an idea](http://effectfully.blogspot.ru/2016/10/insane-descriptions.html) about how non-strictly positive and inductive-inductive data types can be described (it doesn't give you a way to define safe things safely, but this probably can be added). 226 | 227 | - Records can be described (see [`Examples/Data/Product.agda`](Examples/Data/Product.agda`)), but η-laws don't hold for them, because constructors contain `lift refl : tt ≡ tt` and `(p q : tt ≡ tt) -> p ≡ q` doesn't hold definitionally. `μ` is also a `data` rather than `record` (records confuse the termination checker, though, there are `{-# TERMINATING #-}` pragmas anyway), so this breaks η-expansion too. 228 | 229 | - Ornaments may or may not appear later (in the way described in [Unbiased ornaments](http://effectfully.blogspot.com/2016/07/unbiased-ornaments.html)). I don't find them very vital currently. 230 | 231 | - No forcing of indices. [`Lift`](Examples/Data/Lift.agda) can be described, though. 232 | -------------------------------------------------------------------------------- /src/Generic/Core.agda: -------------------------------------------------------------------------------- 1 | -- Everything is strictly positive, but Agda doesn't see this. 2 | {-# OPTIONS --no-positivity-check #-} 3 | 4 | module Generic.Core where 5 | 6 | open import Generic.Lib.Prelude public 7 | 8 | infix 4 _≤ℓ_ 9 | infixr 5 _⇒_ _⊛_ 10 | 11 | _≤ℓ_ : Level -> Level -> Set 12 | α ≤ℓ β = α ⊔ β ≡ β 13 | 14 | mutual 15 | Binder : ∀ {ι} α β γ -> ArgInfo -> ι ⊔ lsuc (α ⊔ β) ≡ γ -> Set ι -> Set γ 16 | Binder α β γ i q I = Coerce q (∃ λ (A : Set α) -> < relevance i > A -> Desc I β) 17 | 18 | data Desc {ι} (I : Set ι) β : Set (ι ⊔ lsuc β) where 19 | var : I -> Desc I β 20 | π : ∀ {α} i 21 | -> (q : α ≤ℓ β) 22 | -> Binder α β _ i (cong (λ αβ -> ι ⊔ lsuc αβ) q) I 23 | -> Desc I β 24 | _⊛_ : Desc I β -> Desc I β -> Desc I β 25 | 26 | pattern DPi i A D = π i refl (coerce (A , D)) 27 | 28 | {-# DISPLAY π i refl (coerce (A , D)) = DPi i A D #-} 29 | 30 | pattern explRelDPi A D = DPi explRelInfo A D 31 | pattern explIrrDPi A D = DPi explIrrInfo A D 32 | pattern implRelDPi A D = DPi implRelInfo A D 33 | pattern implIrrDPi A D = DPi implIrrInfo A D 34 | pattern instRelDPi A D = DPi instRelInfo A D 35 | pattern instIrrDPi A D = DPi instIrrInfo A D 36 | 37 | {-# DISPLAY DPi explRelInfo A D = explRelDPi A D #-} 38 | {-# DISPLAY DPi explIrrInfo A D = explIrrDPi A D #-} 39 | {-# DISPLAY DPi implRelInfo A D = implRelDPi A D #-} 40 | {-# DISPLAY DPi implIrrInfo A D = implIrrDPi A D #-} 41 | {-# DISPLAY DPi instRelInfo A D = instRelDPi A D #-} 42 | {-# DISPLAY DPi instIrrInfo A D = instIrrDPi A D #-} 43 | 44 | _⇒_ : ∀ {ι α β} {I : Set ι} {{q : α ≤ℓ β}} -> Set α -> Desc I β -> Desc I β 45 | _⇒_ {{q}} A D = π (explRelInfo) q (qcoerce (A , λ _ -> D)) 46 | 47 | mutual 48 | ⟦_⟧ : ∀ {ι β} {I : Set ι} -> Desc I β -> (I -> Set β) -> Set β 49 | ⟦ var i ⟧ B = B i 50 | ⟦ π i q C ⟧ B = ⟦ i / C ⟧ᵇ q B 51 | ⟦ D ⊛ E ⟧ B = ⟦ D ⟧ B × ⟦ E ⟧ B 52 | 53 | ⟦_/_⟧ᵇ : ∀ {α ι β γ q} {I : Set ι} i 54 | -> Binder α β γ i q I -> α ≤ℓ β -> (I -> Set β) -> Set β 55 | ⟦ i / coerce (A , D) ⟧ᵇ q B = Coerce′ q $ Pi i A λ x -> ⟦ D x ⟧ B 56 | 57 | mutual 58 | Extend : ∀ {ι β} {I : Set ι} -> Desc I β -> (I -> Set β) -> I -> Set β 59 | Extend (var i) B j = Lift _ (i ≡ j) 60 | Extend (π i q C) B j = Extendᵇ i C q B j 61 | Extend (D ⊛ E) B j = ⟦ D ⟧ B × Extend E B j 62 | 63 | Extendᵇ : ∀ {ι α β γ q} {I : Set ι} i 64 | -> Binder α β γ i q I -> α ≤ℓ β -> (I -> Set β) -> I -> Set β 65 | Extendᵇ i (coerce (A , D)) q B j = Coerce′ q $ ∃ λ x -> Extend (D x) B j 66 | 67 | module _ {ι β} {I : Set ι} (D : Data (Desc I β)) where 68 | mutual 69 | data μ j : Set β where 70 | node : Node D j -> μ j 71 | 72 | Node : Data (Desc I β) -> I -> Set β 73 | Node D j = Any (λ C -> Extend C μ j) (consTypes D) 74 | 75 | mutual 76 | Cons : ∀ {ι β} {I : Set ι} -> (I -> Set β) -> Desc I β -> Set β 77 | Cons B (var i) = B i 78 | Cons B (π i q C) = Consᵇ B i C q 79 | Cons B (D ⊛ E) = ⟦ D ⟧ B -> Cons B E 80 | 81 | Consᵇ : ∀ {ι α β γ q} {I : Set ι} 82 | -> (I -> Set β) -> ∀ i -> Binder α β γ i q I -> α ≤ℓ β -> Set β 83 | Consᵇ B i (coerce (A , D)) q = Coerce′ q $ Pi i A λ x -> Cons B (D x) 84 | 85 | cons : ∀ {ι β} {I : Set ι} {D} -> (D₀ : Data (Desc I β)) -> D ∈ consTypes D₀ -> Cons (μ D₀) D 86 | cons {D = D} D₀ p = go D λ e -> 87 | node (mapAny (consTypes D₀) (λ q -> subst (λ E -> Extend E _ _) q e) p) where 88 | mutual 89 | go : ∀ {ι β} {I : Set ι} {B : I -> Set β} 90 | -> (D : Desc I β) -> (∀ {j} -> Extend D B j -> B j) -> Cons B D 91 | go (var i) k = k lrefl 92 | go (π a q C) k = goᵇ a C k 93 | go (D ⊛ E) k = λ x -> go E (k ∘ _,_ x) 94 | 95 | goᵇ : ∀ {ι α β γ q q′} {I : Set ι} {B : I -> Set β} i 96 | -> (C : Binder α β γ i q′ I) -> (∀ {j} -> Extendᵇ i C q B j -> B j) -> Consᵇ B i C q 97 | goᵇ {q = q} i (coerce (A , D)) k = 98 | coerce′ q $ lamPi i λ x -> go (D x) (k ∘ coerce′ q ∘ _,_ x) 99 | 100 | allCons : ∀ {ι β} {I : Set ι} -> (D : Data (Desc I β)) -> All (Cons (μ D)) (consTypes D) 101 | allCons D = allIn _ (cons D) 102 | 103 | node-inj : ∀ {i β} {I : Set i} {D : Data (Desc I β)} {j} {e₁ e₂ : Node D D j} 104 | -> node {D = D} e₁ ≡ node e₂ -> e₁ ≡ e₂ 105 | node-inj refl = refl 106 | 107 | μ′ : ∀ {β} -> Data (Desc ⊤₀ β) -> Set β 108 | μ′ D = μ D tt 109 | 110 | pos : ∀ {β} -> Desc ⊤₀ β 111 | pos = var tt 112 | 113 | pattern #₀ p = node (inj₁ p) 114 | pattern #₁ p = node (inj₂ (inj₁ p)) 115 | pattern #₂ p = node (inj₂ (inj₂ (inj₁ p))) 116 | pattern #₃ p = node (inj₂ (inj₂ (inj₂ (inj₁ p)))) 117 | pattern #₄ p = node (inj₂ (inj₂ (inj₂ (inj₂ (inj₁ p))))) 118 | pattern #₅ p = node (inj₂ (inj₂ (inj₂ (inj₂ (inj₂ (inj₁ p)))))) 119 | 120 | pattern !#₀ p = node p 121 | pattern !#₁ p = node (inj₂ p) 122 | pattern !#₂ p = node (inj₂ (inj₂ p)) 123 | pattern !#₃ p = node (inj₂ (inj₂ (inj₂ p))) 124 | pattern !#₄ p = node (inj₂ (inj₂ (inj₂ (inj₂ p)))) 125 | pattern !#₅ p = node (inj₂ (inj₂ (inj₂ (inj₂ (inj₂ p))))) 126 | -------------------------------------------------------------------------------- /src/Generic/Function/Elim.agda: -------------------------------------------------------------------------------- 1 | module Generic.Function.Elim where 2 | 3 | open import Generic.Core 4 | 5 | AllAny : ∀ {ι α β γ δ} {I : Set ι} {A : Set α} {C : I -> Set γ} 6 | -> (B : I -> A -> Set β) 7 | -> (∀ x -> (∀ {j} -> B j x -> C j) -> Set δ) 8 | -> (xs : List A) 9 | -> (∀ {j} -> Any (B j) xs -> C j) 10 | -> Set δ 11 | AllAny B D [] k = ⊤ 12 | AllAny B D (x ∷ []) k = D x k 13 | AllAny B D (x ∷ y ∷ xs) k = D x (k ∘ inj₁) × AllAny B D (y ∷ xs) (k ∘ inj₂) 14 | 15 | data VarView {ι β} {I : Set ι} : Desc I β -> Set where 16 | yes-var : ∀ {i} -> VarView (var i) 17 | no-var : ∀ {D} -> VarView D 18 | 19 | varView : ∀ {ι β} {I : Set ι} -> (D : Desc I β) -> VarView D 20 | varView (var i) = yes-var 21 | varView D = no-var 22 | 23 | mutual 24 | Hyp : ∀ {ι β γ} {I : Set ι} {B} 25 | -> (∀ {i} -> B i -> Set γ) -> (D : Desc I β) -> ⟦ D ⟧ B -> Set (β ⊔ γ) 26 | Hyp {β = β} C (var i) y = Lift β (C y) 27 | Hyp C (π i q D) f = Hypᵇ i C D f 28 | Hyp C (D ⊛ E) (x , y) = Hyp C D x × Hyp C E y 29 | 30 | Hypᵇ : ∀ {ι α β γ δ q q′} {I : Set ι} {B} i 31 | -> (∀ {i} -> B i -> Set γ) -> (D : Binder α β δ i q′ I) -> ⟦ i / D ⟧ᵇ q B -> Set (β ⊔ γ) 32 | Hypᵇ {γ = γ} {q = q} i C (coerce (A , D)) f = 33 | Coerce′ (cong (γ ⊔_) q) $ Pi i A λ x -> Hyp C (D x) (appPi i (uncoerce′ q f) x) 34 | 35 | mutual 36 | Elim : ∀ {ι β γ} {I : Set ι} {B} 37 | -> (∀ {i} -> B i -> Set γ) 38 | -> (D : Desc I β) 39 | -> (∀ {j} -> Extend D B j -> B j) 40 | -> Set (β ⊔ γ) 41 | Elim {β = β} C (var i) k = Lift β (C (k lrefl)) 42 | Elim C (π i q D) k = Elimᵇ i C D k 43 | Elim C (D ⊛ E) k with varView D 44 | ... | yes-var = ∀ {x} -> C x -> Elim C E (k ∘ _,_ x) 45 | ... | no-var = ∀ {x} -> Hyp C D x -> Elim C E (k ∘ _,_ x) 46 | 47 | Elimᵇ : ∀ {ι α β γ δ q q′} {I : Set ι} {B} i 48 | -> (∀ {i} -> B i -> Set γ) 49 | -> (D : Binder α β δ i q′ I) 50 | -> (∀ {j} -> Extendᵇ i D q B j -> B j) 51 | -> Set (β ⊔ γ) 52 | Elimᵇ {γ = γ} {q = q} i C (coerce (A , D)) k = 53 | Coerce′ (cong (γ ⊔_) q) $ Pi i A λ x -> Elim C (D x) (k ∘ coerce′ q ∘ _,_ x) 54 | 55 | module _ {ι β γ} {I : Set ι} {D₀ : Data (Desc I β)} (C : ∀ {j} -> μ D₀ j -> Set γ) where 56 | K : Name -> Type -> Type -> (Ds : List (Desc I β)) -> All (const Name) Ds -> Set (ι ⊔ β) 57 | K d a b Ds ns = ∀ {j} -> Node D₀ (packData d a b Ds ns) j -> μ D₀ j 58 | 59 | Elims : ∀ d a b (Ds : List (Desc I β)) ns -> K d a b Ds ns -> Set (β ⊔ γ) 60 | Elims d a b Ds ns = AllAny (λ j D -> Extend D (μ D₀) j) (Elim C) Ds 61 | 62 | module _ (hs : Elims (dataName D₀) 63 | (parsTele D₀) 64 | (indsTele D₀) 65 | (consTypes D₀) 66 | (consNames D₀) 67 | node) where 68 | {-# TERMINATING #-} 69 | mutual 70 | elimHyp : (D : Desc I β) -> (d : ⟦ D ⟧ (μ D₀)) -> Hyp C D d 71 | elimHyp (var i) d = lift (elim d) 72 | elimHyp (π i q D) f = elimHypᵇ i D f 73 | elimHyp (D ⊛ E) (x , y) = elimHyp D x , elimHyp E y 74 | 75 | elimHypᵇ : ∀ {α δ q q′} i 76 | -> (D : Binder α β δ i q′ I) 77 | -> (f : ⟦ i / D ⟧ᵇ q (μ D₀)) 78 | -> Hypᵇ i C D f 79 | elimHypᵇ {q = q} i (coerce (A , D)) f = 80 | coerce′ (cong (_⊔_ γ) q) (lamPi i λ x -> elimHyp (D x) (appPi i (uncoerce′ q f) x)) 81 | 82 | elimExtend : ∀ {j} 83 | -> (D : Desc I β) {k : ∀ {j} -> Extend D (μ D₀) j -> μ D₀ j} 84 | -> Elim C D k 85 | -> (e : Extend D (μ D₀) j) 86 | -> C (k e) 87 | elimExtend (var i) z lrefl = lower z 88 | elimExtend (π i q D) h p = elimExtendᵇ i D h p 89 | elimExtend (D ⊛ E) h (d , e) with varView D 90 | ... | yes-var = elimExtend E (h (elim d)) e 91 | ... | no-var = elimExtend E (h (elimHyp D d)) e 92 | 93 | elimExtendᵇ : ∀ {α δ q q′ j} i 94 | -> (D : Binder α β δ i q′ I) {k : ∀ {j} -> Extendᵇ i D q (μ D₀) j -> μ D₀ j} 95 | -> Elimᵇ i C D k 96 | -> (p : Extendᵇ i D q (μ D₀) j) 97 | -> C (k p) 98 | elimExtendᵇ {q = q} i (coerce (A , D)) h p with p | inspectUncoerce′ q p 99 | ... | _ | (x , e) , refl = elimExtend (D x) (appPi i (uncoerce′ (cong (γ ⊔_) q) h) x) e 100 | 101 | elimAny : ∀ {j} (Ds : List (Desc I β)) d a b ns {k : K d a b Ds ns} 102 | -> Elims d a b Ds ns k -> (a : Node D₀ (packData d a b Ds ns) j) -> C (k a) 103 | elimAny [] d a b tt tt () 104 | elimAny (D ∷ []) d a b (n , ns) h e = elimExtend D h e 105 | elimAny (D ∷ E ∷ Ds) d a b (n , ns) (h , hs) (inj₁ e) = elimExtend D h e 106 | elimAny (D ∷ E ∷ Ds) d a b (n , ns) (h , hs) (inj₂ r) = elimAny (E ∷ Ds) d a b ns hs r 107 | 108 | elim : ∀ {j} -> (d : μ D₀ j) -> C d 109 | elim (node e) = elimAny (consTypes D₀) 110 | (dataName D₀) 111 | (parsTele D₀) 112 | (indsTele D₀) 113 | (consNames D₀) 114 | hs 115 | e 116 | -------------------------------------------------------------------------------- /src/Generic/Function/FoldMono.agda: -------------------------------------------------------------------------------- 1 | module Generic.Function.FoldMono where 2 | 3 | open import Generic.Core 4 | 5 | CurryAll : ∀ {α β} {A : Set α} -> (A -> Set β) -> List A -> Set β -> Set β 6 | CurryAll B [] C = C 7 | CurryAll B (x ∷ xs) C = B x -> CurryAll B xs C 8 | 9 | curryAll : ∀ {α β} {A : Set α} {B : A -> Set β} {C xs} 10 | -> (All B xs -> C) -> CurryAll B xs C 11 | curryAll {xs = []} g = g tt 12 | curryAll {xs = x ∷ xs} g = curryAll ∘ curry g 13 | 14 | mutual 15 | Hyp : ∀ {ι β} {I : Set ι} -> (I -> Set β) -> (D : Desc I β) -> Set β 16 | Hyp C (var i) = C i 17 | Hyp C (π i q D) = Hypᵇ i C D q 18 | Hyp C (D ⊛ E) = Hyp C D × Hyp C E 19 | 20 | Hypᵇ : ∀ {α ι β γ q} {I : Set ι} i 21 | -> (I -> Set β) -> Binder α β γ i q I -> α ≤ℓ β -> Set β 22 | Hypᵇ i C (coerce (A , D)) q = Coerce′ q $ Pi i A λ x -> Hyp C (D x) 23 | 24 | mutual 25 | Fold : ∀ {ι β} {I : Set ι} -> (I -> Set β) -> (D : Desc I β) -> Set β 26 | Fold C (var i) = C i 27 | Fold C (π i q D) = Foldᵇ i C D q 28 | Fold C (D ⊛ E) = Hyp C D -> Fold C E 29 | 30 | Foldᵇ : ∀ {α ι β γ q} {I : Set ι} i 31 | -> (I -> Set β) -> Binder α β γ i q I -> α ≤ℓ β -> Set β 32 | Foldᵇ i C (coerce (A , D)) q = Coerce′ q $ Pi i A λ x -> Fold C (D x) 33 | 34 | module _ {ι β} {I : Set ι} {D₀ : Data (Desc I β)} (C : I -> Set β) where 35 | module _ (hs : All (Fold C) (consTypes D₀)) where 36 | {-# TERMINATING #-} 37 | mutual 38 | foldHyp : (D : Desc I β) -> ⟦ D ⟧ (μ D₀) -> Hyp C D 39 | foldHyp (var i) d = foldMono d 40 | foldHyp (π i q D) f = foldHypᵇ i D f 41 | foldHyp (D ⊛ E) (x , y) = foldHyp D x , foldHyp E y 42 | 43 | foldHypᵇ : ∀ {α γ q q′} i 44 | -> (D : Binder α β γ i q′ I) -> ⟦ i / D ⟧ᵇ q (μ D₀) -> Hypᵇ i C D q 45 | foldHypᵇ {q = q} i (coerce (A , D)) f = 46 | coerce′ q $ lamPi i λ x -> foldHyp (D x) (appPi i (uncoerce′ q f) x) 47 | 48 | foldExtend : ∀ {j} -> (D : Desc I β) -> Fold C D -> Extend D (μ D₀) j -> C j 49 | foldExtend (var i) z lrefl = z 50 | foldExtend (π i q D) h p = foldExtendᵇ i D h p 51 | foldExtend (D ⊛ E) h (d , e) = foldExtend E (h (foldHyp D d)) e 52 | 53 | foldExtendᵇ : ∀ {α γ q q′ j} i 54 | -> (D : Binder α β γ i q′ I) -> Foldᵇ i C D q -> Extendᵇ i D q (μ D₀) j -> C j 55 | foldExtendᵇ {q = q} i (coerce (A , D)) h p with p | inspectUncoerce′ q p 56 | ... | _ | (x , e) , refl = foldExtend (D x) (appPi i (uncoerce′ q h) x) e 57 | 58 | foldAny : ∀ {j} (Ds : List (Desc I β)) d a b ns 59 | -> All (Fold C) Ds -> Node D₀ (packData d a b Ds ns) j -> C j 60 | foldAny [] d a b tt tt () 61 | foldAny (D ∷ []) d a b (_ , ns) (h , tt) e = foldExtend D h e 62 | foldAny (D ∷ E ∷ Ds) d a b (_ , ns) (h , hs) (inj₁ e) = foldExtend D h e 63 | foldAny (D ∷ E ∷ Ds) d a b (_ , ns) (h , hs) (inj₂ r) = foldAny (E ∷ Ds) d a b ns hs r 64 | 65 | foldMono : ∀ {j} -> μ D₀ j -> C j 66 | foldMono (node e) = foldAny (consTypes D₀) 67 | (dataName D₀) 68 | (parsTele D₀) 69 | (indsTele D₀) 70 | (consNames D₀) 71 | hs 72 | e 73 | 74 | curryFoldMono : ∀ {j} -> μ D₀ j -> CurryAll (Fold C) (consTypes D₀) (C j) 75 | curryFoldMono d = curryAll λ hs -> foldMono hs d 76 | -------------------------------------------------------------------------------- /src/Generic/Lib/Category.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Category where 2 | 3 | open import Category.Functor public hiding (Morphism) 4 | open import Category.Applicative public 5 | open import Category.Monad public 6 | 7 | open RawFunctor {{...}} public 8 | open RawApplicative {{...}} hiding (_<$>_; _<&>_; _<$_; zip; zipWith) renaming (_⊛_ to _<*>_) public 9 | open RawMonad {{...}} hiding (pure; _<$>_; _<&>_; _<$_; _⊛_; _<⊛_; _⊛>_; _⊗_; rawFunctor; zip; zipWith) public 10 | 11 | fmap = _<$>_ 12 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/List.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.List where 2 | 3 | open import Data.List.Base hiding ([]) renaming (fromMaybe to maybeToList) public 4 | open List public 5 | 6 | open import Generic.Lib.Intro 7 | open import Generic.Lib.Equality.Propositional 8 | open import Generic.Lib.Decidable 9 | open import Generic.Lib.Category 10 | open import Generic.Lib.Data.Nat 11 | open import Generic.Lib.Data.Maybe 12 | open import Generic.Lib.Data.Sum 13 | open import Generic.Lib.Data.Product 14 | 15 | import Data.List.Categorical as List 16 | 17 | infix 4 _∈_ 18 | 19 | foldr₁ : ∀ {α} {A : Set α} -> (A -> A -> A) -> A -> List A -> A 20 | foldr₁ f z [] = z 21 | foldr₁ f z (x ∷ []) = x 22 | foldr₁ f z (x ∷ y ∷ xs) = f x (foldr₁ f z (y ∷ xs)) 23 | 24 | mapInd : ∀ {α β} {A : Set α} {B : Set β} -> (ℕ -> A -> B) -> List A -> List B 25 | mapInd f [] = [] 26 | mapInd f (x ∷ xs) = f 0 x ∷ mapInd (f ∘ suc) xs 27 | 28 | mapM : ∀ {α β} {A : Set α} {B : Set β} {M : Set β -> Set β} {{mMonad : RawMonad M}} 29 | -> (A -> M B) -> List A -> M (List B) 30 | mapM {{mMonad}} = List.TraversableM.mapM mMonad 31 | 32 | downFromTo : ℕ -> ℕ -> List ℕ 33 | downFromTo n m = map (m +_) (downFrom (n ∸ m)) 34 | 35 | Any : ∀ {α β} {A : Set α} -> (A -> Set β) -> List A -> Set β 36 | Any B [] = ⊥ 37 | Any B (x ∷ []) = B x 38 | Any B (x ∷ xs) = B x ⊎ Any B xs 39 | 40 | _∈_ : ∀ {α} {A : Set α} -> A -> List A -> Set 41 | x ∈ xs = Any (x ≡_) xs 42 | 43 | here : ∀ {α β} {A : Set α} {B : A -> Set β} {x} xs -> B x -> Any B (x ∷ xs) 44 | here [] y = y 45 | here (x ∷ xs) y = inj₁ y 46 | 47 | phere : ∀ {α} {A : Set α} {x : A} xs -> x ∈ x ∷ xs 48 | phere xs = here xs refl 49 | 50 | there : ∀ {α β} {A : Set α} {B : A -> Set β} {x} xs -> Any B xs -> Any B (x ∷ xs) 51 | there [] () 52 | there (x ∷ xs) a = inj₂ a 53 | 54 | mapAny : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} 55 | -> ∀ xs -> (∀ {x} -> B x -> C x) -> Any B xs -> Any C xs 56 | mapAny [] g () 57 | mapAny (x ∷ []) g y = g y 58 | mapAny (x ∷ x′ ∷ xs) g (inj₁ y) = inj₁ (g y) 59 | mapAny (x ∷ x′ ∷ xs) g (inj₂ r) = inj₂ (mapAny (x′ ∷ xs) g r) 60 | 61 | All : ∀ {α β} {A : Set α} -> (A -> Set β) -> List A -> Set β 62 | All B [] = ⊤ 63 | All B (x ∷ xs) = B x × All B xs 64 | 65 | allToList : ∀ {α β} {A : Set α} {B : Set β} {xs : List A} -> All (const B) xs -> List B 66 | allToList {xs = []} tt = [] 67 | allToList {xs = x ∷ xs} (y , ys) = y ∷ allToList ys 68 | 69 | allIn : ∀ {α β} {A : Set α} {B : A -> Set β} -> ∀ xs -> (∀ {x} -> x ∈ xs -> B x) -> All B xs 70 | allIn [] g = tt 71 | allIn (x ∷ xs) g = g (phere xs) , allIn xs (g ∘ there xs) 72 | 73 | mapAll : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} {xs : List A} 74 | -> (∀ {x} -> B x -> C x) -> All B xs -> All C xs 75 | mapAll {xs = []} g tt = tt 76 | mapAll {xs = x ∷ xs} g (y , ys) = g y , mapAll g ys 77 | 78 | unmap : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : Set γ} {xs : List A} 79 | -> (∀ {x} -> B x -> C) -> All B xs -> List C 80 | unmap g = allToList ∘ mapAll g 81 | 82 | mapAllInd : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} {xs} 83 | -> (∀ {x} -> ℕ -> B x -> C x) -> All B xs -> All C xs 84 | mapAllInd {xs = []} g tt = tt 85 | mapAllInd {xs = x ∷ xs} g (y , ys) = g 0 y , mapAllInd (g ∘ suc) ys 86 | 87 | traverseAll : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} {xs : List A} 88 | {F : Set γ -> Set γ} {{fFunctor : RawFunctor F}} {{fApplicative : RawApplicative F}} 89 | -> (∀ {x} -> B x -> F (C x)) -> All B xs -> F (All C xs) 90 | traverseAll {xs = []} g tt = pure tt 91 | traverseAll {xs = x ∷ xs} g (y , ys) = _,_ <$> g y <*> traverseAll g ys 92 | 93 | splitList : ∀ {α β} {A : Set α} {B : A -> Set β} -> List (Σ A B) -> Σ (List A) (All B) 94 | splitList [] = [] , tt 95 | splitList ((x , y) ∷ ps) = pmap (_∷_ x) (_,_ y) (splitList ps) 96 | 97 | lookupAllConst : ∀ {α β} {A : Set α} {B : Set β} {{bEq : Eq B}} {xs : List A} 98 | -> B -> All (const B) xs -> Maybe (∃ (_∈ xs)) 99 | lookupAllConst {xs = []} y tt = nothing 100 | lookupAllConst {xs = x ∷ xs} y (z , ys) = if y == z 101 | then just (_ , phere xs) 102 | else second (there xs) <$> lookupAllConst y ys 103 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Maybe.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Maybe where 2 | 3 | open import Data.Maybe.Base using (Maybe; just; nothing; maybe; maybe′; from-just; fromMaybe) public 4 | 5 | open import Generic.Lib.Intro 6 | open import Generic.Lib.Category 7 | 8 | infixr 0 _?>_ 9 | 10 | instance 11 | MaybeMonad : ∀ {α} -> RawMonad {α} Maybe 12 | MaybeMonad = record 13 | { return = just 14 | ; _>>=_ = Data.Maybe.Base._>>=_ 15 | } 16 | 17 | MaybeApplicative : ∀ {α} -> RawApplicative {α} Maybe 18 | MaybeApplicative = rawIApplicative 19 | 20 | MaybeFunctor : ∀ {α} -> RawFunctor {α} Maybe 21 | MaybeFunctor = rawFunctor 22 | 23 | _?>_ : ∀ {α} {A : Set α} -> Bool -> A -> Maybe A 24 | true ?> x = just x 25 | false ?> x = nothing 26 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Nat.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Nat where 2 | 3 | open import Data.Nat.Base hiding (_⊔_; _^_) public 4 | 5 | open import Generic.Lib.Decidable 6 | 7 | instance 8 | ℕEq : Eq ℕ 9 | ℕEq = viaBase Nat._≟_ where 10 | import Data.Nat as Nat 11 | 12 | foldℕ : ∀ {α} {A : Set α} -> (A -> A) -> A -> ℕ -> A 13 | foldℕ f x 0 = x 14 | foldℕ f x (suc n) = f (foldℕ f x n) 15 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Pow.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Pow where 2 | 3 | open import Generic.Lib.Intro 4 | open import Generic.Lib.Data.Nat 5 | open import Generic.Lib.Data.Product 6 | 7 | infixl 6 _^_ 8 | 9 | _^_ : ∀ {α} -> Set α -> ℕ -> Set α 10 | A ^ 0 = ⊤ 11 | A ^ suc n = A × A ^ n 12 | 13 | elimPow : ∀ {n α} {A : Set α} {b : ∀ {n} -> A ^ n -> Level} 14 | -> (B : ∀ {n} -> (xs : A ^ n) -> Set (b xs)) 15 | -> (∀ {n} {xs : A ^ n} x -> B xs -> B (x , xs)) 16 | -> B tt 17 | -> (xs : A ^ n) 18 | -> B xs 19 | elimPow {0} B f z tt = z 20 | elimPow {suc n} B f z (x , xs) = f x (elimPow B f z xs) 21 | 22 | foldPow : ∀ {n α β} {A : Set α} 23 | -> (B : ℕ -> Set β) -> (∀ {n} -> A -> B n -> B (suc n)) -> B 0 -> A ^ n -> B n 24 | foldPow B f z xs = elimPow (λ {n} _ -> B n) (λ x y -> f x y) z xs 25 | 26 | foldPow₁ : ∀ {n α} {A : Set α} -> (A -> A -> A) -> A ^ suc n -> A 27 | foldPow₁ {0} f (x , []) = x 28 | foldPow₁ {suc n} f (x , xs) = f x (foldPow₁ f xs) 29 | 30 | mapPow : ∀ {n α β} {A : Set α} {B : Set β} -> (A -> B) -> A ^ n -> B ^ n 31 | mapPow f = foldPow (_ ^_) (_,_ ∘ f) tt 32 | 33 | replicatePow : ∀ {α} {A : Set α} n -> A -> A ^ n 34 | replicatePow 0 x = tt 35 | replicatePow (suc n) x = x , replicatePow n x 36 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Product.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Product where 2 | 3 | open import Data.Product renaming (map to pmap; zip to pzip) hiding (map₁; map₂; dmap; _<*>_; assocʳ; assocˡ) public 4 | 5 | open import Generic.Lib.Intro 6 | open import Generic.Lib.Category 7 | 8 | infixl 4 _,ᵢ_ 9 | 10 | first : ∀ {α β γ} {A : Set α} {B : Set β} {C : A -> Set γ} 11 | -> (∀ x -> C x) -> (p : A × B) -> C (proj₁ p) × B 12 | first f (x , y) = f x , y 13 | 14 | second : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : A -> Set γ} 15 | -> (∀ {x} -> B x -> C x) -> Σ A B -> Σ A C 16 | second g (x , y) = x , g y 17 | 18 | -- `B` and `C` are in the same universe. 19 | firstF : ∀ {α β} {A : Set α} {B : Set β} {C : A -> Set β} 20 | {F : Set β -> Set β} {{fFunctor : RawFunctor F}} 21 | -> (∀ x -> F (C x)) -> (p : A × B) -> F (C (proj₁ p) × B) 22 | firstF f (x , y) = flip _,_ y <$> f x 23 | 24 | -- `A` and `C` are in the same universe. 25 | secondF : ∀ {α β} {A : Set α} {B : A -> Set β} {C : A -> Set α} 26 | {F : Set α -> Set α} {{fFunctor : RawFunctor F}} 27 | -> (∀ {x} -> B x -> F (C x)) -> Σ A B -> F (Σ A C) 28 | secondF g (x , y) = _,_ x <$> g y 29 | 30 | record Σᵢ {α β} (A : Set α) (B : .A -> Set β) : Set (α ⊔ β) where 31 | constructor _,ᵢ_ 32 | field 33 | .iproj₁ : A 34 | iproj₂ : B iproj₁ 35 | open Σᵢ public 36 | 37 | ∃ᵢ : ∀ {α β} {A : Set α} -> (.A -> Set β) -> Set (α ⊔ β) 38 | ∃ᵢ = Σᵢ _ 39 | 40 | instance 41 | ,-inst : ∀ {α β} {A : Set α} {B : A -> Set β} {{x : A}} {{y : B x}} -> Σ A B 42 | ,-inst {{x}} {{y}} = x , y 43 | 44 | ,ᵢ-inst : ∀ {α β} {A : Set α} {B : .A -> Set β} {{x : A}} {{y : B x}} -> Σᵢ A B 45 | ,ᵢ-inst {{x}} {{y}} = x ,ᵢ y 46 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Sets.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Sets where 2 | 3 | open import Generic.Lib.Intro 4 | open import Generic.Lib.Data.Nat 5 | open import Generic.Lib.Data.Product 6 | open import Generic.Lib.Data.Pow 7 | 8 | infixl 6 _⊔ⁿ_ 9 | 10 | _⊔ⁿ_ : ∀ {n} -> Level ^ n -> Level -> Level 11 | _⊔ⁿ_ = flip $ foldPow _ _⊔_ 12 | 13 | Sets : ∀ {n} -> (αs : Level ^ n) -> Set (mapPow lsuc αs ⊔ⁿ lzero) 14 | Sets {0} _ = ⊤ 15 | Sets {suc _} (α , αs) = Set α × Sets αs 16 | 17 | FoldSets : ∀ {n β αs} -> Sets {n} αs -> Set β -> Set (αs ⊔ⁿ β) 18 | FoldSets {0} tt B = B 19 | FoldSets {suc _} (A , As) B = A -> FoldSets As B 20 | 21 | HList : ∀ {n} {αs} -> Sets {n} αs -> Set (αs ⊔ⁿ lzero) 22 | HList {0} tt = ⊤ 23 | HList {suc _} (A , As) = A × HList As 24 | 25 | applyFoldSets : ∀ {n β αs} {As : Sets {n} αs} {B : Set β} -> FoldSets As B -> HList As -> B 26 | applyFoldSets {0} y tt = y 27 | applyFoldSets {suc n} f (x , xs) = applyFoldSets (f x) xs 28 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/String.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.String where 2 | 3 | open import Data.String.Base using (String) renaming (toList to toListˢ; _++_ to _++ˢ_) public 4 | 5 | open import Generic.Lib.Decidable 6 | import Data.String as String 7 | 8 | instance 9 | StringEq : Eq String 10 | StringEq = viaBase String._≟_ 11 | -------------------------------------------------------------------------------- /src/Generic/Lib/Data/Sum.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Data.Sum where 2 | 3 | open import Data.Sum hiding (swap) renaming (map to smap) hiding (map₁; map₂; assocʳ; assocˡ; reduce) public 4 | -------------------------------------------------------------------------------- /src/Generic/Lib/Decidable.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Decidable where 2 | 3 | open import Relation.Nullary public 4 | open import Relation.Nullary.Decidable hiding (map) public 5 | open import Relation.Binary using (Decidable) public 6 | 7 | open import Generic.Lib.Intro 8 | open import Generic.Lib.Equality.Propositional 9 | open import Generic.Lib.Equality.Heteroindexed 10 | open import Generic.Lib.Data.Sum 11 | open import Generic.Lib.Data.Product 12 | 13 | open import Relation.Nullary 14 | import Data.String as String 15 | 16 | infix 3 _#_ 17 | 18 | IsSet : ∀ {α} -> Set α -> Set α 19 | IsSet A = Decidable {A = A} _≡_ 20 | 21 | record Eq {α} (A : Set α) : Set α where 22 | infixl 5 _≟_ _==_ 23 | 24 | field _≟_ : IsSet A 25 | 26 | _==_ : A -> A -> Bool 27 | x == y = ⌊ x ≟ y ⌋ 28 | open Eq {{...}} public 29 | 30 | record _↦_ {α} (A B : Set α) : Set α where 31 | constructor packInj 32 | field 33 | to : A -> B 34 | from : B -> A 35 | from-to : from ∘ to ≗ id 36 | 37 | -- Can't make it an instance, because otherwise it unreasonably breaks instance search. 38 | viaInj : ∀ {α} {A B : Set α} {{bEq : Eq B}} -> A ↦ B -> Eq A 39 | viaInj {A = A} {B} inj = record 40 | { _≟_ = flip (via-injection {A = ≡-Setoid A} {B = ≡-Setoid B}) _≟_ $ record 41 | { to = record 42 | { _⟨$⟩_ = to 43 | ; cong = cong to 44 | } 45 | ; injective = λ q -> right (from-to _) (trans (cong from q) (from-to _)) 46 | } 47 | } where open _↦_ inj 48 | 49 | _#_ : ∀ {α} {A : Set α} -> A -> A -> Set 50 | x # y = Dec (x ≡ y) 51 | 52 | delim : ∀ {α π} {A : Set α} {P : Dec A -> Set π} 53 | -> (∀ x -> P (yes x)) -> (∀ c -> P (no c)) -> (d : Dec A) -> P d 54 | delim f g (yes x) = f x 55 | delim f g (no c) = g c 56 | 57 | drec : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> B) -> Dec A -> B 58 | drec = delim 59 | 60 | dmap : ∀ {α β} {A : Set α} {B : Set β} -> (A -> B) -> (¬ A -> ¬ B) -> Dec A -> Dec B 61 | dmap f g = drec (yes ∘ f) (no ∘ g) 62 | 63 | sumM2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 64 | -> (A -> Dec C) -> (B -> Dec C) -> (¬ A -> ¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 65 | sumM2 f g h d e = drec f (λ c -> drec g (h c) e) d 66 | 67 | prodM2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 68 | -> (A -> B -> Dec C) -> (¬ A -> Dec C) -> (¬ B -> Dec C) -> Dec A -> Dec B -> Dec C 69 | prodM2 h f g d e = drec (λ x -> drec (h x) g e) f d 70 | 71 | sumF2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 72 | -> (A -> C) -> (B -> C) -> (¬ A -> ¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 73 | sumF2 f g h = sumM2 (yes ∘ f) (yes ∘ g) (no % ∘ h) 74 | 75 | prodF2 : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 76 | -> (A -> B -> C) -> (¬ A -> ¬ C) -> (¬ B -> ¬ C) -> Dec A -> Dec B -> Dec C 77 | prodF2 h f g = prodM2 (yes % ∘ h) (no ∘ f) (no ∘ g) 78 | 79 | dcong : ∀ {α β} {A : Set α} {B : Set β} {x y} 80 | -> (f : A -> B) -> (f x ≡ f y -> x ≡ y) -> x # y -> f x # f y 81 | dcong f inj = dmap (cong f) (_∘ inj) 82 | 83 | dcong₂ : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 84 | -> (f : A -> B -> C) 85 | -> (f x₁ y₁ ≡ f x₂ y₂ -> x₁ ≡ x₂ × y₁ ≡ y₂) 86 | -> x₁ # x₂ 87 | -> y₁ # y₂ 88 | -> f x₁ y₁ # f x₂ y₂ 89 | dcong₂ f inj = prodF2 (cong₂ f) (λ c -> c ∘ proj₁ ∘ inj) (λ c -> c ∘ proj₂ ∘ inj) 90 | 91 | dhcong₂ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 92 | -> (f : ∀ x -> B x -> C) 93 | -> (f x₁ y₁ ≡ f x₂ y₂ -> [ B ] y₁ ≅ y₂) 94 | -> x₁ # x₂ 95 | -> (∀ y₂ -> y₁ # y₂) 96 | -> f x₁ y₁ # f x₂ y₂ 97 | dhcong₂ f inj (yes refl) q = dcong (f _) (homo ∘ inj) (q _) 98 | dhcong₂ f inj (no c) q = no (c ∘ inds ∘ inj) 99 | 100 | dsubst : ∀ {α β γ} {A : Set α} {x y} 101 | -> (B : A -> Set β) 102 | -> (C : ∀ {x} -> B x -> Set γ) 103 | -> x # y 104 | -> (z : B x) 105 | -> ((z : B y) -> C z) 106 | -> (x ≢ y -> C z) 107 | -> C z 108 | dsubst B C (yes refl) z g h = g z 109 | dsubst B C (no c) z g h = h c 110 | 111 | dsubst′ : ∀ {α β γ} {A : Set α} {C : Set γ} {x y} 112 | -> (B : A -> Set β) -> x # y -> B x -> (B y -> C) -> (x ≢ y -> C) -> C 113 | dsubst′ B = dsubst B _ 114 | 115 | ,-inj : ∀ {α β} {A : Set α} {B : A -> Set β} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂} 116 | -> (x₁ , y₁) ≡ (x₂ , y₂) -> [ B ] y₁ ≅ y₂ 117 | ,-inj refl = irefl 118 | 119 | inj₁-inj : ∀ {α β} {A : Set α} {B : Set β} {x₁ x₂ : A} 120 | -> inj₁ {B = B} x₁ ≡ inj₁ x₂ -> x₁ ≡ x₂ 121 | inj₁-inj refl = refl 122 | 123 | inj₂-inj : ∀ {α β} {A : Set α} {B : Set β} {y₁ y₂ : B} 124 | -> inj₂ {A = A} y₁ ≡ inj₂ y₂ -> y₁ ≡ y₂ 125 | inj₂-inj refl = refl 126 | 127 | -- _<,>ᵈ_ : ∀ {α β} {A : Set α} {B : Set β} {x₁ x₂ : A} {y₁ y₂ : B} 128 | -- -> x₁ # x₂ -> y₁ # y₂ -> x₁ , y₁ # x₂ , y₂ 129 | -- _<,>ᵈ_ = dcong₂ _,_ (inds-homo ∘ ,-inj) 130 | 131 | -- _<,>ᵈᵒ_ : ∀ {α β} {A : Set α} {B : A -> Set β} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂} 132 | -- -> x₁ # x₂ -> (∀ y₂ -> y₁ # y₂) -> x₁ , y₁ # x₂ , y₂ 133 | -- _<,>ᵈᵒ_ = dhcong₂ _,_ ,-inj 134 | 135 | decSum : ∀ {α β} {A : Set α} {B : Set β} 136 | -> IsSet A -> IsSet B -> IsSet (A ⊎ B) 137 | decSum f g (inj₁ x₁) (inj₁ x₂) = dcong inj₁ inj₁-inj (f x₁ x₂) 138 | decSum f g (inj₂ y₁) (inj₂ y₂) = dcong inj₂ inj₂-inj (g y₁ y₂) 139 | decSum f g (inj₁ x₁) (inj₂ y₂) = no λ() 140 | decSum f g (inj₂ y₁) (inj₁ x₂) = no λ() 141 | 142 | decProd : ∀ {α β} {A : Set α} {B : A -> Set β} 143 | -> IsSet A -> (∀ {x} -> IsSet (B x)) -> IsSet (Σ A B) 144 | decProd f g (x₁ , y₁) (x₂ , y₂) = dhcong₂ _,_ ,-inj (f x₁ x₂) (g y₁) 145 | 146 | module _ where 147 | import Relation.Binary.PropositionalEquality as B 148 | 149 | liftBase : ∀ {α} {A : Set α} {x y : A} -> x B.≡ y -> x ≡ y 150 | liftBase B.refl = refl 151 | 152 | lowerBase : ∀ {α} {A : Set α} {x y : A} -> x ≡ y -> x B.≡ y 153 | lowerBase refl = B.refl 154 | 155 | viaBase : ∀ {α} {A : Set α} -> Decidable (B._≡_ {A = A}) -> Eq A 156 | viaBase d = record 157 | { _≟_ = flip (via-injection {A = ≡-Setoid _} {B = B.setoid _}) d $ record 158 | { to = record 159 | { _⟨$⟩_ = id 160 | ; cong = lowerBase 161 | } 162 | ; injective = liftBase 163 | } 164 | } 165 | -------------------------------------------------------------------------------- /src/Generic/Lib/Equality/Coerce.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Equality.Coerce where 2 | 3 | open import Generic.Lib.Intro 4 | open import Generic.Lib.Equality.Propositional 5 | open import Generic.Lib.Decidable 6 | open import Generic.Lib.Data.Product 7 | 8 | Coerce′ : ∀ {α β} -> α ≡ β -> Set α -> Set β 9 | Coerce′ refl = id 10 | 11 | coerce′ : ∀ {α β} {A : Set α} -> (q : α ≡ β) -> A -> Coerce′ q A 12 | coerce′ refl = id 13 | 14 | uncoerce′ : ∀ {α β} {A : Set α} -> (q : α ≡ β) -> Coerce′ q A -> A 15 | uncoerce′ refl = id 16 | 17 | inspectUncoerce′ : ∀ {α β} {A : Set α} 18 | -> (q : α ≡ β) -> (p : Coerce′ q A) -> ∃ λ x -> p ≡ coerce′ q x 19 | inspectUncoerce′ refl x = x , refl 20 | 21 | split : ∀ {α β γ δ} {A : Set α} {B : A -> Set β} {C : Set γ} 22 | -> (q : α ⊔ β ≡ δ) -> Coerce′ q (Σ A B) -> (∀ x -> B x -> C) -> C 23 | split q p g = uncurry g (uncoerce′ q p) 24 | 25 | decCoerce′ : ∀ {α β} {A : Set α} -> (q : α ≡ β) -> IsSet A -> IsSet (Coerce′ q A) 26 | decCoerce′ refl = id 27 | 28 | data Coerce {β} : ∀ {α} -> α ≡ β -> Set α -> Set β where 29 | coerce : ∀ {A} -> A -> Coerce refl A 30 | 31 | qcoerce : ∀ {α β} {A : Set α} {q : α ≡ β} -> A -> Coerce q A 32 | qcoerce {q = refl} = coerce 33 | -------------------------------------------------------------------------------- /src/Generic/Lib/Equality/Congn.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Equality.Congn where 2 | 3 | open import Generic.Lib.Intro 4 | open import Generic.Lib.Equality.Propositional 5 | open import Generic.Lib.Data.Product 6 | open import Generic.Lib.Data.Nat 7 | open import Generic.Lib.Data.Pow 8 | open import Generic.Lib.Data.Sets 9 | 10 | zip≡ : ∀ {n αs} {As : Sets {n} αs} -> HList As -> HList As -> Sets (replicatePow n lzero) 11 | zip≡ {0} tt tt = tt 12 | zip≡ {suc _} (x , xs) (y , ys) = (x ≡ y) , zip≡ xs ys 13 | 14 | congn : ∀ {β} {B : Set β} n {αs} {As : Sets {n} αs} {xs ys : HList As} 15 | -> (f : FoldSets As B) -> FoldSets (zip≡ xs ys) (applyFoldSets f xs ≡ applyFoldSets f ys) 16 | congn 0 y = refl 17 | congn (suc n) f refl = congn n (f _) 18 | 19 | private 20 | module Test where 21 | cong′ : ∀ {α β} {A : Set α} {B : Set β} {x y} -> (f : A -> B) -> x ≡ y -> f x ≡ f y 22 | cong′ = congn 1 23 | 24 | cong₂′ : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 25 | -> (g : A -> B -> C) -> x₁ ≡ x₂ -> y₁ ≡ y₂ -> g x₁ y₁ ≡ g x₂ y₂ 26 | cong₂′ = congn 2 27 | -------------------------------------------------------------------------------- /src/Generic/Lib/Equality/Heteroindexed.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Equality.Heteroindexed where 2 | 3 | open import Data.Product 4 | 5 | open import Generic.Lib.Equality.Propositional 6 | 7 | data [_]_≅_ {ι α} {I : Set ι} {i} (A : I -> Set α) (x : A i) : ∀ {j} -> A j -> Set where 8 | irefl : [ A ] x ≅ x 9 | 10 | inds : ∀ {ι α} {I : Set ι} {A : I -> Set α} {i j} {x : A i} {y : A j} 11 | -> [ A ] x ≅ y -> i ≡ j 12 | inds irefl = refl 13 | 14 | homo : ∀ {ι α} {I : Set ι} {A : I -> Set α} {i} {x y : A i} 15 | -> [ A ] x ≅ y -> x ≡ y 16 | homo irefl = refl 17 | 18 | inds-homo : ∀ {ι α} {I : Set ι} {A : Set α} {i j : I} {x y : A} 19 | -> [_]_≅_ {i = i} (λ _ -> A) x {j} y -> i ≡ j × x ≡ y 20 | inds-homo irefl = refl , refl 21 | -------------------------------------------------------------------------------- /src/Generic/Lib/Equality/Propositional.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Equality.Propositional where 2 | 3 | open import Level 4 | open import Relation.Binary 5 | open import Data.Empty 6 | 7 | infix 3 _≡_ _≢_ _≗_ 8 | 9 | data _≡_ {α} {A : Set α} (x : A) : A -> Set where 10 | instance refl : x ≡ x 11 | 12 | pattern lrefl = lift refl 13 | 14 | _≢_ : ∀ {α} {A : Set α} -> A -> A -> Set 15 | x ≢ y = x ≡ y -> ⊥ 16 | 17 | _≗_ : ∀ {α β} {A : Set α} {B : A -> Set β} -> (∀ x -> B x) -> (∀ x -> B x) -> Set α 18 | f ≗ g = ∀ x -> f x ≡ g x 19 | 20 | sym : ∀ {α} {A : Set α} {x y : A} -> x ≡ y -> y ≡ x 21 | sym refl = refl 22 | 23 | trans : ∀ {α} {A : Set α} {x y z : A} -> x ≡ y -> y ≡ z -> x ≡ z 24 | trans refl refl = refl 25 | 26 | left : ∀ {α} {A : Set α} {x y z : A} -> y ≡ x -> z ≡ x -> y ≡ z 27 | left refl refl = refl 28 | 29 | right : ∀ {α} {A : Set α} {x y z : A} -> x ≡ y -> x ≡ z -> y ≡ z 30 | right refl refl = refl 31 | 32 | subst : ∀ {α β} {A : Set α} {x y} -> (B : A -> Set β) -> x ≡ y -> B x -> B y 33 | subst B refl z = z 34 | 35 | cong : ∀ {α β} {A : Set α} {B : Set β} {x y} -> (f : A -> B) -> x ≡ y -> f x ≡ f y 36 | cong f refl = refl 37 | 38 | cong₂ : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} {x₁ x₂ y₁ y₂} 39 | -> (g : A -> B -> C) -> x₁ ≡ x₂ -> y₁ ≡ y₂ -> g x₁ y₁ ≡ g x₂ y₂ 40 | cong₂ g refl refl = refl 41 | 42 | ≡-Setoid : ∀ {α} -> Set α -> Setoid α zero 43 | ≡-Setoid A = record 44 | { Carrier = A 45 | ; _≈_ = _≡_ 46 | ; isEquivalence = record 47 | { refl = refl 48 | ; sym = sym 49 | ; trans = trans 50 | } 51 | } 52 | -------------------------------------------------------------------------------- /src/Generic/Lib/Intro.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Intro where 2 | 3 | open import Level renaming (zero to lzero; suc to lsuc) public 4 | open import Function public 5 | open import Data.Bool.Base hiding (_<_; _≤_) public 6 | 7 | infixl 10 _% 8 | infixl 2 _>>>_ 9 | 10 | data ⊥ {α} : Set α where 11 | record ⊤ {α} : Set α where 12 | instance constructor tt 13 | 14 | ⊥₀ = ⊥ {lzero} 15 | ⊤₀ = ⊤ {lzero} 16 | 17 | tt₀ : ⊤₀ 18 | tt₀ = tt 19 | 20 | _% = _∘_ 21 | 22 | _>>>_ : ∀ {α β γ} {A : Set α} {B : A -> Set β} {C : ∀ {x} -> B x -> Set γ} 23 | -> (f : ∀ x -> B x) -> (∀ {x} -> (y : B x) -> C y) -> ∀ x -> C (f x) 24 | (f >>> g) x = g (f x) 25 | -------------------------------------------------------------------------------- /src/Generic/Lib/Prelude.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Prelude where 2 | 3 | open import Generic.Lib.Intro public 4 | open import Generic.Lib.Equality.Propositional public 5 | open import Generic.Lib.Equality.Coerce public 6 | open import Generic.Lib.Equality.Heteroindexed public 7 | open import Generic.Lib.Equality.Congn public 8 | open import Generic.Lib.Decidable public 9 | open import Generic.Lib.Category public 10 | open import Generic.Lib.Data.Nat public 11 | open import Generic.Lib.Data.String public 12 | open import Generic.Lib.Data.Maybe public 13 | open import Generic.Lib.Data.Sum public 14 | open import Generic.Lib.Data.Product public 15 | open import Generic.Lib.Data.List public 16 | open import Generic.Lib.Data.Pow public 17 | open import Generic.Lib.Data.Sets public 18 | open import Generic.Lib.Reflection.Core public 19 | open import Generic.Lib.Reflection.Fold public 20 | -------------------------------------------------------------------------------- /src/Generic/Lib/Reflection/Core.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Reflection.Core where 2 | 3 | open import Agda.Builtin.Reflection using (withNormalisation; Relevance; Visibility; clause) public 4 | open import Reflection 5 | renaming (visible to expl; hidden to impl; instance′ to inst; 6 | relevant to rel; irrelevant to irr; pi to absPi; lam to absLam; def to appDef) 7 | hiding (Arg-info; var; con; meta; visibility; relevance; _≟_; return; _>>=_; _>>_) public 8 | open import Reflection.Argument.Information using (ArgInfo; visibility; relevance) public 9 | import Reflection.Name 10 | open Term using () renaming (var to appVar; con to appCon; meta to appMeta) public 11 | open Pattern using () renaming (var to patVar; con to patCon) public 12 | open Literal using () renaming (meta to litMeta) public 13 | open Sort public 14 | 15 | open import Generic.Lib.Intro 16 | open import Generic.Lib.Equality.Propositional 17 | open import Generic.Lib.Decidable 18 | open import Generic.Lib.Category 19 | open import Generic.Lib.Data.Nat 20 | open import Generic.Lib.Data.String 21 | open import Generic.Lib.Data.Maybe 22 | open import Generic.Lib.Data.Product 23 | open import Generic.Lib.Data.List 24 | 25 | open import Data.Vec using (toList) 26 | open import Data.Vec.N-ary using (N-ary; curryⁿ) 27 | 28 | infixr 5 _‵→_ 29 | infixl 3 _·_ 30 | 31 | listCurryⁿ : ∀ {α β} {A : Set α} {B : Set β} n -> (List A -> B) -> N-ary n A B 32 | listCurryⁿ n f = curryⁿ {n = n} (f ∘ toList) 33 | 34 | named : String -> String 35 | named s = if s == "_" then "x" else s 36 | 37 | record Reify {α} (A : Set α) : Set α where 38 | field reify : A -> Term 39 | 40 | macro 41 | reflect : A -> Term -> TC _ 42 | reflect = unify ∘ reify 43 | open Reify {{...}} public 44 | 45 | pattern pureVar n = appVar n [] 46 | pattern pureCon c = appCon c [] 47 | pattern pureDef f = appDef f [] 48 | pattern pureMeta m = appMeta m [] 49 | 50 | {-# DISPLAY appVar i [] = pureVar i #-} 51 | {-# DISPLAY appCon c [] = pureCon c #-} 52 | {-# DISPLAY appDef f [] = pureDef f #-} 53 | {-# DISPLAY appMeta m [] = pureMeta m #-} 54 | 55 | pattern explInfo r = arg-info expl r 56 | pattern implInfo r = arg-info impl r 57 | pattern instInfo r = arg-info inst r 58 | 59 | {-# DISPLAY arg-info expl r = explInfo r #-} 60 | {-# DISPLAY arg-info impl r = implInfo r #-} 61 | {-# DISPLAY arg-info inst r = instInfo r #-} 62 | 63 | pattern explRelInfo = explInfo rel 64 | pattern explIrrInfo = explInfo irr 65 | pattern implRelInfo = implInfo rel 66 | pattern implIrrInfo = implInfo irr 67 | pattern instRelInfo = instInfo rel 68 | pattern instIrrInfo = instInfo irr 69 | 70 | {-# DISPLAY explInfo rel = explRelInfo #-} 71 | {-# DISPLAY explInfo irr = explIrrInfo #-} 72 | {-# DISPLAY implInfo rel = implRelInfo #-} 73 | {-# DISPLAY implInfo irr = implIrrInfo #-} 74 | {-# DISPLAY instInfo rel = instRelInfo #-} 75 | {-# DISPLAY instInfo irr = instIrrInfo #-} 76 | 77 | pattern explArg r x = arg (explInfo r) x 78 | pattern implArg r x = arg (implInfo r) x 79 | pattern instArg r x = arg (instInfo r) x 80 | 81 | {-# DISPLAY arg (explInfo r) = explArg r #-} 82 | {-# DISPLAY arg (implInfo r) = implArg r #-} 83 | {-# DISPLAY arg (instInfo r) = instArg r #-} 84 | 85 | pattern explRelArg x = explArg rel x 86 | pattern implRelArg x = implArg rel x 87 | pattern instRelArg x = instArg rel x 88 | 89 | {-# DISPLAY explArg rel x = explRelArg x #-} 90 | {-# DISPLAY implArg rel x = implRelArg x #-} 91 | {-# DISPLAY instArg rel x = instRelArg x #-} 92 | 93 | pattern pi s a b = absPi a (abs s b) 94 | 95 | {-# DISPLAY absPi a (abs s b) = pi s a b #-} 96 | 97 | pattern explPi r s a b = pi s (explArg r a) b 98 | pattern implPi r s a b = pi s (implArg r a) b 99 | pattern instPi r s a b = pi s (instArg r a) b 100 | 101 | {-# DISPLAY pi (explArg r a) s b = explPi r s a b #-} 102 | {-# DISPLAY pi (implArg r a) s b = implPi r s a b #-} 103 | {-# DISPLAY pi (instArg r a) s b = instPi r s a b #-} 104 | 105 | pattern explRelPi s a b = explPi rel a s b 106 | pattern explIrrPi s a b = explPi irr a s b 107 | pattern implRelPi s a b = implPi rel a s b 108 | pattern implIrrPi s a b = implPi irr a s b 109 | pattern instRelPi s a b = instPi rel a s b 110 | pattern instIrrPi s a b = instPi irr a s b 111 | 112 | {-# DISPLAY explPi rel a s b = explRelPi s a b #-} 113 | {-# DISPLAY explPi irr a s b = explIrrPi s a b #-} 114 | {-# DISPLAY implPi rel a s b = implRelPi s a b #-} 115 | {-# DISPLAY implPi irr a s b = implIrrPi s a b #-} 116 | {-# DISPLAY instPi rel a s b = instRelPi s a b #-} 117 | {-# DISPLAY instPi irr a s b = instIrrPi s a b #-} 118 | 119 | pattern lam v s t = absLam v (abs s t) 120 | 121 | {-# DISPLAY absLam v (abs s t) = lam v s t #-} 122 | 123 | pattern explLam s t = lam expl s t 124 | pattern implLam s t = lam impl s t 125 | pattern instLam s t = lam inst s t 126 | 127 | {-# DISPLAY lam expl s t = explLam s t #-} 128 | {-# DISPLAY lam impl s t = implLam s t #-} 129 | {-# DISPLAY lam inst s t = instLam s t #-} 130 | 131 | pattern _‵→_ a b = pi "_" (explRelArg a) b 132 | 133 | -- No longer parses for whatever reason. 134 | -- {-# DISPLAY pi "_" (explRelArg a) b = a ‵→ b #-} 135 | 136 | mutual 137 | <_>_ : ∀ {α} -> Relevance -> Set α -> Set α 138 | <_>_ = flip RelValue 139 | 140 | data RelValue {α} (A : Set α) : Relevance -> Set α where 141 | relv : A -> < rel > A 142 | irrv : .A -> < irr > A 143 | 144 | elimRelValue : ∀ {r α π} {A : Set α} 145 | -> (P : ∀ {r} -> < r > A -> Set π) 146 | -> (∀ x -> P (relv x)) 147 | -> (∀ .x -> P (irrv x)) 148 | -> (x : < r > A) 149 | -> P x 150 | elimRelValue P f g (relv x) = f x 151 | elimRelValue P f g (irrv x) = g x 152 | 153 | unrelv : ∀ {α} {A : Set α} -> < rel > A -> A 154 | unrelv (relv x) = x 155 | 156 | -- Is it possible to handle this in some other way that doesn't require a postulate? 157 | -- See the `appRel` function below. Or is the postulate fine? 158 | postulate 159 | .unirrv : ∀ {α} {A : Set α} -> < irr > A -> A 160 | 161 | <_>_~>_ : ∀ {α β} -> Relevance -> Set α -> Set β -> Set (α ⊔ β) 162 | < rel > A ~> B = A -> B 163 | < irr > A ~> B = .A -> B 164 | 165 | lamRel : ∀ {r α β} {A : Set α} {B : Set β} -> (< r > A -> B) -> < r > A ~> B 166 | lamRel {rel} f = λ x -> f (relv x) 167 | lamRel {irr} f = λ x -> f (irrv x) 168 | 169 | -- The laziness is intentional. 170 | appRel : ∀ {r α β} {A : Set α} {B : Set β} -> (< r > A ~> B) -> < r > A -> B 171 | appRel {rel} f rx = f (unrelv rx) 172 | appRel {irr} f rx = f (unirrv rx) 173 | 174 | Pi : ∀ {α β} i -> (A : Set α) -> (< relevance i > A -> Set β) -> Set (α ⊔ β) 175 | Pi explRelInfo A B = (x : A) -> B (relv x) 176 | Pi explIrrInfo A B = . (x : A) -> B (irrv x) 177 | Pi implRelInfo A B = {x : A} -> B (relv x) 178 | Pi implIrrInfo A B = . {x : A} -> B (irrv x) 179 | Pi instRelInfo A B = {{x : A}} -> B (relv x) 180 | Pi instIrrInfo A B = .{{x : A}} -> B (irrv x) 181 | 182 | lamPi : ∀ {α β} {A : Set α} i {B : < relevance i > A -> Set β} -> (∀ x -> B x) -> Pi i A B 183 | lamPi explRelInfo f = λ x -> f (relv x) 184 | lamPi explIrrInfo f = λ x -> f (irrv x) 185 | lamPi implRelInfo f = f _ 186 | lamPi implIrrInfo f = f _ 187 | lamPi instRelInfo f = f _ 188 | lamPi instIrrInfo f = f _ 189 | 190 | appPi : ∀ {α β} {A : Set α} i {B : < relevance i > A -> Set β} -> Pi i A B -> ∀ x -> B x 191 | appPi explRelInfo f (relv x) = f x 192 | appPi explIrrInfo f (irrv x) = f x 193 | appPi implRelInfo y (relv x) = y 194 | appPi implIrrInfo y (irrv x) = y 195 | appPi instRelInfo y (relv x) = y {{x}} 196 | appPi instIrrInfo y (irrv x) = y {{x}} 197 | 198 | RelEq : ∀ {α} -> Relevance -> Set α -> Set α 199 | RelEq rel A = Eq A 200 | RelEq irr A = ⊤ 201 | 202 | vis : {A : Set} -> (A -> List (Arg Term) -> Term) -> A -> List Term -> Term 203 | vis k x = k x ∘ map explRelArg 204 | 205 | vis# : ∀ {A : Set} n -> (A -> List (Arg Term) -> Term) -> A -> N-ary n Term Term 206 | vis# n k = listCurryⁿ n ∘ vis k 207 | 208 | isRelevant : Relevance -> Bool 209 | isRelevant rel = true 210 | isRelevant irr = false 211 | 212 | argInfo : ∀ {α} {A : Set α} -> Arg A -> _ 213 | argInfo (arg i x) = i 214 | 215 | argVal : ∀ {α} {A : Set α} -> Arg A -> A 216 | argVal (arg i x) = x 217 | 218 | unExpl : ∀ {α} {A : Set α} -> Arg A -> Maybe A 219 | unExpl (explArg r x) = just x 220 | unExpl _ = nothing 221 | 222 | absName : ∀ {α} {A : Set α} -> Abs A -> String 223 | absName (abs s x) = s 224 | 225 | absVal : ∀ {α} {A : Set α} -> Abs A -> A 226 | absVal (abs s x) = x 227 | 228 | patVars : List String -> List (Arg Pattern) 229 | patVars = map (explRelArg ∘ patVar ∘ named) 230 | 231 | record Data {α} (A : Set α) : Set α where 232 | no-eta-equality 233 | constructor packData 234 | field 235 | dataName : Name 236 | parsTele : Type 237 | indsTele : Type 238 | consTypes : List A 239 | consNames : All (const Name) consTypes 240 | open Data public 241 | 242 | instance 243 | NameEq : Eq Name 244 | NameEq = viaBase Reflection.Name._≟_ 245 | 246 | EqRelValue : ∀ {α r} {A : Set α} {{aEq : RelEq r A}} -> Eq (< r > A) 247 | EqRelValue {A = A} {{aEq}} = record 248 | { _≟_ = go 249 | } where 250 | relv-inj : {x y : A} -> relv x ≡ relv y -> x ≡ y 251 | relv-inj refl = refl 252 | 253 | go : ∀ {r} {{aEq : RelEq r A}} -> IsSet (< r > A) 254 | go (relv x) (relv y) = dcong relv relv-inj (x ≟ y) 255 | go (irrv x) (irrv y) = yes refl 256 | 257 | ArgFunctor : ∀ {α} -> RawFunctor {α} Arg 258 | ArgFunctor = record 259 | { _<$>_ = λ{ f (arg i x) -> arg i (f x) } 260 | } 261 | 262 | AbsFunctor : ∀ {α} -> RawFunctor {α} Abs 263 | AbsFunctor = record 264 | { _<$>_ = λ{ f (abs s x) -> abs s (f x) } 265 | } 266 | 267 | TCMonad : ∀ {α} -> RawMonad {α} TC 268 | TCMonad = record 269 | { return = Reflection.return 270 | ; _>>=_ = Reflection._>>=_ 271 | } 272 | 273 | TCApplicative : ∀ {α} -> RawApplicative {α} TC 274 | TCApplicative = rawIApplicative 275 | 276 | TCFunctor : ∀ {α} -> RawFunctor {α} TC 277 | TCFunctor = rawFunctor 278 | 279 | keep : (ℕ -> ℕ) -> ℕ -> ℕ 280 | keep ι 0 = 0 281 | keep ι (suc n) = suc (ι n) 282 | 283 | {-# TERMINATING #-} 284 | mutual 285 | ren : (ℕ -> ℕ) -> Term -> Term 286 | ren ι (appVar v xs) = appVar (ι v) (rens ι xs) 287 | ren ι (appCon c xs) = appCon c (rens ι xs) 288 | ren ι (appDef f xs) = appDef f (rens ι xs) 289 | ren ι (lam v s t) = lam v s (ren (keep ι) t) 290 | ren ι (pat-lam cs xs) = undefined where postulate undefined : _ 291 | ren ι (pi s a b) = pi s (ren ι <$> a) (ren (keep ι) b) 292 | ren ι (agda-sort s) = agda-sort (renSort ι s) 293 | ren ι (lit l) = lit l 294 | ren ι (appMeta x xs) = appMeta x (rens ι xs) 295 | ren ι unknown = unknown 296 | 297 | rens : (ℕ -> ℕ) -> List (Arg Term) -> List (Arg Term) 298 | rens ι = map (fmap (ren ι)) 299 | 300 | renSort : (ℕ -> ℕ) -> Sort -> Sort 301 | renSort ι (set t) = set (ren ι t) 302 | renSort ι (lit n) = lit n 303 | renSort ι unknown = unknown 304 | 305 | shiftBy : ℕ -> Term -> Term 306 | shiftBy = ren ∘ _+_ 307 | 308 | shift : Term -> Term 309 | shift = shiftBy 1 310 | 311 | unshiftBy : ℕ -> Term -> Term 312 | unshiftBy n = ren (_∸ n) 313 | 314 | isSomeName : Name -> Term -> Bool 315 | isSomeName n (appDef m _) = n == m 316 | isSomeName n (appCon m _) = n == m 317 | isSomeName n t = false 318 | 319 | {-# TERMINATING #-} 320 | mutual 321 | mapName : (ℕ -> List (Arg Term) -> Term) -> Name -> Term -> Term 322 | mapName f n (appVar v xs) = appVar v (mapNames f n xs) 323 | mapName f n (appCon m xs) = (if n == m then f 0 else appCon m) (mapNames f n xs) 324 | mapName f n (appDef m xs) = (if n == m then f 0 else appDef m) (mapNames f n xs) 325 | mapName f n (lam v s t) = lam v s (mapName (f ∘ suc) n t) 326 | mapName f n (pat-lam cs xs) = undefined where postulate undefined : _ 327 | mapName f n (pi s a b) = pi s (mapName f n <$> a) (mapName (f ∘ suc) n b) 328 | mapName f n (agda-sort s) = agda-sort (mapNameSort f n s) 329 | mapName f n (lit l) = lit l 330 | mapName f n (appMeta x xs) = appMeta x (mapNames f n xs) 331 | mapName f n unknown = unknown 332 | 333 | mapNames : (ℕ -> List (Arg Term) -> Term) -> Name -> List (Arg Term) -> List (Arg Term) 334 | mapNames f n = map (fmap (mapName f n)) 335 | 336 | mapNameSort : (ℕ -> List (Arg Term) -> Term) -> Name -> Sort -> Sort 337 | mapNameSort f n (set t) = set (mapName f n t) 338 | mapNameSort f n (lit l) = lit l 339 | mapNameSort f n unknown = unknown 340 | 341 | explsOnly : List (Arg Term) -> List Term 342 | explsOnly = mapMaybe unExpl 343 | 344 | initType : Type -> Type 345 | initType (pi s a b) = pi s a (initType b) 346 | initType b = unknown 347 | 348 | lastType : Type -> Type 349 | lastType (pi s a b) = lastType b 350 | lastType b = b 351 | 352 | -- These two should return just `Type` like everything else. 353 | takePis : ℕ -> Type -> Maybe Type 354 | takePis 0 a = just unknown 355 | takePis (suc n) (pi s a b) = pi s a <$> takePis n b 356 | takePis _ _ = nothing 357 | 358 | dropPis : ℕ -> Type -> Maybe Type 359 | dropPis 0 a = just a 360 | dropPis (suc n) (pi s a b) = dropPis n b 361 | dropPis _ _ = nothing 362 | 363 | monoLastType : Type -> Type 364 | monoLastType = go 0 where 365 | go : ℕ -> Type -> Type 366 | go n (pi s a b) = go (suc n) b 367 | go n b = unshiftBy n b 368 | 369 | appendType : Type -> Type -> Type 370 | appendType (pi s a b) c = pi s a (appendType b c) 371 | appendType b c = c 372 | 373 | explLamsBy : Type -> Term -> Term 374 | explLamsBy (explPi r s a b) t = explLam s (explLamsBy b t) 375 | explLamsBy (pi s a b) t = explLamsBy b t 376 | explLamsBy b t = t 377 | 378 | implicitize : Type -> Type 379 | implicitize (explPi r s a b) = implPi r s a (implicitize b) 380 | implicitize (pi s a b) = pi s a (implicitize b) 381 | implicitize b = b 382 | 383 | leadImpls : Type -> List (Abs Term) 384 | leadImpls (implPi r s a b) = abs s a ∷ leadImpls b 385 | leadImpls b = [] 386 | 387 | pisToAbsArgTypes : Type -> List (Abs (Arg Type)) 388 | pisToAbsArgTypes (pi s a b) = abs s a ∷ pisToAbsArgTypes b 389 | pisToAbsArgTypes b = [] 390 | 391 | explPisToAbsTypes : Type -> List (Abs Type) 392 | explPisToAbsTypes (explPi r s a b) = abs s a ∷ explPisToAbsTypes b 393 | explPisToAbsTypes (pi s a b) = explPisToAbsTypes b 394 | explPisToAbsTypes b = [] 395 | 396 | explPisToNames : Type -> List String 397 | explPisToNames = map absName ∘ explPisToAbsTypes 398 | 399 | countPis : Type -> ℕ 400 | countPis = length ∘ pisToAbsArgTypes 401 | 402 | countExplPis : Type -> ℕ 403 | countExplPis = length ∘ explPisToAbsTypes 404 | 405 | pisToAbsArgVars : ℕ -> Type -> List (Abs (Arg Term)) 406 | pisToAbsArgVars (suc n) (pi s (arg i a) b) = abs s (arg i (pureVar n)) ∷ pisToAbsArgVars n b 407 | pisToAbsArgVars n b = [] 408 | 409 | pisToArgVars : ℕ -> Type -> List (Arg Term) 410 | pisToArgVars = map absVal % ∘ pisToAbsArgVars 411 | 412 | explPisToAbsVars : ℕ -> Type -> List (Abs Term) 413 | explPisToAbsVars (suc n) (explPi r s a b) = abs s (pureVar n) ∷ explPisToAbsVars n b 414 | explPisToAbsVars (suc n) (pi s a b) = explPisToAbsVars n b 415 | explPisToAbsVars n b = [] 416 | 417 | throw : ∀ {α} {A : Set α} -> String -> TC A 418 | throw s = typeError (strErr s ∷ []) 419 | 420 | panic : ∀ {α} {A : Set α} -> String -> TC A 421 | panic s = throw $ "panic: " ++ˢ s 422 | 423 | -- I'll merge these later. 424 | macro 425 | sate : Name -> Term -> TC _ 426 | sate f ?r = 427 | getType f >>= λ a -> 428 | let res = λ app -> quoteTC (vis# (countExplPis a) app f) >>= unify ?r in 429 | getDefinition f >>= λ 430 | { (data-cons _) -> res appCon 431 | ; _ -> res appDef 432 | } 433 | 434 | sateMacro : Name -> Term -> TC _ 435 | sateMacro f ?r = 436 | getType f >>= λ a -> 437 | quoteTC (vis# (pred (countExplPis a)) appDef f) >>= unify ?r 438 | 439 | _·_ : Term -> Term -> Term 440 | _·_ = sate _$_ 441 | 442 | unshift′ : Term -> Term 443 | unshift′ t = explLam "_" t · sate tt₀ 444 | 445 | -- A note for myself: `foldℕ (sate lsuc) (sate lzero) n` is not `reify n`: 446 | -- it's damn `lsuc` -- not `suc`. 447 | termLevelOf : Term -> Maybe Term 448 | termLevelOf (agda-sort (set t)) = just t 449 | termLevelOf (agda-sort (lit n)) = just (foldℕ (sate lsuc) (sate lzero) n) 450 | termLevelOf (agda-sort unknown) = just unknown 451 | termLevelOf _ = nothing 452 | 453 | instance 454 | TermReify : Reify Term 455 | TermReify = record 456 | { reify = id 457 | } 458 | 459 | NameReify : Reify Name 460 | NameReify = record 461 | { reify = lit ∘′ name 462 | } 463 | 464 | VisibilityReify : Reify Visibility 465 | VisibilityReify = record 466 | { reify = λ 467 | { expl -> sate expl 468 | ; impl -> sate impl 469 | ; inst -> sate inst 470 | } 471 | } 472 | 473 | RelevanceReify : Reify Relevance 474 | RelevanceReify = record 475 | { reify = λ 476 | { rel -> sate rel 477 | ; irr -> sate irr 478 | } 479 | } 480 | 481 | ArgInfoReify : Reify ArgInfo 482 | ArgInfoReify = record 483 | { reify = λ{ (arg-info v r) -> sate arg-info (reify v) (reify r) } 484 | } 485 | 486 | ProdReify : ∀ {α β} {A : Set α} {B : A -> Set β} 487 | {{aReify : Reify A}} {{bReify : ∀ {x} -> Reify (B x)}} -> Reify (Σ A B) 488 | ProdReify = record 489 | { reify = uncurry λ x y -> sate _,_ (reify x) (reify y) 490 | } 491 | 492 | ℕReify : Reify ℕ 493 | ℕReify = record 494 | { reify = foldℕ (sate suc) (sate zero) 495 | } 496 | 497 | ListReify : ∀ {α} {A : Set α} {{aReify : Reify A}} -> Reify (List A) 498 | ListReify = record 499 | { reify = foldr (sate _∷_ ∘ reify) (sate []) 500 | } 501 | 502 | AllReify : ∀ {α β} {A : Set α} {B : A -> Set β} {xs} {{bReify : ∀ {x} -> Reify (B x)}} 503 | -> Reify (All B xs) 504 | AllReify {B = B} {{bReify}} = record 505 | { reify = go _ 506 | } where 507 | go : ∀ xs -> All B xs -> Term 508 | go [] tt = sate tt₀ 509 | go (x ∷ xs) (y , ys) = sate _,_ (reify {{bReify}} y) (go xs ys) 510 | 511 | toTuple : List Term -> Term 512 | toTuple = foldr₁ (sate _,_) (sate tt₀) 513 | 514 | curryBy : Type -> Term -> Term 515 | curryBy = go 0 where 516 | go : ℕ -> Type -> Term -> Term 517 | go n (pi s (arg (arg-info v r) a) b) t = lam v s $ go (suc n) b t 518 | go n _ t = shiftBy n t · toTuple (map pureVar (downFrom n)) 519 | 520 | explUncurryBy : Type -> Term -> Term 521 | explUncurryBy a f = explLam "x" $ appDef (quote id) (explArg rel (shift f) ∷ go a (pureVar 0)) where 522 | go : Term -> Term -> List (Arg Term) 523 | go (explPi r s a b@(pi _ _ _)) p = explArg r (sate proj₁ p) ∷ go b (sate proj₂ p) 524 | go (pi s a b@(pi _ _ _)) p = go b (sate proj₂ p) 525 | go (explPi r s a b) x = explArg r x ∷ [] 526 | go _ t = [] 527 | 528 | defineTerm : Name -> Term -> TC _ 529 | defineTerm n t = 530 | getType n >>= λ a -> 531 | defineFun n (clause (map (implRelArg ∘ patVar ∘ named ∘ absName) (leadImpls a)) t ∷ []) 532 | 533 | -- Able to normalize a Setω. 534 | normalize : Term -> TC Term 535 | normalize (pi s (arg i a) b) = 536 | pi s ∘ arg i <$> normalize a <*> extendContext (arg i a) (normalize b) 537 | normalize t = normalise t 538 | 539 | getNormType : Name -> TC Type 540 | getNormType = getType >=> normalize 541 | 542 | inferNormType : Term -> TC Type 543 | inferNormType = inferType >=> normalize 544 | 545 | getData : Name -> TC (Data Type) 546 | getData d = getNormType d >>= λ ab -> getDefinition d >>= λ 547 | { (data-type p cs) -> 548 | mapM (λ c -> _,_ c ∘ dropPis p <$> getNormType c) cs >>= λ mans -> 549 | case takePis p ab ⊗ (dropPis p ab ⊗ (mapM (uncurry λ c ma -> flip _,_ c <$> ma) mans)) of λ 550 | { nothing -> panic "getData: data" 551 | ; (just (a , b , acs)) -> return ∘ uncurry (packData d a b) $ splitList acs 552 | } 553 | ; (record-type c _) -> getNormType c >>= dropPis (countPis ab) >>> λ 554 | { nothing -> panic "getData: record" 555 | ; (just a′) -> return $ packData d (initType ab) (lastType ab) (a′ ∷ []) (c , tt) 556 | } 557 | ; _ -> throw "not a data" 558 | } 559 | 560 | macro 561 | TypeOf : Term -> Term -> TC _ 562 | TypeOf t ?r = inferNormType t >>= unify ?r 563 | 564 | runTC : ∀ {α} {A : Set α} -> TC A -> Term -> TC _ 565 | runTC a ?r = bindTC a quoteTC >>= unify ?r 566 | -------------------------------------------------------------------------------- /src/Generic/Lib/Reflection/Fold.agda: -------------------------------------------------------------------------------- 1 | module Generic.Lib.Reflection.Fold where 2 | 3 | open import Generic.Lib.Intro 4 | open import Generic.Lib.Decidable 5 | open import Generic.Lib.Category 6 | open import Generic.Lib.Data.Nat 7 | open import Generic.Lib.Data.String 8 | open import Generic.Lib.Data.Maybe 9 | open import Generic.Lib.Data.List 10 | open import Generic.Lib.Reflection.Core 11 | 12 | foldTypeOf : Data Type -> Type 13 | foldTypeOf (packData d a b cs ns) = appendType iab ∘ pi "π" π ∘ pi "P" P $ cs ʰ→ from ‵→ to where 14 | infixr 5 _ʰ→_ 15 | i = countPis a 16 | j = countPis b 17 | ab = appendType a b 18 | iab = implicitize ab 19 | π = implRelArg (quoteTerm Level) 20 | P = explRelArg ∘ appendType (shiftBy (suc j) b) ∘ agda-sort ∘ set $ pureVar j 21 | hyp = mapName (λ p -> appVar p ∘ drop i) d ∘ shiftBy (2 + j) 22 | _ʰ→_ = flip $ foldr (λ c r -> hyp c ‵→ shift r) 23 | from = appDef d (pisToArgVars (2 + i + j) ab) 24 | to = appVar 1 (pisToArgVars (3 + j) b) 25 | 26 | foldClausesOf : Name -> Data Type -> List Clause 27 | foldClausesOf f (packData d a b cs ns) = allToList $ mapAllInd (λ {a} n -> clauseOf n a) ns where 28 | k = length cs 29 | 30 | clauseOf : ℕ -> Type -> Name -> Clause 31 | clauseOf i c n = clause lhs rhs where 32 | args = explPisToNames c 33 | j = length args 34 | lhs = patVars ("P" ∷ unmap (λ n -> "f" ++ˢ showName n) ns) 35 | ∷ʳ explRelArg (patCon n (patVars args)) 36 | 37 | tryHyp : ℕ -> ℕ -> Type -> Maybe Term 38 | tryHyp m l (explPi r s a b) = explLam s <$> tryHyp (suc m) l b 39 | tryHyp m l (pi s a b) = tryHyp m l b 40 | tryHyp m l (appDef e _) = d == e ?> vis appDef f (pars ∷ʳ rarg) where 41 | pars = map (λ i -> pureVar (m + i)) $ downFromTo (suc k + j) j 42 | rarg = vis appVar (m + l) $ map pureVar (downFrom m) 43 | tryHyp m l b = nothing 44 | 45 | hyps : ℕ -> Type -> List Term 46 | hyps (suc l) (explPi r s a b) = fromMaybe (pureVar l) (tryHyp 0 l a) ∷ hyps l b 47 | hyps l (pi s a b) = hyps l b 48 | hyps l b = [] 49 | 50 | rhs = vis appVar (k + j ∸ suc i) (hyps j c) 51 | 52 | -- i = 1: f 53 | -- j = 3: x t r 54 | -- k = 2: g f 55 | -- 5 4 3 2 1 0 3 2 4 3 1 1 0 6 5 2 1 0 56 | -- fold P g f (cons x t r) = f x (fold g f t) (λ y u -> fold g f (r y u)) 57 | 58 | defineFold : Name -> Data Type -> TC _ 59 | defineFold f D = 60 | declareDef (explRelArg f) (foldTypeOf D) >> 61 | defineFun f (foldClausesOf f D) 62 | 63 | deriveFold : Name -> Data Type -> TC Name 64 | deriveFold d D = 65 | freshName ("fold" ++ˢ showName d) >>= λ f -> 66 | f <$ defineFold f D 67 | 68 | deriveFoldTo : Name -> Name -> TC _ 69 | deriveFoldTo f = getData >=> defineFold f 70 | -------------------------------------------------------------------------------- /src/Generic/Main.agda: -------------------------------------------------------------------------------- 1 | module Generic.Main where 2 | 3 | open import Generic.Core public 4 | open import Generic.Function.FoldMono public 5 | open import Generic.Property.Eq public 6 | open import Generic.Reflection.ReadData public 7 | open import Generic.Reflection.DeriveEq public 8 | -------------------------------------------------------------------------------- /src/Generic/Property/Eq.agda: -------------------------------------------------------------------------------- 1 | module Generic.Property.Eq where 2 | 3 | open import Generic.Core 4 | 5 | SemEq : ∀ {i β} {I : Set i} -> Desc I β -> Set 6 | SemEq (var i) = ⊤ 7 | SemEq (π i q C) = ⊥ 8 | SemEq (D ⊛ E) = SemEq D × SemEq E 9 | 10 | mutual 11 | ExtendEq : ∀ {i β} {I : Set i} -> Desc I β -> Set β 12 | ExtendEq (var i) = ⊤ 13 | ExtendEq (π i q C) = ExtendEqᵇ i C q 14 | ExtendEq (D ⊛ E) = SemEq D × ExtendEq E 15 | 16 | ExtendEqᵇ : ∀ {ι α β γ q} {I : Set ι} i -> Binder α β γ i q I -> α ≤ℓ β -> Set β 17 | ExtendEqᵇ (arg-info v r) (coerce (A , D)) q = Coerce′ q $ RelEq r A × ∀ {x} -> ExtendEq (D x) 18 | 19 | instance 20 | {-# TERMINATING #-} -- Why? 21 | DataEq : ∀ {i β} {I : Set i} {D : Data (Desc I β)} {j} 22 | {{eqD : All ExtendEq (consTypes D)}} -> Eq (μ D j) 23 | DataEq {ι} {β} {I} {D₀} = record { _≟_ = decMu } where 24 | mutual 25 | decSem : ∀ D {{eqD : SemEq D}} -> IsSet (⟦ D ⟧ (μ D₀)) 26 | decSem (var i) d₁ d₂ = decMu d₁ d₂ 27 | decSem (π i q C) {{()}} 28 | decSem (D ⊛ E) {{eqD , eqE}} p₁ p₂ = 29 | decProd (decSem D {{eqD}}) (decSem E {{eqE}}) p₁ p₂ 30 | 31 | decExtend : ∀ {j} D {{eqD : ExtendEq D}} -> IsSet (Extend D (μ D₀) j) 32 | decExtend (var i) lrefl lrefl = yes refl 33 | decExtend (π i q C) p₁ p₂ = decExtendᵇ i C q p₁ p₂ 34 | decExtend (D ⊛ E) {{eqD , eqE}} p₁ p₂ = 35 | decProd (decSem D {{eqD}}) (decExtend E {{eqE}}) p₁ p₂ 36 | 37 | decExtendᵇ : ∀ {α γ q j} i (C : Binder α β γ i q I) q {{eqC : ExtendEqᵇ i C q}} 38 | -> IsSet (Extendᵇ i C q (μ D₀) j) 39 | decExtendᵇ (arg-info v r) (coerce (A , D)) q {{eqC}} p₁ p₂ = 40 | split q eqC λ eqA eqD -> 41 | decCoerce′ q (decProd (_≟_ {{EqRelValue {{eqA}}}}) (decExtend (D _) {{eqD}})) p₁ p₂ 42 | 43 | decAny : ∀ {j} (Ds : List (Desc I β)) {{eqD : All ExtendEq Ds}} 44 | -> ∀ d a b ns -> IsSet (Node D₀ (packData d a b Ds ns) j) 45 | decAny [] d a b tt () () 46 | decAny (D ∷ []) {{eqD , _}} d a b (_ , ns) e₁ e₂ = decExtend D {{eqD}} e₁ e₂ 47 | decAny (D ∷ E ∷ Ds) {{eqD , eqDs}} d a b (_ , ns) s₁ s₂ = 48 | decSum (decExtend D {{eqD}}) (decAny (E ∷ Ds) {{eqDs}} d a b ns) s₁ s₂ 49 | 50 | decMu : ∀ {j} -> IsSet (μ D₀ j) 51 | decMu (node e₁) (node e₂) = dcong node node-inj $ decAny (consTypes D₀) 52 | (dataName D₀) 53 | (parsTele D₀) 54 | (indsTele D₀) 55 | (consNames D₀) 56 | e₁ 57 | e₂ 58 | -------------------------------------------------------------------------------- /src/Generic/Property/Reify.agda: -------------------------------------------------------------------------------- 1 | module Generic.Property.Reify where 2 | 3 | open import Generic.Core 4 | 5 | data ExplView : Visibility -> Set where 6 | yes-expl : ExplView expl 7 | no-expl : ∀ {v} -> ExplView v 8 | 9 | explView : ∀ v -> ExplView v 10 | explView expl = yes-expl 11 | explView v = no-expl 12 | 13 | ExplMaybe : ∀ {α} -> Visibility -> Set α -> Set α 14 | ExplMaybe v A with explView v 15 | ... | yes-expl = A 16 | ... | no-expl = ⊤ 17 | 18 | caseExplMaybe : ∀ {α π} {A : Set α} (P : Visibility -> Set π) 19 | -> (A -> P expl) -> (∀ {v} -> P v) -> ∀ v -> ExplMaybe v A -> P v 20 | caseExplMaybe P f y v x with explView v 21 | ... | yes-expl = f x 22 | ... | no-expl = y 23 | 24 | Prod : ∀ {α} β -> Visibility -> Set α -> Set (α ⊔ β) -> Set (α ⊔ β) 25 | Prod β v A B with explView v 26 | ... | yes-expl = A × B 27 | ... | no-expl = B 28 | 29 | uncurryProd : ∀ {α β γ} {A : Set α} {B : Set (α ⊔ β)} {C : Set γ} v 30 | -> Prod β v A B -> (ExplMaybe v A -> B -> C) -> C 31 | uncurryProd v p g with explView v | p 32 | ... | yes-expl | (x , y) = g x y 33 | ... | no-expl | y = g tt y 34 | 35 | SemReify : ∀ {i β} {I : Set i} -> Desc I β -> Set 36 | SemReify (var i) = ⊤ 37 | SemReify (π i q C) = ⊥ 38 | SemReify (D ⊛ E) = SemReify D × SemReify E 39 | 40 | mutual 41 | ExtendReify : ∀ {i β} {I : Set i} -> Desc I β -> Set β 42 | ExtendReify (var i) = ⊤ 43 | ExtendReify (π i q C) = ExtendReifyᵇ i C q 44 | ExtendReify (D ⊛ E) = SemReify D × ExtendReify E 45 | 46 | ExtendReifyᵇ : ∀ {ι α β γ q} {I : Set ι} i -> Binder α β γ i q I -> α ≤ℓ β -> Set β 47 | ExtendReifyᵇ {β = β} i (coerce (A , D)) q = Coerce′ q $ 48 | Prod β (visibility i) (Reify A) (∀ {x} -> ExtendReify (D x)) 49 | 50 | -- Can't reify an irrelevant thing into its relevant representation. 51 | postulate 52 | reifyᵢ : ∀ {α} {A : Set α} {{aReify : Reify A}} -> .A -> Term 53 | 54 | instance 55 | {-# TERMINATING #-} 56 | DataReify : ∀ {i β} {I : Set i} {D : Data (Desc I β)} {j} 57 | {{reD : All ExtendReify (consTypes D)}} -> Reify (μ D j) 58 | DataReify {ι} {β} {I} {D₀} = record { reify = reifyMu } where 59 | mutual 60 | reifySem : ∀ D {{reD : SemReify D}} -> ⟦ D ⟧ (μ D₀) -> Term 61 | reifySem (var i) d = reifyMu d 62 | reifySem (π i q C) {{()}} 63 | reifySem (D ⊛ E) {{reD , reE}} (x , y) = 64 | sate _,_ (reifySem D {{reD}} x) (reifySem E {{reE}} y) 65 | 66 | reifyExtend : ∀ {j} D {{reD : ExtendReify D}} -> Extend D (μ D₀) j -> List Term 67 | reifyExtend (var i) lrefl = [] 68 | reifyExtend (π i q C) p = reifyExtendᵇ i C q p 69 | reifyExtend (D ⊛ E) {{reD , reE}} (x , e) = 70 | reifySem D {{reD}} x ∷ reifyExtend E {{reE}} e 71 | 72 | reifyExtendᵇ : ∀ {α γ q j} i (C : Binder α β γ i q I) q′ {{reC : ExtendReifyᵇ i C q′}} 73 | -> Extendᵇ i C q′ (μ D₀) j -> List Term 74 | reifyExtendᵇ i (coerce (A , D)) q {{reC}} p = 75 | uncurryProd (visibility i) (uncoerce′ q reC) λ mreA reD -> split q p λ x e -> 76 | let es = reifyExtend (D x) {{reD}} e in 77 | caseExplMaybe _ (λ reA -> elimRelValue _ (reify {{reA}}) (reifyᵢ {{reA}}) x ∷ es) 78 | es 79 | (visibility i) 80 | mreA 81 | 82 | reifyDesc : ∀ {j} D {{reD : ExtendReify D}} -> Name -> Extend D (μ D₀) j -> Term 83 | reifyDesc D n e = vis appCon n (reifyExtend D e) 84 | 85 | reifyAny : ∀ {j} (Ds : List (Desc I β)) {{reD : All ExtendReify Ds}} 86 | -> ∀ d a b ns -> Node D₀ (packData d a b Ds ns) j -> Term 87 | reifyAny [] d a b tt () 88 | reifyAny (D ∷ []) {{reD , _}} d a b (n , ns) e = reifyDesc D {{reD}} n e 89 | reifyAny (D ∷ E ∷ Ds) {{reD , reDs}} d a b (n , ns) (inj₁ e) = reifyDesc D {{reD}} n e 90 | reifyAny (D ∷ E ∷ Ds) {{reD , reDs}} d a b (n , ns) (inj₂ r) = 91 | reifyAny (E ∷ Ds) {{reDs}} d a b ns r 92 | 93 | reifyMu : ∀ {j} -> μ D₀ j -> Term 94 | reifyMu (node e) = reifyAny (consTypes D₀) 95 | (dataName D₀) 96 | (parsTele D₀) 97 | (indsTele D₀) 98 | (consNames D₀) 99 | e 100 | -------------------------------------------------------------------------------- /src/Generic/Reflection/DeriveEq.agda: -------------------------------------------------------------------------------- 1 | module Generic.Reflection.DeriveEq where 2 | 3 | open import Generic.Core 4 | open import Generic.Function.FoldMono 5 | open import Generic.Reflection.ReadData 6 | 7 | fromToClausesOf : Data Type -> Name -> List Clause 8 | fromToClausesOf (packData d a b cs ns) f = unmap (λ {a} -> clauseOf a) ns where 9 | vars : ℕ -> ℕ -> Type -> List (Maybe (String × ℕ) × ℕ) 10 | vars (suc i) j (explPi r s a b) = if isSomeName d a 11 | then (just (s , j) , i) ∷ vars i (suc j) b 12 | else (nothing , i) ∷ vars i j b 13 | vars i j (pi s a b) = vars i j b 14 | vars i j b = [] 15 | 16 | clauseOf : Type -> Name -> Clause 17 | clauseOf c n = clause lhs rhs where 18 | es = explPisToNames c 19 | i = length es 20 | mxs = vars i 0 c 21 | xs = mapMaybe proj₁ mxs 22 | k = length xs 23 | lhs = explRelArg (patCon n (patVars es)) ∷ [] 24 | lams = λ t -> foldr (explLam ∘ proj₁) t xs 25 | each = λ m i -> maybe (proj₂ >>> λ j -> pureVar (k ∸ suc j)) (pureVar (i + k)) m 26 | args = map (uncurry each) mxs 27 | grow = lams (vis appCon n args) 28 | rs = mapMaybe (uncurry λ m i -> vis# 1 appDef f (pureVar i) <$ m) mxs 29 | rhs = vis appDef (quote congn) $ reify k ∷ grow ∷ rs 30 | 31 | toTypeOf : Data Type -> Name -> Type 32 | toTypeOf (packData d a b cs ns) d′ = let ab = appendType a b; k = countPis ab in 33 | appendType (implicitize ab) $ appDef d (pisToArgVars k ab) ‵→ appDef d′ (pisToArgVars (suc k) ab) 34 | 35 | fromTypeOf : Data Type -> Name -> Type 36 | fromTypeOf (packData d a b cs ns) d′ = let ab = appendType a b; k = countPis ab in 37 | appendType (implicitize ab) $ appDef d′ (pisToArgVars k ab) ‵→ appDef d (pisToArgVars (suc k) ab) 38 | 39 | fromToTypeOf : Data Type -> Name -> Name -> Name -> Type 40 | fromToTypeOf (packData d a b cs ns) d′ to from = let ab = appendType a b; k = countPis ab in 41 | appendType (implicitize ab) ∘ pi "x" (explRelArg (appDef d (pisToArgVars k ab))) $ 42 | sate _≡_ (vis# 1 appDef from (vis# 1 appDef to (pureVar 0))) (pureVar 0) 43 | 44 | injTypeOf : Data Type -> Name -> Type 45 | injTypeOf (packData d a b cs ns) d′ = 46 | let ab = appendType a b 47 | k = countPis ab 48 | avs = pisToArgVars k ab 49 | in appendType (implicitize ab) $ sate _↦_ (appDef d avs) (appDef d′ avs) 50 | 51 | deriveDesc : Name -> Data Type -> TC Name 52 | deriveDesc d D = 53 | freshName (showName d ++ˢ "′") >>= λ d′ -> 54 | getType d >>= λ a -> 55 | declareDef (explRelArg d′) a >> 56 | d′ <$ (quoteData D >>= defineTerm d′) 57 | 58 | deriveTo : Data Type -> Name -> Name -> TC Name 59 | deriveTo D d′ fd = 60 | freshName ("to" ++ˢ showName d′) >>= λ to -> 61 | declareDef (explRelArg to) (toTypeOf D d′) >> 62 | to <$ defineTerm to (sateMacro gcoerce (pureDef fd)) 63 | 64 | deriveFrom : Data Type -> Name -> TC Name 65 | deriveFrom D d′ = 66 | freshName ("from" ++ˢ showName d′) >>= λ from -> 67 | declareDef (explRelArg from) (fromTypeOf D d′) >> 68 | from <$ defineTerm from (guncoercePure D) 69 | 70 | deriveFromTo : Data Type -> Name -> Name -> Name -> TC Name 71 | deriveFromTo D d′ to from = 72 | freshName ("fromTo" ++ˢ showName d′) >>= λ from-to -> 73 | declareDef (explRelArg from-to) (fromToTypeOf D d′ to from) >> 74 | from-to <$ defineFun from-to (fromToClausesOf D from-to) 75 | 76 | deriveInj : Data Type -> Name -> Name -> Name -> Name -> TC Name 77 | deriveInj D d′ to from from-to = 78 | freshName ("inj" ++ˢ showName d′) >>= λ inj -> 79 | declareDef (explRelArg inj) (injTypeOf D d′) >> 80 | inj <$ defineTerm inj (sate packInj (pureDef to) (pureDef from) (pureDef from-to)) 81 | 82 | deriveEqTo : Name -> Name -> TC _ 83 | deriveEqTo e d = 84 | getData d >>= λ D -> 85 | deriveDesc d D >>= λ d′ -> 86 | deriveFold d D >>= λ f -> 87 | deriveTo D d′ f >>= λ to -> 88 | deriveFrom D d′ >>= λ from -> 89 | deriveFromTo D d′ to from >>= λ from-to -> 90 | deriveInj D d′ to from from-to >>= λ inj -> 91 | defineTerm e (sate viaInj (pureDef inj)) 92 | -------------------------------------------------------------------------------- /src/Generic/Reflection/ReadData.agda: -------------------------------------------------------------------------------- 1 | module Generic.Reflection.ReadData where 2 | 3 | open import Generic.Core 4 | open import Generic.Function.FoldMono 5 | 6 | ‵π : ArgInfo -> String -> Term -> Term -> Term 7 | ‵π i s a b = sate π (reify i) (sate refl) ∘ sate coerce ∘ sate _,_ a $ 8 | appDef (quote appRel) (implRelArg (reify (relevance i)) ∷ explRelArg (explLam s b) ∷ []) 9 | 10 | quoteHyp : Name -> ℕ -> Type -> Maybe (Maybe Term) 11 | quoteHyp d p (pi s (arg i a) b) = 12 | quoteHyp d p a >>= maybe (const nothing) (fmap (‵π i s a) <$> quoteHyp d p b) 13 | quoteHyp d p t@(appDef n is) = just $ d == n ?> sate var (toTuple ∘ map argVal $ drop p is) 14 | quoteHyp d p t = just nothing 15 | 16 | quoteDesc : Name -> ℕ -> Type -> Maybe Term 17 | quoteDesc d p (pi s (arg i a) b) = 18 | (λ ma' b' -> maybe (λ a' -> sate _⊛_ a' (unshift′ b')) (‵π i s a b') ma') 19 | <$> quoteHyp d p a <*> quoteDesc d p b 20 | quoteDesc d p t = join $ quoteHyp d p t 21 | 22 | -- We didn't need that `β` previously. Why do we need it now? 23 | -- Or, perhaps, why didn't we need it earlier? It kinda makes sense. 24 | quoteData : Data Type -> TC Term 25 | quoteData (packData d a b cs ns) = 26 | case termLevelOf (monoLastType b) ⊗ mapM (quoteDesc d (countPis a)) cs of λ 27 | { nothing -> throw "can't read a data type" 28 | ; (just (β , cs′)) -> (λ qa qb -> explLamsBy a ∘ curryBy b ∘ appDef (quote μ) 29 | $ implRelArg unknown 30 | ∷ implRelArg β 31 | ∷ explRelArg (sate packData (reify d) qa qb (reify cs′) (reify ns)) 32 | ∷ []) 33 | <$> quoteTC a <*> quoteTC b 34 | } 35 | 36 | onFinalMu : ∀ {α} {A : Set α} -> (∀ {ι β} {I : Set ι} -> Data (Desc I β) -> TC A) -> Type -> TC A 37 | onFinalMu k = monoLastType >>> λ 38 | { (appDef (quote μ) (implRelArg qι ∷ implRelArg qβ ∷ implRelArg qI ∷ explRelArg qD ∷ _)) -> 39 | bindTC (unquoteTC qι) λ ι -> 40 | bindTC (unquoteTC qβ) λ β -> 41 | bindTC (unquoteTC qI) λ I -> 42 | bindTC (unquoteTC qD) λ D -> k {ι} {β} {I} D 43 | ; _ -> 44 | throw "the return type must be μ" 45 | } 46 | 47 | guncoercePure : Data Type -> Term 48 | guncoercePure (packData d a b cs ns) = explLam "x" ∘ vis appDef (quote curryFoldMono) 49 | $ explUncurryBy b (vis appDef d (replicate (countExplPis a) unknown)) 50 | ∷ pureVar 0 51 | ∷ unmap pureCon ns 52 | 53 | macro 54 | readData : Name -> Term -> TC _ 55 | readData d ?r = getData d >>= quoteData >>= unify ?r 56 | 57 | gcoerce : Name -> Term -> TC _ 58 | gcoerce fd ?r = inferNormType ?r >>= onFinalMu λ{ D@(packData _ _ b _ _) -> 59 | quoteTC (μ D) >>= λ μD -> 60 | traverseAll quoteTC (allCons D) >>= λ cs′ -> 61 | unify ?r $ vis appDef fd (curryBy b μD ∷ allToList cs′) 62 | } 63 | 64 | guncoerce : ∀ {ι β} {I : Set ι} {D : Data (Desc I β)} {j} -> μ D j -> Term -> TC _ 65 | guncoerce {D = packData d a b cs ns} e ?r = 66 | quoteTC e >>= λ qe -> unify ?r ∘ vis appDef (quote curryFoldMono) 67 | $ explUncurryBy b (vis appDef d (replicate (countExplPis a) unknown)) 68 | ∷ qe 69 | ∷ unmap pureCon ns 70 | -------------------------------------------------------------------------------- /src/Generic/Test.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test where 2 | 3 | import Generic.Test.Data 4 | import Generic.Test.DeriveEq 5 | import Generic.Test.Elim 6 | import Generic.Test.Eq 7 | import Generic.Test.Experiment 8 | import Generic.Test.ReadData 9 | import Generic.Test.Reify 10 | -------------------------------------------------------------------------------- /src/Generic/Test/Data.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data where 2 | 3 | open import Generic.Test.Data.Fin public 4 | open import Generic.Test.Data.Lift public 5 | open import Generic.Test.Data.Product public 6 | open import Generic.Test.Data.List public 7 | open import Generic.Test.Data.Maybe public 8 | open import Generic.Test.Data.Vec public 9 | open import Generic.Test.Data.W public 10 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/Fin.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.Fin where 2 | 3 | open import Generic.Main 4 | 5 | import Data.Fin as Fin 6 | 7 | Fin : ℕ -> Set 8 | Fin = readData Fin.Fin 9 | 10 | pattern fzero {n} = #₀ (relv n , lrefl) 11 | pattern fsuc {n} i = !#₁ (relv n , i , lrefl) 12 | 13 | elimFin : ∀ {n π} 14 | -> (P : ∀ {n} -> Fin n -> Set π) 15 | -> (∀ {n} {i : Fin n} -> P i -> P (fsuc i)) 16 | -> (∀ {n} -> P {suc n} fzero) 17 | -> (i : Fin n) 18 | -> P i 19 | elimFin P f x fzero = x 20 | elimFin P f x (fsuc i) = f (elimFin P f x i) 21 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/Lift.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.Lift where 2 | 3 | open import Generic.Main as Main hiding (Lift; lift; lower) 4 | 5 | Lift : ∀ {α} β -> Set α -> Set (α ⊔ β) 6 | Lift = readData Main.Lift 7 | 8 | pattern lift x = !#₀ (relv x , lrefl) 9 | 10 | lower : ∀ {α} {A : Set α} β -> Lift β A -> A 11 | lower β (lift x) = x 12 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/List.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.List where 2 | 3 | module _ where 4 | open import Generic.Main as Main hiding ([]; _∷_) renaming (List to StdList) 5 | 6 | infixr 5 _∷_ _∷′_ 7 | 8 | List : ∀ {α} -> Set α -> Set α 9 | List = readData StdList 10 | 11 | pattern [] = #₀ lrefl 12 | pattern _∷_ x xs = !#₁ (relv x , xs , lrefl) 13 | 14 | _∷′_ : ∀ {α} {A : Set α} -> A -> List A -> List A 15 | _∷′_ = _∷_ 16 | 17 | elimList : ∀ {α π} {A : Set α} 18 | -> (P : List A -> Set π) 19 | -> (∀ {xs} x -> P xs -> P (x ∷ xs)) 20 | -> P [] 21 | -> ∀ xs 22 | -> P xs 23 | elimList P f z [] = z 24 | elimList P f z (x ∷ xs) = f x (elimList P f z xs) 25 | 26 | toStdList : ∀ {α} {A : Set α} -> List A -> StdList A 27 | toStdList xs = guncoerce xs 28 | 29 | -- The entire content of `Data.List.Base` (modulo `Generic.Test.Data.Maybe` instead of 30 | -- `Data.Maybe.Base` and _∷_ was renamed to _∷′_ in some places) 31 | 32 | open import Data.Nat.Base using (ℕ; zero; suc; _+_; _*_) 33 | open import Data.Sum as Sum using (_⊎_; inj₁; inj₂) 34 | open import Data.Bool.Base using (Bool; false; true; not; _∧_; _∨_; if_then_else_) 35 | open import Generic.Test.Data.Maybe 36 | open import Data.Product as Prod using (_×_; _,_) 37 | open import Function 38 | 39 | infixr 5 _++_ 40 | 41 | [_] : ∀ {a} {A : Set a} → A → List A 42 | [ x ] = x ∷ [] 43 | 44 | _++_ : ∀ {a} {A : Set a} → List A → List A → List A 45 | [] ++ ys = ys 46 | (x ∷ xs) ++ ys = x ∷ (xs ++ ys) 47 | 48 | -- Snoc. 49 | 50 | infixl 5 _∷ʳ_ 51 | 52 | _∷ʳ_ : ∀ {a} {A : Set a} → List A → A → List A 53 | xs ∷ʳ x = xs ++ [ x ] 54 | 55 | null : ∀ {a} {A : Set a} → List A → Bool 56 | null [] = true 57 | null (x ∷ xs) = false 58 | 59 | -- * List transformations 60 | 61 | map : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → List A → List B 62 | map f [] = [] 63 | map f (x ∷ xs) = f x ∷ map f xs 64 | 65 | replicate : ∀ {a} {A : Set a} → (n : ℕ) → A → List A 66 | replicate zero x = [] 67 | replicate (suc n) x = x ∷ replicate n x 68 | 69 | zipWith : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} 70 | → (A → B → C) → List A → List B → List C 71 | zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys 72 | zipWith f _ _ = [] 73 | 74 | zip : ∀ {a b} {A : Set a} {B : Set b} → List A → List B → List (A × B) 75 | zip = zipWith (_,_) 76 | 77 | intersperse : ∀ {a} {A : Set a} → A → List A → List A 78 | intersperse x [] = [] 79 | intersperse x (y ∷ []) = [ y ] 80 | intersperse x (y ∷ z ∷ zs) = y ∷ x ∷ intersperse x (z ∷ zs) 81 | 82 | -- * Reducing lists (folds) 83 | 84 | foldr : ∀ {a b} {A : Set a} {B : Set b} → (A → B → B) → B → List A → B 85 | foldr c n [] = n 86 | foldr c n (x ∷ xs) = c x (foldr c n xs) 87 | 88 | foldl : ∀ {a b} {A : Set a} {B : Set b} → (A → B → A) → A → List B → A 89 | foldl c n [] = n 90 | foldl c n (x ∷ xs) = foldl c (c n x) xs 91 | 92 | -- ** Special folds 93 | 94 | concat : ∀ {a} {A : Set a} → List (List A) → List A 95 | concat = foldr _++_ [] 96 | 97 | concatMap : ∀ {a b} {A : Set a} {B : Set b} → 98 | (A → List B) → List A → List B 99 | concatMap f = concat ∘ map f 100 | 101 | and : List Bool → Bool 102 | and = foldr _∧_ true 103 | 104 | or : List Bool → Bool 105 | or = foldr _∨_ false 106 | 107 | any : ∀ {a} {A : Set a} → (A → Bool) → List A → Bool 108 | any p = or ∘ map p 109 | 110 | all : ∀ {a} {A : Set a} → (A → Bool) → List A → Bool 111 | all p = and ∘ map p 112 | 113 | sum : List ℕ → ℕ 114 | sum = foldr _+_ 0 115 | 116 | product : List ℕ → ℕ 117 | product = foldr _*_ 1 118 | 119 | length : ∀ {a} {A : Set a} → List A → ℕ 120 | length = foldr (λ _ → suc) 0 121 | 122 | import Data.List.Base 123 | 124 | reverse : ∀ {a} {A : Set a} → List A → List A 125 | reverse = foldl (λ rev x → x ∷′ rev) [] 126 | 127 | -- * Building lists 128 | 129 | -- ** Scans 130 | 131 | scanr : ∀ {a b} {A : Set a} {B : Set b} → 132 | (A → B → B) → B → List A → List B 133 | scanr f e [] = e ∷ [] 134 | scanr f e (x ∷ xs) with scanr f e xs 135 | ... | [] = [] -- dead branch 136 | ... | y ∷ ys = f x y ∷ y ∷ ys 137 | 138 | scanl : ∀ {a b} {A : Set a} {B : Set b} → 139 | (A → B → A) → A → List B → List A 140 | scanl f e [] = e ∷ [] 141 | scanl f e (x ∷ xs) = e ∷ scanl f (f e x) xs 142 | 143 | -- ** Unfolding 144 | 145 | -- Unfold. Uses a measure (a natural number) to ensure termination. 146 | 147 | unfold : ∀ {a b} {A : Set a} (B : ℕ → Set b) 148 | (f : ∀ {n} → B (suc n) → Maybe (A × B n)) → 149 | ∀ {n} → B n → List A 150 | unfold B f {n = zero} s = [] 151 | unfold B f {n = suc n} s with f s 152 | ... | nothing = [] 153 | ... | just (x , s') = x ∷ unfold B f s' 154 | 155 | -- downFrom 3 = 2 ∷ 1 ∷ 0 ∷ []. 156 | 157 | downFrom : ℕ → List ℕ 158 | downFrom n = unfold Singleton f (wrap n) 159 | where 160 | data Singleton : ℕ → Set where 161 | wrap : (n : ℕ) → Singleton n 162 | 163 | f : ∀ {n} → Singleton (suc n) → Maybe (ℕ × Singleton n) 164 | f {n} (wrap .(suc n)) = just (n , wrap n) 165 | 166 | -- ** Conversions 167 | 168 | fromMaybe : ∀ {a} {A : Set a} → Maybe A → List A 169 | fromMaybe (just x) = [ x ] 170 | fromMaybe nothing = [] 171 | 172 | -- * Sublists 173 | 174 | -- ** Extracting sublists 175 | 176 | take : ∀ {a} {A : Set a} → ℕ → List A → List A 177 | take zero xs = [] 178 | take (suc n) [] = [] 179 | take (suc n) (x ∷ xs) = x ∷ take n xs 180 | 181 | drop : ∀ {a} {A : Set a} → ℕ → List A → List A 182 | drop zero xs = xs 183 | drop (suc n) [] = [] 184 | drop (suc n) (x ∷ xs) = drop n xs 185 | 186 | splitAt : ∀ {a} {A : Set a} → ℕ → List A → (List A × List A) 187 | splitAt zero xs = ([] , xs) 188 | splitAt (suc n) [] = ([] , []) 189 | splitAt (suc n) (x ∷ xs) with splitAt n xs 190 | ... | (ys , zs) = (x ∷ ys , zs) 191 | 192 | takeWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A 193 | takeWhile p [] = [] 194 | takeWhile p (x ∷ xs) with p x 195 | ... | true = x ∷ takeWhile p xs 196 | ... | false = [] 197 | 198 | dropWhile : ∀ {a} {A : Set a} → (A → Bool) → List A → List A 199 | dropWhile p [] = [] 200 | dropWhile p (x ∷ xs) with p x 201 | ... | true = dropWhile p xs 202 | ... | false = x ∷ xs 203 | 204 | span : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A) 205 | span p [] = ([] , []) 206 | span p (x ∷ xs) with p x 207 | ... | true = Prod.map (_∷′_ x) id (span p xs) 208 | ... | false = ([] , x ∷ xs) 209 | 210 | break : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A) 211 | break p = span (not ∘ p) 212 | 213 | inits : ∀ {a} {A : Set a} → List A → List (List A) 214 | inits [] = [] ∷ [] 215 | inits (x ∷ xs) = [] ∷ map (_∷′_ x) (inits xs) 216 | 217 | tails : ∀ {a} {A : Set a} → List A → List (List A) 218 | tails [] = [] ∷ [] 219 | tails (x ∷ xs) = (x ∷ xs) ∷ tails xs 220 | 221 | infixl 5 _∷ʳ'_ 222 | 223 | data InitLast {a} {A : Set a} : List A → Set a where 224 | []ʳ : InitLast [] 225 | _∷ʳ'_ : (xs : List A) (x : A) → InitLast (xs ∷ʳ x) 226 | 227 | initLast : ∀ {a} {A : Set a} (xs : List A) → InitLast xs 228 | initLast [] = []ʳ 229 | initLast (x ∷ xs) with initLast xs 230 | initLast (x ∷ .[]) | []ʳ = [] ∷ʳ' x 231 | initLast (x ∷ .(ys ∷ʳ y)) | ys ∷ʳ' y = (x ∷ ys) ∷ʳ' y 232 | 233 | -- * Searching lists 234 | 235 | -- ** Searching with a predicate 236 | 237 | -- A generalised variant of filter. 238 | 239 | gfilter : ∀ {a b} {A : Set a} {B : Set b} → 240 | (A → Maybe B) → List A → List B 241 | gfilter p [] = [] 242 | gfilter p (x ∷ xs) with p x 243 | ... | just y = y ∷ gfilter p xs 244 | ... | nothing = gfilter p xs 245 | 246 | filter : ∀ {a} {A : Set a} → (A → Bool) → List A → List A 247 | filter p = gfilter (λ x → if p x then just′ x else nothing) 248 | 249 | partition : ∀ {a} {A : Set a} → (A → Bool) → List A → (List A × List A) 250 | partition p [] = ([] , []) 251 | partition p (x ∷ xs) with p x | partition p xs 252 | ... | true | (ys , zs) = (x ∷ ys , zs) 253 | ... | false | (ys , zs) = (ys , x ∷ zs) 254 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/Maybe.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.Maybe where 2 | 3 | open import Generic.Main as Main hiding (Maybe; just; nothing) 4 | 5 | Maybe : ∀ {α} -> Set α -> Set α 6 | Maybe = readData Main.Maybe 7 | 8 | pattern nothing = #₀ lrefl 9 | pattern just x = !#₁ (relv x , lrefl) 10 | 11 | just′ : ∀ {α} {A : Set α} -> A -> Maybe A 12 | just′ = just 13 | 14 | nothing′ : ∀ {α} {A : Set α} -> Maybe A 15 | nothing′ = nothing 16 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/Product.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.Product where 2 | 3 | open import Generic.Main as Main hiding (Σ; proj₁; proj₂; _,′_) renaming (_,_ to _,′_) 4 | 5 | infixr 4 _,_ 6 | 7 | Σ : ∀ {α β} -> (A : Set α) -> (A -> Set β) -> Set (α ⊔ β) 8 | Σ = readData Main.Σ 9 | 10 | pattern _,_ x y = !#₀ (relv x ,′ relv y ,′ lrefl) 11 | 12 | proj₁ : ∀ {α β} {A : Set α} {B : A -> Set β} -> Σ A B -> A 13 | proj₁ (x , y) = x 14 | 15 | proj₂ : ∀ {α β} {A : Set α} {B : A -> Set β} -> (p : Σ A B) -> B (proj₁ {B = B} p) 16 | proj₂ (x , y) = y 17 | 18 | ηo : ∀ {α β} {A : Set α} {B : A -> Set β} -> (p : Σ A B) -> p ≡ proj₁ {B = B} p , proj₂ {B = B} p 19 | ηo (x , y) = refl 20 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/Vec.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.Vec where 2 | 3 | open import Generic.Main 4 | 5 | import Data.Vec as Vec 6 | 7 | infixr 5 _∷ᵥ_ 8 | 9 | Vec : ∀ {α} -> Set α -> ℕ -> Set α 10 | Vec = readData Vec.Vec 11 | 12 | -- []ᵥ : ∀ {α} {A : Set α} -> Vec A 0 13 | pattern []ᵥ = #₀ lrefl 14 | 15 | -- _∷ᵥ_ : ∀ {n α} {A : Set α} -> A -> Vec A n -> Vec A (suc n) 16 | pattern _∷ᵥ_ {n} x xs = !#₁ (relv n , relv x , xs , lrefl) 17 | 18 | elimVec : ∀ {n α π} {A : Set α} 19 | -> (P : ∀ {n} -> Vec A n -> Set π) 20 | -> (∀ {n} {xs : Vec A n} x -> P xs -> P (x ∷ᵥ xs)) 21 | -> P []ᵥ 22 | -> (xs : Vec A n) 23 | -> P xs 24 | elimVec P f z []ᵥ = z 25 | elimVec P f z (x ∷ᵥ xs) = f x (elimVec P f z xs) 26 | -------------------------------------------------------------------------------- /src/Generic/Test/Data/W.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Data.W where 2 | 3 | open import Generic.Main 4 | 5 | data W′ {α β} (A : Set α) (B : A -> Set β) : Set (α ⊔ β) where 6 | sup′ : ∀ {x} -> (B x -> W′ A B) -> W′ A B 7 | 8 | W : ∀ {α β} -> (A : Set α) -> (A -> Set β) -> Set (α ⊔ β) 9 | W = readData W′ 10 | 11 | pattern sup x g = !#₀ (relv x , g , lrefl) 12 | 13 | elimW : ∀ {α β π} {A : Set α} {B : A -> Set β} 14 | -> (P : W A B -> Set π) 15 | -> (∀ {x} {g : B x -> W A B} -> (∀ y -> P (g y)) -> P (sup x g)) 16 | -> ∀ w 17 | -> P w 18 | elimW {B = B} P h (sup x g) = h (λ y -> elimW {B = B} P h (g y)) 19 | -------------------------------------------------------------------------------- /src/Generic/Test/DeriveEq.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.DeriveEq where 2 | 3 | open import Generic.Main 4 | 5 | open import Data.Vec using (Vec) renaming ([] to []ᵥ; _∷_ to _∷ᵥ_) 6 | 7 | module DeriveEqStar where 8 | open import Relation.Binary 9 | open import Relation.Binary.Construct.Closure.ReflexiveTransitive 10 | 11 | instance StarEq : ∀ {i t} {I : Set i} {T : Rel I t} {i j} 12 | {{iEq : Eq I}} {{tEq : ∀ {i j} -> Eq (T i j)}} -> Eq (Star T i j) 13 | unquoteDef StarEq = deriveEqTo StarEq (quote Star) 14 | 15 | module DeriveEqVec where 16 | instance VecEq : ∀ {n α} {A : Set α} {{aEq : Eq A}} -> Eq (Vec A n) 17 | unquoteDef VecEq = deriveEqTo VecEq (quote Vec) 18 | 19 | xs : Vec ℕ 3 20 | xs = 2 ∷ᵥ 4 ∷ᵥ 1 ∷ᵥ []ᵥ 21 | 22 | test₁ : xs ≟ xs ≡ yes refl 23 | test₁ = refl 24 | 25 | test₂ : xs ≟ (2 ∷ᵥ 4 ∷ᵥ 2 ∷ᵥ []ᵥ) ≡ no _ 26 | test₂ = refl 27 | 28 | module DeriveEqD where 29 | data D {α β} (A : Set α) (B : A -> Set β) : ∀ {n x} -> Vec (B x) n -> ℕ -> Set (α ⊔ β) where 30 | c₁ : ∀ {x n} (ys : Vec (B x) n) m -> .A -> D A B ys m 31 | c₂ : ∀ {x n m y} {ys zs : Vec (B x) n} 32 | -> D A B (y ∷ᵥ ys) 0 -> Vec A m -> D A B ys (suc n) -> D A B zs n 33 | 34 | -- `VecEq` is in scope. 35 | instance DEq : ∀ {α β} {A : Set α} {B : A -> Set β} {n m x} {ys : Vec (B x) n} 36 | {{aEq : Eq A}} {{bEq : ∀ {x} -> Eq (B x)}} -> Eq (D A B ys m) 37 | unquoteDef DEq = deriveEqTo DEq (quote D) 38 | 39 | -- -- Seems like the problem is that irrelevance and metavariables resolution do not play well. 40 | -- module DeriveEqE where 41 | -- data E {α} (A : Set α) : ∀ {n} -> .(Vec A n) -> Set α where 42 | -- c₁ : ∀ {n} -> .(xs : Vec A n) -> E A xs 43 | 44 | -- instance EEq : ∀ {α n} {A : Set α} .{xs : Vec A n} -> Eq (E A xs) 45 | -- unquoteDef EEq = deriveEqTo EEq (quote E) 46 | -- -- Variable xs is declared irrelevant, so it cannot be used here 47 | -- -- when checking that the expression xs has type _B_76 A _ n₁ _ n₁ 48 | 49 | module DeriveEqF where 50 | data F {α} (A : Set α) : Set α where 51 | c₁ : ∀ {n} -> .(Vec A n) -> F A 52 | 53 | instance FEq : ∀ {α} {A : Set α} -> Eq (F A) 54 | unquoteDef FEq = deriveEqTo FEq (quote F) 55 | -------------------------------------------------------------------------------- /src/Generic/Test/Elim.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Elim where 2 | 3 | open import Generic.Core 4 | open import Generic.Function.Elim 5 | open import Generic.Test.Data.Vec 6 | 7 | -- Is it possible to get rid of these `lift`s? 8 | elimVec′ : ∀ {n α π} {A : Set α} 9 | -> (P : ∀ {n} -> Vec A n -> Set π) 10 | -> (∀ {n} {xs : Vec A n} x -> P xs -> P (x ∷ᵥ xs)) 11 | -> P []ᵥ 12 | -> (xs : Vec A n) 13 | -> P xs 14 | elimVec′ P f z = elim P (lift z , λ x r -> lift (f x r)) 15 | -------------------------------------------------------------------------------- /src/Generic/Test/Eq.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Eq where 2 | 3 | open import Generic.Main hiding (List; []; _∷_) 4 | open import Generic.Test.Data.Fin 5 | open import Generic.Test.Data.List 6 | open import Generic.Test.Data.Vec 7 | 8 | xs : Vec (List (Fin 4)) 3 9 | xs = (fsuc fzero ∷ fzero ∷ []) 10 | ∷ᵥ (fsuc (fsuc fzero) ∷ []) 11 | ∷ᵥ (fzero ∷ fsuc (fsuc (fsuc fzero)) ∷ []) 12 | ∷ᵥ []ᵥ 13 | 14 | test : xs ≟ xs ≡ yes refl 15 | test = refl 16 | -------------------------------------------------------------------------------- /src/Generic/Test/Experiment.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --type-in-type #-} 2 | 3 | module Generic.Test.Experiment where 4 | 5 | open import Generic.Lib.Prelude 6 | 7 | data Desc (I : Set) : Set₁ where 8 | ret : I -> Desc I 9 | π : (A : Set) -> (A -> Desc I) -> Desc I 10 | _⊕_ : Desc I -> Desc I -> Desc I 11 | 12 | data RecDesc (I : Set) : Set₁ where 13 | rec : ((I -> Set) -> Desc I) -> RecDesc I 14 | 15 | ⟦_⟧ : ∀ {I} -> Desc I -> (I -> Set) -> I -> Set 16 | ⟦ ret i ⟧ B j = i ≡ j 17 | ⟦ π A D ⟧ B j = ∃ λ x -> ⟦ D x ⟧ B j 18 | ⟦ D ⊕ E ⟧ B j = ⟦ D ⟧ B j ⊎ ⟦ E ⟧ B j 19 | 20 | ⟦_⟧ʳ : ∀ {I} -> RecDesc I -> (I -> Set) -> I -> Set 21 | ⟦ rec K ⟧ʳ B j = ⟦ K B ⟧ B j 22 | 23 | {-# NO_POSITIVITY_CHECK #-} 24 | data μ {I} (D : RecDesc I) j : Set where 25 | node : ⟦ D ⟧ʳ (μ D) j -> μ D j 26 | 27 | DescEq : ∀ {I} -> Desc I -> Set 28 | DescEq (ret i) = ⊤ 29 | DescEq (π A D) = Eq A × ∀ x -> DescEq (D x) 30 | DescEq (D ⊕ E) = DescEq D × DescEq E 31 | 32 | RecDescEq : ∀ {I} -> RecDesc I -> Set₁ 33 | RecDescEq (rec K) = ∀ {B} {{eqB : ∀ {i} -> Eq (B i)}} -> DescEq (K B) 34 | 35 | instance 36 | {-# TERMINATING #-} 37 | DataEq : ∀ {I} {D : RecDesc I} {j} {{recDescEq : RecDescEq D}} -> Eq (μ D j) 38 | DataEq {D = D} {j} = record 39 | { _≟_ = eqMu 40 | } where 41 | node-inj : ∀ {I D j x y} -> node {I} {D} {j} x ≡ node y -> x ≡ y 42 | node-inj refl = refl 43 | 44 | eqSem : ∀ {j} D₀ D {{descEq : DescEq D}} -> IsSet (⟦ D ⟧ (μ D₀) j) 45 | eqSem D₀ (ret i) refl refl = yes refl 46 | eqSem D₀ (π A D) {{eqA , eqD}} p₁ p₂ = 47 | decProd (_≟_ {{eqA}}) (eqSem D₀ (D _) {{eqD _}}) p₁ p₂ 48 | eqSem D₀ (D ⊕ E) {{eqD , eqE}} s₁ s₂ = 49 | decSum (eqSem D₀ D {{eqD}}) (eqSem D₀ E {{eqE}}) s₁ s₂ 50 | 51 | eqMu : ∀ {D} {{recDescEq : RecDescEq D}} -> IsSet (μ D j) 52 | eqMu {D@(rec K)} {{recDescEq}} (node d₁) (node d₂) = 53 | dcong node node-inj (eqSem D _ {{recDescEq {{DataEq {D = D}}}}} d₁ d₂) 54 | 55 | module Example1 where 56 | open import Generic.Main using (deriveEqTo) 57 | 58 | instance 59 | ListEq : {A : Set} {{eqA : Eq A}} -> Eq (List A) 60 | unquoteDef ListEq = deriveEqTo ListEq (quote List) 61 | 62 | data Rose (A : Set) : Set where 63 | rose : A -> List (Rose A) -> Rose A 64 | 65 | Rose′ : Set -> Set 66 | Rose′ A = flip μ tt $ rec λ Rose′ -> π A λ _ -> π (List (Rose′ tt)) λ _ -> ret tt 67 | 68 | pattern rose′ x rs = node (x , rs , refl) 69 | 70 | {-# TERMINATING #-} 71 | Rose→Rose′ : ∀ {A} -> Rose A -> Rose′ A 72 | Rose→Rose′ (rose x rs) = rose′ x (map Rose→Rose′ rs) 73 | 74 | {-# TERMINATING #-} 75 | Rose′→Rose : ∀ {A} -> Rose′ A -> Rose A 76 | Rose′→Rose (rose′ x rs) = rose x (map Rose′→Rose rs) 77 | 78 | arose : Rose′ ℕ 79 | arose = rose′ 4 (rose′ 1 (rose′ 6 [] ∷ []) ∷ rose′ 2 [] ∷ []) 80 | 81 | find : {A : Set} {{x : A}} -> A 82 | find {{x}} = x 83 | 84 | -- -- Instance search cannot be used to find elements in an explicit function type 85 | -- test : _≟_ {{DataEq {{λ {_} -> find}}}} arose arose ≡ yes refl 86 | -- test = refl 87 | 88 | module Example2 where 89 | -- If `A` is a parameter, then this definition would be strictly positive, 90 | -- but I intentionally made it an index, because we can't make `A` a parameter 91 | -- in the described version of `D`. 92 | {-# NO_POSITIVITY_CHECK #-} 93 | data D : Set -> Set where 94 | ret : ∀ {A} -> A -> D A 95 | rec : ∀ {A} -> D (D A) -> D A 96 | 97 | D′ : Set -> Set 98 | D′ = μ $ rec λ D′ -> (π Set λ A -> π A λ _ -> ret A) 99 | ⊕ (π Set λ A -> π (D′ (D′ A)) λ _ -> ret A) 100 | 101 | pattern ret′ x = node (inj₁ (_ , x , refl)) 102 | pattern rec′ r = node (inj₂ (_ , r , refl)) 103 | 104 | Dⁿ : ℕ -> Set -> Set 105 | Dⁿ 0 A = A 106 | Dⁿ (suc n) A = D (Dⁿ n A) 107 | 108 | D′ⁿ : ℕ -> Set -> Set 109 | D′ⁿ 0 A = A 110 | D′ⁿ (suc n) A = D′ (D′ⁿ n A) 111 | 112 | ⟨D→D′⟩ⁿ : ∀ {A} n -> D (Dⁿ n A) -> D′ (D′ⁿ n A) 113 | ⟨D→D′⟩ⁿ 0 (ret x) = ret′ x 114 | ⟨D→D′⟩ⁿ (suc n) (ret r) = ret′ (⟨D→D′⟩ⁿ n r) 115 | ⟨D→D′⟩ⁿ n (rec r) = rec′ (⟨D→D′⟩ⁿ (suc n) r) 116 | 117 | ⟨D′→D⟩ⁿ : ∀ {A} n -> D′ (D′ⁿ n A) -> D (Dⁿ n A) 118 | ⟨D′→D⟩ⁿ 0 (ret′ x) = ret x 119 | ⟨D′→D⟩ⁿ (suc n) (ret′ r) = ret (⟨D′→D⟩ⁿ n r) 120 | ⟨D′→D⟩ⁿ n (rec′ r) = rec (⟨D′→D⟩ⁿ (suc n) r) 121 | 122 | D→D′ : ∀ {A} -> D A -> D′ A 123 | D→D′ = ⟨D→D′⟩ⁿ 0 124 | 125 | D′→D : ∀ {A} -> D′ A -> D A 126 | D′→D = ⟨D′→D⟩ⁿ 0 127 | -------------------------------------------------------------------------------- /src/Generic/Test/ReadData.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.ReadData where 2 | 3 | open import Generic.Main 4 | 5 | data D {α β} (A : Set α) (B : ℕ -> Set β) : ∀ {n} -> B n -> List ℕ -> Set (α ⊔ β) where 6 | c₁ : ∀ {n} (y : B n) xs -> A -> D A B y xs 7 | c₂ : ∀ {y : B 0} -> (∀ {n} (y : B n) {{xs}} -> D A B y xs) -> .(List A) -> D A B y [] 8 | 9 | -- No longer works: Failed to solve the following constraints: Has bigger sort: Setω 10 | -- D′ : TypeOf D 11 | D′ : ∀ {α β} -> (A : Set α) -> (B : ℕ -> Set β) -> ∀ {n} -> B n -> List ℕ -> Set (α ⊔ β) 12 | D′ = readData D 13 | 14 | unquoteDecl foldD = deriveFoldTo foldD (quote D) 15 | 16 | inj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D A B y xs -> D′ A B y xs 17 | inj = gcoerce foldD 18 | 19 | outj : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D′ A B y xs -> D A B y xs 20 | outj d = guncoerce d 21 | 22 | pattern c₁′ {n} y xs x = #₀ (relv n , relv y , relv xs , relv x , lrefl) 23 | pattern c₂′ {y} r ys = !#₁ (relv y , r , irrv ys , lrefl) 24 | 25 | inj′ : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D A B y xs -> D′ A B y xs 26 | inj′ (c₁ y xs x) = c₁′ y xs x 27 | inj′ (c₂ r ys) = c₂′ (λ y -> inj′ (r y)) ys 28 | 29 | outj′ : ∀ {α β} {A : Set α} {B : ℕ -> Set β} {n xs} {y : B n} -> D′ A B y xs -> D A B y xs 30 | outj′ (c₁′ y xs x) = c₁ y xs x 31 | outj′ (c₂′ r ys) = c₂ (λ y -> outj′ (r y)) ys 32 | -------------------------------------------------------------------------------- /src/Generic/Test/Reify.agda: -------------------------------------------------------------------------------- 1 | module Generic.Test.Reify where 2 | 3 | open import Generic.Core 4 | open import Generic.Property.Reify 5 | open import Generic.Test.Data.Fin 6 | open import Generic.Test.Data.Vec 7 | 8 | open import Data.Fin renaming (Fin to StdFin) 9 | open import Data.Vec renaming (Vec to StdVec) 10 | 11 | xs : Vec (Fin 4) 3 12 | xs = fsuc (fsuc (fsuc fzero)) ∷ᵥ fzero ∷ᵥ fsuc fzero ∷ᵥ []ᵥ 13 | 14 | xs′ : StdVec (StdFin 4) 3 15 | xs′ = suc (suc (suc zero)) ∷ zero ∷ (suc zero) ∷ [] 16 | 17 | test : reflect xs ≡ xs′ 18 | test = refl 19 | --------------------------------------------------------------------------------