agda-prelude/Prelude.agda
2025-04-14 20:20:43 +02:00

220 lines
7.2 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

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

module Prelude where
open import Function public
open import Agda.Builtin.FromNat public
open import Agda.Builtin.FromNeg public
open import Agda.Builtin.FromString public
module Level where
open import Agda.Primitive public using (Level ; lzero ; lsuc ; _⊔_)
open import Level public using (Lift ; lift ; lower)
import Level.Literals as Lit
instance liftᴵ : {a } {A : Set a} A Lift A
liftᴵ = lift it
instance number : Number Level
number = Lit.Level
open Level public using (Level ; lzero ; lsuc ; _⊔_)
module where
open import Data.Unit public hiding (module )
open import Data.Unit.Properties public
open public using ( ; tt)
module where
open import Data.Empty public hiding (⊥-elim)
open import Data.Empty.Irrelevant public
open public using ( ; ⊥-elim)
module where
open import Data.Nat public hiding (module )
open import Data.Nat.Properties public
import Data.Nat.Literals as Lit
instance number = Lit.number
open public using ( ; zero ; suc)
module Fin where
import Data.Fin hiding (module Fin)
open Data.Fin public
open import Data.Fin.Properties public
open import Data.Fin.Literals public renaming (number to number)
instance number : {n} Number (Fin n)
number = number $-
open Fin public using (Fin ; zero ; suc ; #_)
module Bool where
open import Data.Bool public hiding (module Bool)
open import Data.Bool.Properties public
open Bool public using (Bool ; true ; false ; if_then_else_)
module where
open import Data.Sum public hiding (module _⊎_)
open import Data.Sum.Properties public
open public using (_⊎_ ; inj₁ ; inj₂)
module Σ where
open import Data.Product public hiding (module Σ)
open import Data.Product.Properties public
open Σ public using (Σ ; Σ-syntax ; _×_ ; ; ∃-syntax ; ; ∄-syntax ;
_,_ ; -,_ ; proj₁ ; proj₂)
module List where
open import Data.List public hiding (module List)
open import Data.List.Properties public
open List public using (List ; [] ; _∷_)
module Vec where
open import Data.Vec public hiding (module Vec)
open import Data.Vec.Properties public
open Vec public using (Vec ; [] ; _∷_)
module Maybe where
open import Data.Maybe public hiding (module Maybe)
open import Data.Maybe.Properties public
open Maybe public using (Maybe ; nothing ; just)
module Rel where
open import Relation.Nullary public renaming
(Irrelevant to Irrelevant₀ ; Stable to Stable₀ ;
WeaklyDecidable to WeaklyDecidable₀)
open import Relation.Nullary.Decidable public
open import Relation.Unary public hiding (⌊_⌋)
renaming (Irrelevant to Irrelevant₁ ; Recomputable to Recomputable₁ ;
Stable to Stable₁ ; WeaklyDecidable to WeaklyDecidable₁ ;
Universal to Universal₁ ; Decidable to Decidable₁ ; Empty to Empty₁ ;
_⇒_ to _⇒₁_)
open import Relation.Binary public
renaming (Irrelevant to Irrelevant₂ ; Recomputable to Recomputable₂ ;
Stable to Stable₂ ; WeaklyDecidable to WeaklyDecidable₂ ;
Universal to Universal₂ ; Decidable to Decidable₂ ; Empty to Empty₂ ;
_⇒_ to _⇒₂_)
module _ where
private variable
ℓ₁ ℓ₂ ℓ₃ ℓ₄ ℓ₅ ℓ₆ ℓ₇ : Level
A B C D E F G : Set _
_Preserves₃_⟶_⟶_⟶_ :
(A B C D) Rel A ℓ₁ Rel B ℓ₂ Rel C ℓ₃ Rel D ℓ₄ Set _
f Preserves₃ _ᵃ_ _ᵇ_ _ᶜ_ _ᵈ_ =
{a a b b c c}
a ∼ᵃ a b ∼ᵇ b c ∼ᶜ c
f a b c ∼ᵈ f a b c
_Preserves₄_⟶_⟶_⟶_⟶_ :
(A B C D E)
Rel A ℓ₁ Rel B ℓ₂ Rel C ℓ₃ Rel D ℓ₄ Rel E ℓ₅ Set _
f Preserves₄ _ᵃ_ _ᵇ_ _ᶜ_ _ᵈ_ _ᵉ_ =
{a a b b c c d d}
a ∼ᵃ a b ∼ᵇ b c ∼ᶜ c d ∼ᵈ d
f a b c d ∼ᵉ f a b c d
_Preserves₅_⟶_⟶_⟶_⟶_⟶_ :
(A B C D E F)
Rel A ℓ₁ Rel B ℓ₂ Rel C ℓ₃ Rel D ℓ₄ Rel E ℓ₅ Rel F ℓ₆ Set _
f Preserves₅ _ᵃ_ _ᵇ_ _ᶜ_ _ᵈ_ _ᵉ_ _ᶠ_ =
{a a b b c c d d e e}
a ∼ᵃ a b ∼ᵇ b c ∼ᶜ c d ∼ᵈ d e ∼ᵉ e
f a b c d e ∼ᶠ f a b c d e
_Preserves₆_⟶_⟶_⟶_⟶_⟶_⟶_ :
(A B C D E F G)
Rel A ℓ₁ Rel B ℓ₂ Rel C ℓ₃ Rel D ℓ₄ Rel E ℓ₅ Rel F ℓ₆
Rel G ℓ₇ Set _
𝑓 Preserves₆ _ᵃ_ _ᵇ_ _ᶜ_ _ᵈ_ _ᵉ_ _ᶠ_ _ᵍ_ =
{a a b b c c d d e e f f}
a ∼ᵃ a b ∼ᵇ b c ∼ᶜ c d ∼ᵈ d e ∼ᵉ e f ∼ᶠ f
𝑓 a b c d e f ∼ᵍ 𝑓 a b c d e f
open Rel public
using (Dec ; yes ; no ; ¬_ ; True ; False ; ⌊_⌋ ;
Pred ; Decidable₁ ; Rel ; Decidable₂)
open import Algebra.Core public
module Algebra where
open import Algebra.Definitions as D public
open import Algebra.Structures as S public
open import Algebra.Bundles as B public
module WithEq {a } {A : Set a} ( : Rel A ) where
open D public
open S public
open B public
module where
open import Relation.Binary.PropositionalEquality public hiding (module _≡_)
At : {a} (A : Set a) Rel A _
At A = _≡_ {A = A}
open public using (_≡_ ; _≢_ ; refl) renaming (At to ≡-At)
module IsDecidable where
record IsDecidable {a r} {A : Set a} (R : Rel A r) : Set (a r) where
field dec : Decidable₂ R
open IsDecidable public
private variable
a b : Level
A B : Set a
F : A Set b
n :
Eq : Set a Set _
Eq = IsDecidable ≡-At
_≟_ : Eq A Decidable₂ (≡-At A)
_≟_ = dec
instance
eq- : Eq
eq- .dec _ _ = yes refl
eq-⊥ : Eq
eq-⊥ .dec () ()
eq- : Eq
eq- .dec = ._≟_
eq-fin : Eq (Fin n)
eq-fin .dec = Fin._≟_
eq-bool : Eq Bool
eq-bool .dec = Bool._≟_
eq-⊎ : Eq A Eq B Eq (A B)
eq-⊎ .dec = ⊎.≡-dec dec dec
eq-prod : Eq A Rel.∀[ Eq F ] Eq (Σ A F)
eq-prod .dec = Σ.≡-dec dec dec
eq-maybe : Eq A Eq (Maybe A)
eq-maybe .dec = Maybe.≡-dec dec
eq-list : Eq A Eq (List A)
eq-list .dec = List.≡-dec dec
eq-vec : Eq A Eq (Vec A n)
eq-vec .dec = Vec.≡-dec dec
open IsDecidable public
using (IsDecidable ; dec ; Eq ; _≟_)
hiding (module IsDecidable)
module Uninhabited where
record Uninhabited {a} (A : Set a) : Set a where
field absurd : ¬ A
open Uninhabited
private variable
a b : Level
A B : Set a
absurd : Uninhabited A A B
absurd x with () absurd x