Setoid.agda 8.7 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318
  1. module Setoid where
  2. module Logic where
  3. infix 4 _/\_
  4. -- infix 2 _\/_
  5. data True : Set where
  6. tt : True
  7. data False : Set where
  8. data _/\_ (P Q : Set) : Set where
  9. andI : P -> Q -> P /\ Q
  10. -- Not allowed if we have proof irrelevance
  11. -- data _\/_ (P Q : Set) : Set where
  12. -- orIL : P -> P \/ Q
  13. -- orIR : Q -> P \/ Q
  14. module Setoid where
  15. data Setoid : Set1 where
  16. setoid : (A : Set)
  17. -> (_==_ : A -> A -> Set)
  18. -> (refl : (x : A) -> x == x)
  19. -> (sym : (x y : A) -> x == y -> y == x)
  20. -> (trans : (x y z : A) -> x == y -> y == z -> x == z)
  21. -> Setoid
  22. El : Setoid -> Set
  23. El (setoid A _ _ _ _) = A
  24. module Projections where
  25. eq : (A : Setoid) -> El A -> El A -> Set
  26. eq (setoid _ e _ _ _) = e
  27. refl : (A : Setoid) -> {x : El A} -> eq A x x
  28. refl (setoid _ _ r _ _) = r _
  29. sym : (A : Setoid) -> {x y : El A} -> (h : eq A x y) -> eq A y x
  30. sym (setoid _ _ _ s _) = s _ _
  31. trans : (A : Setoid) -> {x y z : El A} -> eq A x y -> eq A y z -> eq A x z
  32. trans (setoid _ _ _ _ t) = t _ _ _
  33. module Equality (A : Setoid) where
  34. infix 6 _==_
  35. _==_ : El A -> El A -> Set
  36. _==_ = Projections.eq A
  37. refl : {x : El A} -> x == x
  38. refl = Projections.refl A
  39. sym : {x y : El A} -> x == y -> y == x
  40. sym = Projections.sym A
  41. trans : {x y z : El A} -> x == y -> y == z -> x == z
  42. trans = Projections.trans A
  43. module EqChain (A : Setoid.Setoid) where
  44. infixl 5 _===_ _=-=_
  45. infix 8 _since_
  46. open Setoid
  47. private open module EqA = Equality A
  48. eqProof>_ : (x : El A) -> x == x
  49. eqProof> x = refl
  50. _=-=_ : (x : El A) -> {y : El A} -> x == y -> x == y
  51. x =-= eq = eq
  52. _===_ : {x y z : El A} -> x == y -> y == z -> x == z
  53. _===_ = trans
  54. _since_ : {x : El A} -> (y : El A) -> x == y -> x == y
  55. _ since eq = eq
  56. module Fun where
  57. open Logic
  58. open Setoid
  59. infixr 10 _=>_ _==>_
  60. open Setoid.Projections using (eq)
  61. data _=>_ (A B : Setoid) : Set where
  62. lam : (f : El A -> El B)
  63. -> ({x y : El A} -> eq A x y
  64. -> eq B (f x) (f y)
  65. )
  66. -> A => B
  67. app : {A B : Setoid} -> (A => B) -> El A -> El B
  68. app (lam f _) = f
  69. cong : {A B : Setoid} -> (f : A => B) -> {x y : El A} ->
  70. eq A x y -> eq B (app f x) (app f y)
  71. cong (lam _ resp) = resp
  72. data EqFun {A B : Setoid}(f g : A => B) : Set where
  73. eqFunI : ({x y : El A} -> eq A x y -> eq B (app f x) (app g y)) ->
  74. EqFun f g
  75. eqFunE : {A B : Setoid} -> {f g : A => B} -> {x y : El A} ->
  76. EqFun f g -> eq A x y -> eq B (app f x) (app g y)
  77. eqFunE (eqFunI h) = h
  78. _==>_ : Setoid -> Setoid -> Setoid
  79. A ==> B = setoid (A => B) EqFun r s t
  80. where
  81. module Proof where
  82. open module EqChainB = EqChain B
  83. module EqA = Equality A
  84. open module EqB = Equality B
  85. -- either abstract or --proof-irrelevance needed
  86. -- (we don't want to compare the proofs for equality)
  87. -- abstract
  88. r : (f : A => B) -> EqFun f f
  89. r f = eqFunI (\xy -> cong f xy)
  90. s : (f g : A => B) -> EqFun f g -> EqFun g f
  91. s f g fg =
  92. eqFunI (\{x}{y} xy ->
  93. app g x =-= app g y since cong g xy
  94. === app f x since sym (eqFunE fg xy)
  95. === app f y since cong f xy
  96. )
  97. t : (f g h : A => B) -> EqFun f g -> EqFun g h -> EqFun f h
  98. t f g h fg gh =
  99. eqFunI (\{x}{y} xy ->
  100. app f x =-= app g y since eqFunE fg xy
  101. === app g x since cong g (EqA.sym xy)
  102. === app h y since eqFunE gh xy
  103. )
  104. open Proof
  105. infixl 100 _$_
  106. _$_ : {A B : Setoid} -> El (A ==> B) -> El A -> El B
  107. _$_ = app
  108. lam2 : {A B C : Setoid} ->
  109. (f : El A -> El B -> El C) ->
  110. ({x x' : El A} -> eq A x x' ->
  111. {y y' : El B} -> eq B y y' -> eq C (f x y) (f x' y')
  112. ) -> El (A ==> B ==> C)
  113. lam2 {A} f h = lam (\x -> lam (\y -> f x y)
  114. (\y -> h EqA.refl y))
  115. (\x -> eqFunI (\y -> h x y))
  116. where
  117. module EqA = Equality A
  118. lam3 : {A B C D : Setoid} ->
  119. (f : El A -> El B -> El C -> El D) ->
  120. ({x x' : El A} -> eq A x x' ->
  121. {y y' : El B} -> eq B y y' ->
  122. {z z' : El C} -> eq C z z' -> eq D (f x y z) (f x' y' z')
  123. ) -> El (A ==> B ==> C ==> D)
  124. lam3 {A} f h =
  125. lam (\x -> lam2 (\y z -> f x y z)
  126. (\y z -> h EqA.refl y z))
  127. (\x -> eqFunI (\y -> eqFunI (\z -> h x y z)))
  128. where
  129. module EqA = Equality A
  130. eta : {A B : Setoid} -> (f : El (A ==> B)) ->
  131. eq (A ==> B) f (lam (\x -> f $ x) (\xy -> cong f xy))
  132. eta f = eqFunI (\xy -> cong f xy)
  133. id : {A : Setoid} -> El (A ==> A)
  134. id = lam (\x -> x) (\x -> x)
  135. {- Now it looks okay. But it's incredibly slow! Proof irrelevance makes it
  136. go fast again... The problem is equality checking of (function type)
  137. setoids which without proof irrelevance checks equality of the proof that
  138. EqFun is an equivalence relation. It's not clear why using lam3 involves
  139. so many more equality checks than using lam. Making the proofs abstract
  140. makes the problem go away.
  141. -}
  142. compose : {A B C : Setoid} -> El ((B ==> C) ==> (A ==> B) ==> (A ==> C))
  143. compose =
  144. lam3 (\f g x -> f $ (g $ x))
  145. (\f g x -> eqFunE f (eqFunE g x))
  146. _∘_ : {A B C : Setoid} -> El (B ==> C) -> El (A ==> B) -> El (A ==> C)
  147. f ∘ g = compose $ f $ g
  148. const : {A B : Setoid} -> El (A ==> B ==> A)
  149. const = lam2 (\x y -> x) (\x y -> x)
  150. module Nat where
  151. open Logic
  152. open Setoid
  153. open Fun
  154. infixl 10 _+_
  155. data Nat : Set where
  156. zero : Nat
  157. suc : Nat -> Nat
  158. module NatSetoid where
  159. eqNat : Nat -> Nat -> Set
  160. eqNat zero zero = True
  161. eqNat zero (suc _) = False
  162. eqNat (suc _) zero = False
  163. eqNat (suc n) (suc m) = eqNat n m
  164. data EqNat (n m : Nat) : Set where
  165. eqnat : eqNat n m -> EqNat n m
  166. uneqnat : {n m : Nat} -> EqNat n m -> eqNat n m
  167. uneqnat (eqnat x) = x
  168. r : (x : Nat) -> eqNat x x
  169. r zero = tt
  170. r (suc n) = r n
  171. -- reflexivity of EqNat
  172. rf : (n : Nat) -> EqNat n n
  173. rf = \ x -> eqnat (r x)
  174. s : (x y : Nat) -> eqNat x y -> eqNat y x
  175. s zero zero _ = tt
  176. s (suc n) (suc m) h = s n m h
  177. s zero (suc _) ()
  178. s (suc _) zero ()
  179. -- symmetry of EqNat
  180. sy : (x y : Nat) -> EqNat x y -> EqNat y x
  181. sy = \x y h -> eqnat (s x y (uneqnat h))
  182. t : (x y z : Nat) -> eqNat x y -> eqNat y z -> eqNat x z
  183. t zero zero z xy yz = yz
  184. t (suc x) (suc y) (suc z) xy yz = t x y z xy yz
  185. t zero (suc _) _ () _
  186. t (suc _) zero _ () _
  187. t (suc _) (suc _) zero _ ()
  188. -- transitivity of EqNat
  189. tr : (x y z : Nat) -> EqNat x y -> EqNat y z -> EqNat x z
  190. tr = \x y z xy yz -> eqnat (t x y z (uneqnat xy) (uneqnat yz))
  191. NAT : Setoid
  192. NAT = setoid Nat NatSetoid.EqNat NatSetoid.rf NatSetoid.sy NatSetoid.tr
  193. _+_ : Nat -> Nat -> Nat
  194. zero + m = m
  195. suc n + m = suc (n + m)
  196. plus : El (NAT ==> NAT ==> NAT)
  197. plus = lam2 (\n m -> n + m) eqPlus
  198. where
  199. module EqNat = Equality NAT
  200. open EqNat
  201. open NatSetoid
  202. eqPlus : {n n' : Nat} -> n == n' -> {m m' : Nat} -> m == m' -> n + m == n' + m'
  203. eqPlus {zero} {zero} _ mm = mm
  204. eqPlus {suc n} {suc n'} (eqnat nn) {m}{m'} (eqnat mm) =
  205. eqnat (uneqnat (eqPlus{n}{n'} (eqnat nn)
  206. {m}{m'} (eqnat mm)
  207. ) )
  208. eqPlus {zero} {suc _} (eqnat ()) _
  209. eqPlus {suc _} {zero} (eqnat ()) _
  210. module List where
  211. open Logic
  212. open Setoid
  213. data List (A : Set) : Set where
  214. nil : List A
  215. _::_ : A -> List A -> List A
  216. LIST : Setoid -> Setoid
  217. LIST A = setoid (List (El A)) eqList r s t
  218. where
  219. module EqA = Equality A
  220. open EqA
  221. eqList : List (El A) -> List (El A) -> Set
  222. eqList nil nil = True
  223. eqList nil (_ :: _) = False
  224. eqList (_ :: _) nil = False
  225. eqList (x :: xs) (y :: ys) = x == y /\ eqList xs ys
  226. r : (x : List (El A)) -> eqList x x
  227. r nil = tt
  228. r (x :: xs) = andI refl (r xs)
  229. s : (x y : List (El A)) -> eqList x y -> eqList y x
  230. s nil nil h = h
  231. s (x :: xs) (y :: ys) (andI xy xys) = andI (sym xy) (s xs ys xys)
  232. s nil (_ :: _) ()
  233. s (_ :: _) nil ()
  234. t : (x y z : List (El A)) -> eqList x y -> eqList y z -> eqList x z
  235. t nil nil zs _ h = h
  236. t (x :: xs) (y :: ys) (z :: zs) (andI xy xys) (andI yz yzs) =
  237. andI (trans xy yz) (t xs ys zs xys yzs)
  238. t nil (_ :: _) _ () _
  239. t (_ :: _) nil _ () _
  240. t (_ :: _) (_ :: _) nil _ ()
  241. open Fun