1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-12-13 13:53:52 +01:00

Added Agda code for some HITs

This commit is contained in:
Niels
2017-05-22 16:46:58 +02:00
parent 27d78e1e75
commit b85976a96d
12 changed files with 900 additions and 0 deletions

View File

@@ -0,0 +1,63 @@
{-# OPTIONS --without-K --rewriting #-}
open import HoTT
module Integers where
private
data Integers : Set where
z : Integers
S : Integers -> Integers
P : Integers -> Integers
Ints : Set
Ints = Integers
nul : Ints
nul = z
Succ : Ints -> Ints
Succ = S
Pred : Ints -> Ints
Pred = P
postulate
invl : (x : Integers) -> P(S x) == x
invr : (x : Integers) -> S(P x) == x
trunc : is-set Ints
Zind : (Y : Integers -> Set)
-> (zY : Y z)
-> (SY : (x : Integers) -> Y x -> Y(S x))
-> (PY : (x : Integers) -> Y x -> Y(P x))
-> (invYl : (x : Integers) (y : Y x) -> PathOver Y (invl x) (PY (S x) (SY x y)) y)
-> (invYr : (x : Integers) (y : Y x) -> PathOver Y (invr x) (SY (P x) (PY x y)) y)
-> (t : (x : Integers) -> is-set (Y x))
-> (x : Integers) -> Y x
Zind Y zY SY PY invYl invYr t z = zY
Zind Y zY SY PY invYl invYr t (S x) = SY x (Zind Y zY SY PY invYl invYr t x)
Zind Y zY SY PY invYl invYr t (P x) = PY x (Zind Y zY SY PY invYl invYr t x)
postulate
Zind-βinvl :
(Y : Integers -> Set)
-> (zY : Y z)
-> (SY : (x : Integers) -> Y x -> Y(S x))
-> (PY : (x : Integers) -> Y x -> Y(P x))
-> (invYl : (x : Integers) (y : Y x) -> PathOver Y (invl x) (PY (S x) (SY x y)) y)
-> (invYr : (x : Integers) (y : Y x) -> PathOver Y (invr x) (SY (P x) (PY x y)) y)
-> (t : (x : Integers) -> is-set (Y x))
-> (x : Integers)
-> apd (Zind Y zY SY PY invYl invYr t) (invl x) == invYl x (Zind Y zY SY PY invYl invYr t x)
Zind-βinvr :
(Y : Integers -> Set)
-> (zY : Y z)
-> (SY : (x : Integers) -> Y x -> Y(S x))
-> (PY : (x : Integers) -> Y x -> Y(P x))
-> (invYl : (x : Integers) (y : Y x) -> PathOver Y (invl x) (PY (S x) (SY x y)) y)
-> (invYr : (x : Integers) (y : Y x) -> PathOver Y (invr x) (SY (P x) (PY x y)) y)
-> (t : (x : Integers) -> is-set (Y x))
-> (x : Integers)
-> apd (Zind Y zY SY PY invYl invYr t) (invr x) == invYr x (Zind Y zY SY PY invYl invYr t x)