Skip to content
This repository has been archived by the owner on Nov 6, 2020. It is now read-only.

Commit

Permalink
Sn and revamped tools for higher-order paths
Browse files Browse the repository at this point in the history
(1) Sn and its dependent elimination rule
(2) Revamped library for higher-order paths
  • Loading branch information
favonia committed Apr 24, 2012
1 parent 609a4af commit a3795e7
Show file tree
Hide file tree
Showing 6 changed files with 269 additions and 113 deletions.
File renamed without changes.
122 changes: 53 additions & 69 deletions Path/Higher-order.agda
Expand Up @@ -9,17 +9,15 @@ module Path.Higher-order where
open import Prelude
open import Path
open import Path.Lemmas
open import Space.Nat.Lemmas

-- TODO Generic boring paths for the case that B is a constant function.
-- TODO Build a nicer interface for Env. (Although the hidden definition
-- the one suitable for induction...) Make sure definitional equalities
-- hold for all finite cases.
-- TODO Generalized cong[dep]-const lemma

------------------------------------------------------------------------
-- Higher-order loop spaces

Ω : {ℓ} n {A : Set ℓ} A Set
Ω-refl : {ℓ} n {A : Set ℓ} (base : A) Ω n base
private Ω-refl : {ℓ} n {A : Set ℓ} (base : A) Ω n base

Ω 0 {A} _ = A
Ω (suc n) {A} base = Ω-refl n base ≡ Ω-refl n base
Expand All @@ -30,71 +28,57 @@ open import Path.Lemmas
------------------------------------------------------------------------
-- Higher-order path spaces

-- This definition is not suitable for inductive definition of cong[dep]⇑
--Paths⇑ : ∀ {ℓ} n → Set ℓ → Set ℓ
--Paths⇑ 0 A = A
--Paths⇑ (suc i) A = Σ A (λ x → Σ A (λ y → Paths⇑ i (x ≡ y)))

-- Env 0 = tt
-- Env 1 = (tt , (x₁ , y₁))
-- Env 2 = ((tt , (x₁ , y₁)) , (x₂ , y₂))
-- Env 3 = (((tt , (x₁ , y₁)) , (x₂ , y₂)) , (x₃ , y₃))
private
Env⇑ : {ℓ} n Set Set
Path⇑ : {ℓ} n (A : Set ℓ) Env⇑ n A Set

Env⇑ {ℓ} 0 A = ↑ ℓ ⊤
Env⇑ {ℓ} 1 A = A × A
Env⇑ (suc (suc i)) A = Σ (Env⇑ i A) (λ e Path⇑ i A e × Path⇑ i A e)

Path⇑ 0 A _ = A
Path⇑ 1 A (x , y) = x ≡ y
Path⇑ (suc (suc i)) A (_ , (x , y)) = x ≡ y

--usable but useless
--Paths⇑ : ∀ {ℓ} n → Set ℓ → Set ℓ
--Paths⇑ n A = Σ (Env⇑ n A) (Path⇑ n A)

-- These interfaces will be changed
Cong[dep]-env⇑ : {ℓ} n Set Set
Cong[dep]-env⇑ = Env⇑

Cong[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) (f : (x : A) B x)
(e : Env⇑ n A) (p : Path⇑ n A e) Set ℓ₂
-- Endpoints[d] 0 = tt
-- Endpoints[d] 1 = (tt, x₁ , y₁)
-- Endpoints[d] 2 = ((tt, x₁ , y₂) , x₂ , y₂)

-- This definition is perfect for induction but probably not user-friendly
Endpoints⇑ : {ℓ} n Set Set
Path⇑ : {ℓ} n {A : Set ℓ} Endpoints⇑ n A Set

Endpoints⇑ {ℓ} 0 A = ↑ ℓ ⊤
Endpoints⇑ (suc n) A = Σ (Endpoints⇑ n A) (λ c Path⇑ n c × Path⇑ n c)

Path⇑ 0 {A} _ = A
Path⇑ (suc n) (_ , x , y) = x ≡ y

-- This definition is terrible for induction, but seems nicer to use?
Endpoints′⇑ : {ℓ} n Set Set
Endpoints′⇑ {ℓ} 0 A = ↑ ℓ ⊤
Endpoints′⇑ (suc n) A = Σ A (λ x Σ A ( λ y Endpoints′⇑ n (x ≡ y)))

-- A helper function to convert the above definition
endpoints′⇒endpoints⇑ : {ℓ} n m {A : Set ℓ} (cA : Endpoints⇑ n A) Endpoints′⇑ m (Path⇑ n cA) Endpoints⇑ (m + n) A
endpoints′⇒endpoints⇑ n 0 c _ = c
endpoints′⇒endpoints⇑ n (suc m) {A} c (x , y , p) =
subst (λ n Endpoints⇑ n A) (n+suc m n) $ endpoints′⇒endpoints⇑ (suc n) m (_ , (x , y)) p

-- A section that parametrized by a high-order path
-- Endpoints[d] 0 = tt
-- Endpoints[d] 1 = (tt, bx₁ , by₁)
-- Endpoints[d] 2 = ((tt, bx₁ , by₂) , bx₂ , by₂)
Endpoints[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) Endpoints⇑ n A Set ℓ₂
Path[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) {cA : Endpoints⇑ n A} Path⇑ n cA Endpoints[dep]⇑ n B cA Set ℓ₂

Endpoints[dep]⇑ {ℓ₁} {ℓ₂} 0 _ _ = ↑ ℓ₂ ⊤
Endpoints[dep]⇑ (suc n) B (cA , x , y) = Σ (Endpoints[dep]⇑ n B cA)
(λ c Path[dep]⇑ n B x c × Path[dep]⇑ n B y c)

Path[dep]⇑ 0 B a _ = B a
Path[dep]⇑ (suc n) B a (cB , bx , by) = subst (λ x Path[dep]⇑ n B x cB) a bx ≡ by

-- Higher-order cong(ruence) that takes high-order paths
-- Reuse Path[dep] to avoid all sorts of problems in Space.Sphere
cong-endpoints[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) (f : (x : A) B x)
(cA : Endpoints⇑ n A) Endpoints[dep]⇑ n B cA
cong[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) (f : (x : A) B x)
(e : Env⇑ n A) (p : Path⇑ n A e) Cong[dep]⇑ n B f _ p
(cA : Endpoints⇑ n A) (pA : Path⇑ n cA) Path[dep]⇑ n B pA (cong-endpoints[dep]⇑ n B f cA)

Cong[dep]⇑ 0 B f _ p = B p
Cong[dep]⇑ 1 B f (x , y) p = subst B p (f x) ≡ (f y)
Cong[dep]⇑ (suc (suc i)) B f (_ , (x , y)) p = subst (Cong[dep]⇑ i B f _) p (cong[dep]⇑ i B f _ x) ≡ cong[dep]⇑ i B f _ y
cong-endpoints[dep]⇑ 0 B f _ = lift tt
cong-endpoints[dep]⇑ (suc n) B f (cA , x , y) = (cong-endpoints[dep]⇑ n B f cA ,
cong[dep]⇑ n B f cA x ,
cong[dep]⇑ n B f cA y)

cong[dep]⇑ 0 B f _ p = f p
cong[dep]⇑ 1 B f (x , y) p = cong[dep] B f p
cong[dep]⇑ (suc (suc i)) B f (_ , (x , y)) p = cong[dep] (Cong[dep]⇑ i B f _) (cong[dep]⇑ i B f _) p

-- Env 0 = tt
-- Env 1 = (tt , ((ax₁ , bx₁) , (ay₁ , by₁)))
-- Env 2 = ((tt , ((ax₁ , bx₁) , (ay₁ , by₂))) , ((ax₂ , bx₂) , (ay₂ , by₂)))
private
Env[d]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) Set (ℓ₁ ⊔ ℓ₂)
PathA[d]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) Env[d]⇑ n B Set ℓ₁
PathB[d]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) (e : Env[d]⇑ n B) PathA[d]⇑ n B e Set ℓ₂

Env[d]⇑ {ℓ₁} {ℓ₂} 0 B = ↑ (ℓ₁ ⊔ ℓ₂) ⊤
Env[d]⇑ (suc i) B = Σ (Env[d]⇑ i B) (λ e Σ (PathA[d]⇑ i B e) (PathB[d]⇑ i B e) ×
Σ (PathA[d]⇑ i B e) (PathB[d]⇑ i B e))

PathA[d]⇑ 0 {A} _ _ = A
PathA[d]⇑ (suc i) _ (_ , ((ax , _) , (ay , _))) = ax ≡ ay

PathB[d]⇑ 0 B _ a = B a
PathB[d]⇑ (suc i) B (_ , ((_ , bx) , (_ , by))) a = subst (PathB[d]⇑ i B _) a bx ≡ by

-- This interface will be changed
PathB[dep]⇑-env : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) Set (ℓ₁ ⊔ ℓ₂)
PathB[dep]⇑-env = Env[d]⇑

-- This interface will be changed
PathB[dep]⇑ : {ℓ₁ ℓ₂} n {A : Set ℓ₁} (B : A Set ℓ₂) (e : Env[d]⇑ n B) PathA[d]⇑ n B e Set ℓ₂
PathB[dep]⇑ = PathB[d]⇑
cong[dep]⇑ 0 B f t p = f p
cong[dep]⇑ (suc n) B f (cA , (x , y)) p = cong[dep] (λ x Path[dep]⇑ n B x (cong-endpoints[dep]⇑ n B f cA)) (cong[dep]⇑ n B f cA) p

12 changes: 5 additions & 7 deletions Path/Omega2-abelian.agda
Expand Up @@ -23,20 +23,18 @@ module Path.Omega2-abelian {ℓ} {A : Set ℓ} (base : A) where
open import Prelude hiding (_∘_)
open import Path
open import Path.Lemmas
open import Path.Higher-order

private
Ω₁A = base ≡ base
Ω₂A = refl base ≡ refl base

lemma₁ : (p : Ω₂A) cong (λ p′ trans p′ (refl base)) p ≡ p
lemma₁ : (p : Ω 2 base) cong (λ p′ trans p′ (refl base)) p ≡ p
lemma₁ p = elim″ (λ {x} p cong (λ p trans p (refl _)) p ≡ trans (trans-reflʳ x) p) (refl _) p

lemma₂ : {ℓ₁ ℓ₂ ℓ₃} {A : Set ℓ₁} {B : Set ℓ₂} {C : Set ℓ₃}
(f : A B C) {x y : A} {u v : B}
(p : x ≡ y) (q : u ≡ v) cong₂ f p q ≡ cong₂′ f p q
lemma₂ f p q = elim (λ {_ _} p cong₂ f p q ≡ cong₂′ f p q) (λ _ sym $ trans-reflʳ _) p

abelian : (p q : Ω₂A) trans p q ≡ trans q p
abelian : (p q : Ω 2 base) trans p q ≡ trans q p
abelian p q =
trans p q ≡⟨ cong (trans p) $ sym $ cong-id q ⟩
trans p (cong id q) ≡⟨ cong (flip trans (cong id q)) $ sym $ lemma₁ p ⟩
Expand Down Expand Up @@ -93,14 +91,14 @@ private
_∙_ = cong₂ trans
unit = refl (refl base)

trans-unital : IsUnital Ω₂A _∘_ unit
trans-unital : IsUnital 2 base) _∘_ unit
trans-unital = record
{ unitʳ = trans-reflʳ
; unitˡ = λ _ refl _
--; assoc = trans-assoc
}

cong₂-trans-unital : IsUnital Ω₂A _∙_ unit
cong₂-trans-unital : IsUnital 2 base) _∙_ unit
cong₂-trans-unital = record
{ unitʳ = unitʳ
; unitˡ = unitˡ
Expand Down
10 changes: 8 additions & 2 deletions README.agda
Expand Up @@ -5,10 +5,12 @@
------------------------------------------------------------------------

-- Copyright (c) 2012 Favonia
-- Copyright (c) 2011-2012 Nils Anders Danielsson
-- released under BSD-like license (See LICENSE)

-- A large portion of code was shamelessly copied from Nils' library
-- A large portion of code was shamelessly copied from Nils Anders Danielsson'
-- library released under BSD-like license (See LICENCE-Nils)
-- http://www.cse.chalmers.se/~nad/repos/equality/
-- Copyright (c) 2011-2012 Nils Anders Danielsson

{-# OPTIONS --without-K #-}

Expand Down Expand Up @@ -78,6 +80,10 @@ import Space.Integer
-- Definition of intervals
import Space.Interval

-- Some basic facts about Nat
-- (Definition of Nat is in the Prelude)
import Space.Nat.Lemmas

-- Definition of spheres (base + loop)
import Space.Sphere

Expand Down
20 changes: 20 additions & 0 deletions Space/Nat/Lemmas.agda
@@ -0,0 +1,20 @@
------------------------------------------------------------------------
-- Basic facts about natural numbers.
------------------------------------------------------------------------

-- Copyright (c) 2012 Favonia

{-# OPTIONS --without-K #-}

module Space.Nat.Lemmas where

open import Prelude
open import Path

n+0 : n n + 0 ≡ n
n+0 0 = refl _
n+0 (suc n) = cong suc $ n+0 n

n+suc : n m n + suc m ≡ suc n + m
n+suc 0 m = refl _
n+suc (suc n) m = cong suc $ n+suc n m

0 comments on commit a3795e7

Please sign in to comment.