agda-spa/Chain.agda
Danila Fedorin 46c084d24c Add the beginnings of a formalization of chains
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
2023-08-19 14:22:03 -07:00

25 lines
1.0 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 Chain where
open import Data.Nat as Nat using (; suc; _+_; _≤_)
open import Data.Product using (_×_; Σ; _,_)
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
module _ {a} {A : Set a} (_R_ : A A Set a) where
data Chain : A A Set a where
done : {a : A} Chain a a 0
step : {a₁ a₂ a₃ : A} {n : } a₁ R a₂ Chain a₂ a₃ n Chain a₁ a₃ (suc n)
concat : {a₁ a₂ a₃ : A} {n₁ n₂ : } Chain a₁ a₂ n₁ Chain a₂ a₃ n₂ Chain a₁ a₃ (n₁ + n₂)
concat done a₂a₃ = a₂a₃
concat (step a₁Ra aa₂) a₂a₃ = step a₁Ra (concat aa₂ a₂a₃)
empty-≡ : {a₁ a₂ : A} Chain a₁ a₂ 0 a₁ a₂
empty-≡ done = refl
Bounded : Set a
Bounded bound = {a₁ a₂ : A} {n : } Chain a₁ a₂ n n bound
Height : Set a
Height height = (Σ (A × A) (λ (a₁ , a₂) Chain a₁ a₂ height) × Bounded height)