2024-02-10 16:51:43 -08:00
|
|
|
|
module Utils where
|
|
|
|
|
|
2024-03-07 21:53:45 -08:00
|
|
|
|
open import Agda.Primitive using () renaming (_⊔_ to _⊔ℓ_)
|
2024-02-11 14:16:42 -08:00
|
|
|
|
open import Data.Nat using (ℕ; suc)
|
2024-02-10 16:51:43 -08:00
|
|
|
|
open import Data.List using (List; []; _∷_; _++_)
|
|
|
|
|
open import Data.List.Membership.Propositional using (_∈_)
|
2024-02-11 12:45:33 -08:00
|
|
|
|
open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
|
2024-02-10 16:51:43 -08:00
|
|
|
|
open import Data.List.Relation.Unary.Any using (Any; here; there) -- TODO: re-export these with nicer names from map
|
2024-02-11 12:45:33 -08:00
|
|
|
|
open import Relation.Binary.PropositionalEquality using (_≡_; sym; refl)
|
2024-02-10 16:51:43 -08:00
|
|
|
|
open import Relation.Nullary using (¬_)
|
|
|
|
|
|
|
|
|
|
data Unique {c} {C : Set c} : List C → Set c where
|
|
|
|
|
empty : Unique []
|
|
|
|
|
push : ∀ {x : C} {xs : List C}
|
|
|
|
|
→ All (λ x' → ¬ x ≡ x') xs
|
|
|
|
|
→ Unique xs
|
|
|
|
|
→ Unique (x ∷ xs)
|
|
|
|
|
|
|
|
|
|
Unique-append : ∀ {c} {C : Set c} {x : C} {xs : List C} →
|
|
|
|
|
¬ x ∈ xs → Unique xs → Unique (xs ++ (x ∷ []))
|
|
|
|
|
Unique-append {c} {C} {x} {[]} _ _ = push [] empty
|
|
|
|
|
Unique-append {c} {C} {x} {x' ∷ xs'} x∉xs (push x'≢ uxs') =
|
|
|
|
|
push (help x'≢) (Unique-append (λ x∈xs' → x∉xs (there x∈xs')) uxs')
|
|
|
|
|
where
|
|
|
|
|
x'≢x : ¬ x' ≡ x
|
|
|
|
|
x'≢x x'≡x = x∉xs (here (sym x'≡x))
|
|
|
|
|
|
|
|
|
|
help : {l : List C} → All (λ x'' → ¬ x' ≡ x'') l → All (λ x'' → ¬ x' ≡ x'') (l ++ (x ∷ []))
|
|
|
|
|
help {[]} _ = x'≢x ∷ []
|
|
|
|
|
help {e ∷ es} (x'≢e ∷ x'≢es) = x'≢e ∷ help x'≢es
|
|
|
|
|
|
|
|
|
|
All¬-¬Any : ∀ {p c} {C : Set c} {P : C → Set p} {l : List C} → All (λ x → ¬ P x) l → ¬ Any P l
|
|
|
|
|
All¬-¬Any {l = x ∷ xs} (¬Px ∷ _) (here Px) = ¬Px Px
|
|
|
|
|
All¬-¬Any {l = x ∷ xs} (_ ∷ ¬Pxs) (there Pxs) = All¬-¬Any ¬Pxs Pxs
|
2024-02-11 12:45:33 -08:00
|
|
|
|
|
2024-03-01 19:08:11 -08:00
|
|
|
|
All-single : ∀ {p c} {C : Set c} {P : C → Set p} {c : C} {l : List C} → All P l → c ∈ l → P c
|
|
|
|
|
All-single {c = c} {l = x ∷ xs} (p ∷ ps) (here refl) = p
|
|
|
|
|
All-single {c = c} {l = x ∷ xs} (p ∷ ps) (there c∈xs) = All-single ps c∈xs
|
|
|
|
|
|
2024-02-11 12:45:33 -08:00
|
|
|
|
All-x∈xs : ∀ {a} {A : Set a} (xs : List A) → All (λ x → x ∈ xs) xs
|
|
|
|
|
All-x∈xs [] = []
|
|
|
|
|
All-x∈xs (x ∷ xs') = here refl ∷ map there (All-x∈xs xs')
|
2024-02-11 14:16:42 -08:00
|
|
|
|
|
|
|
|
|
iterate : ∀ {a} {A : Set a} (n : ℕ) → (f : A → A) → A → A
|
|
|
|
|
iterate 0 _ a = a
|
|
|
|
|
iterate (suc n) f a = f (iterate n f a)
|
2024-03-07 21:53:45 -08:00
|
|
|
|
|
|
|
|
|
data Pairwise {a} {b} {c} {A : Set a} {B : Set b} (P : A → B → Set c) : List A → List B → Set (a ⊔ℓ b ⊔ℓ c) where
|
|
|
|
|
[] : Pairwise P [] []
|
|
|
|
|
_∷_ : ∀ {x : A} {y : B} {xs : List A} {ys : List B} →
|
|
|
|
|
P x y → Pairwise P xs ys →
|
|
|
|
|
Pairwise P (x ∷ xs) (y ∷ ys)
|