From 9a8dc821469836bce507e276059502ee921949e7 Mon Sep 17 00:00:00 2001 From: Chloe Brown Date: Thu, 23 Sep 2021 13:46:05 +0100 Subject: Define Motzkin numbers. --- src/Motzkin.agda | 147 +++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 147 insertions(+) create mode 100644 src/Motzkin.agda (limited to 'src/Motzkin.agda') diff --git a/src/Motzkin.agda b/src/Motzkin.agda new file mode 100644 index 0000000..d5aa67e --- /dev/null +++ b/src/Motzkin.agda @@ -0,0 +1,147 @@ + +module Motzkin where + +open import Data.Nat using (ℕ; zero; suc; _+_; _≤?_; z≤n) +open import Data.Nat.Properties using (≤-totalOrder) + +open import Data.Empty using (⊥-elim) +open import Data.Fin using (Fin; zero; suc; toℕ; raise) +open import Data.List using (List; []; _∷_; length; filter; cartesianProductWith; allFin; map) +open import Data.List.Membership.Propositional using (_∈_) +open import Data.List.Relation.Binary.Pointwise as Pointwise using (Pointwise; []; _∷_) +open import Data.List.Relation.Binary.Suffix as Suffix using (base; next) +open import Data.List.Relation.Unary.All as All using (All; []; _∷_) +open import Data.List.Relation.Unary.All.Properties as AllP using (all-filter) +open import Data.List.Relation.Unary.Any as Any using (Any; here; there) +open import Data.List.Relation.Unary.Any.Properties as AnyP using (lookup-result) +open import Data.List.Relation.Unary.Linked using ([]; [-]; _∷_) +open import Data.List.Relation.Unary.Sorted.TotalOrder ≤-totalOrder using (Sorted; sorted?) +open import Data.Product using (_×_; _,_; ∃₂; proj₁; proj₂; map₁) +open import Data.Product.Properties as Prod +open import Data.Sum using (_⊎_; [_,_]) +open import Data.Vec as Vec using (Vec; []; _∷_; _++_; replicate; updateAt; toList) +open import Data.Vec.Properties using (map-id) +open import Function using (_∘_; case_of_; id) +open import Function.Equivalence using (_⇔_; equivalence) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; cong; subst; setoid) +open import Relation.Nullary using (Dec; yes; no; ¬_) +import Relation.Nullary.Decidable as Dec +open import Relation.Nullary.Implication using (_→-dec_) +open import Relation.Unary using (Pred; Decidable) + +_≋_ : ∀ {n} → Rel (List (Fin n)) _ +_≋_ = Pointwise _≡_ + +Suffix : ∀ {n} → Rel (List (Fin n)) _ +Suffix = Suffix.Suffix _≡_ + +counts : ∀ {n} → (xs : List (Fin n)) → Vec ℕ n +counts [] = replicate 0 +counts (x ∷ xs) = updateAt x suc (counts xs) + +sorted-0s : ∀ {n} → (Sorted ∘ toList {n = n} ∘ replicate) 0 +sorted-0s {zero} = [] +sorted-0s {suc zero} = [-] +sorted-0s {suc (suc n)} = z≤n ∷ sorted-0s + +≋⇒≡ : ∀ {n} {xs ys : List (Fin n)} → xs ≋ ys → xs ≡ ys +≋⇒≡ [] = refl +≋⇒≡ (x≡y ∷ xs≋ys) = subst (λ _ → _ ≡ _ ∷ _) x≡y (cong (_ ∷_) (≋⇒≡ xs≋ys)) + +P : ∀ {n} → Pred (List (Fin n)) _ +P xs = ∀ ys → Suffix ys xs → (Sorted ∘ toList ∘ counts) ys + +data Motzkin {n} : List (Fin n) → Set where + base : Motzkin [] + next : ∀ {x xs} → (Sorted ∘ toList ∘ counts) (x ∷ xs) → Motzkin xs → Motzkin (x ∷ xs) + +P⇔Motzkin : ∀ {n} xs → P {n} xs ⇔ Motzkin xs +P⇔Motzkin xs = equivalence (⇒ xs) ⇐ + where + ⇒ : ∀ xs → P xs → Motzkin xs + ⇒ [] pxs = base + ⇒ (x ∷ xs) pxs = next (pxs (x ∷ xs) (Suffix.refl refl)) (⇒ xs (λ ys ys