univ.agda 14 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380
  1. {-# OPTIONS --no-positivity-check #-}
  2. module univ where
  3. open import Base
  4. open import Nat
  5. import Logic.ChainReasoning
  6. module Chain
  7. {A : Set}( _==_ : A -> A -> Set)
  8. (refl : {x : A} -> x == x)
  9. (trans : {x y z : A} -> x == y -> y == z -> x == z) =
  10. Logic.ChainReasoning.Mono.Homogenous _==_ (\x -> refl) (\x y z -> trans)
  11. -- mutual inductive recursive definition of S and the functions _=S_, El, eq,
  12. -- and all the proofs on these functions
  13. mutual
  14. infix 40 _==_ _=S_ _=Fam_
  15. infixr 80 _<<_
  16. infixl 80 _>>_
  17. infixl 150 _!_
  18. data S : Set where
  19. nat : S
  20. pi : (A : S)(F : Fam A) -> S
  21. sigma : (A : S)(F : Fam A) -> S
  22. data Fam (A : S) : Set where
  23. fam : (F : El A -> S) -> Map _==_ _=S_ F -> Fam A
  24. _=S'_ : rel S
  25. nat =S' pi _ _ = False
  26. nat =S' sigma _ _ = False
  27. pi _ _ =S' nat = False
  28. pi _ _ =S' sigma _ _ = False
  29. sigma _ _ =S' nat = False
  30. sigma _ _ =S' pi _ _ = False
  31. nat =S' nat = True
  32. pi A F =S' pi B G =
  33. (B =S A) * \ B=A -> F =Fam G >> B=A
  34. sigma A F =S' sigma B G =
  35. (A =S B) * \A=B -> F >> A=B =Fam G
  36. data _=S_ (A B : S) : Set where
  37. eqS : A =S' B -> A =S B
  38. El' : S -> Set
  39. El' nat = Nat
  40. El' (pi A F) =
  41. ((x : El A) -> El (F ! x)) * \f ->
  42. {x y : El A}(x=y : x == y) -> f x == pFam F x=y << f y
  43. El' (sigma A F) =
  44. El A * \x -> El (F ! x)
  45. data El (A : S) : Set where
  46. el : El' A -> El A
  47. _=='_ : {A : S} -> rel (El A)
  48. _=='_ {nat} (el x) (el y) = x =N y
  49. _=='_ {pi A F} (el < f , pf >) (el < g , pg >) =
  50. (x : El A) -> f x == g x
  51. _=='_ {sigma A F} (el < x , Fx >) (el < y , Fy >) =
  52. x == y * \x=y -> Fx == pFam F x=y << Fy
  53. data _==_ {A : S}(x y : El A) : Set where
  54. eq : x ==' y -> x == y
  55. _=Fam_ : {A : S} -> rel (Fam A)
  56. F =Fam G = (x : El _) -> F ! x =S G ! x
  57. _!_ : {A : S} -> Fam A -> El A -> S
  58. fam F _ ! x = F x
  59. pFam : {A : S}(F : Fam A) -> Map _==_ _=S_ (_!_ F)
  60. pFam (fam F pF) = pF
  61. -- Families are contravariant so they cast in the other direction.
  62. _>>_ : {A B : S} -> Fam A -> A =S B -> Fam B
  63. _>>_ {A}{B} F A=B = fam G pG
  64. where
  65. G : El B -> S
  66. G x = F ! (A=B << x)
  67. pG : Map _==_ _=S_ G
  68. pG x=y = pFam F (p<< A=B x=y)
  69. pfiFam : {A B : S}(p q : A =S B)(F : Fam A) -> F >> p =Fam F >> q
  70. pfiFam p q F x = pFam F (pfi p q x)
  71. _<<_ : {A B : S} -> A =S B -> El B -> El A
  72. _<<_ {nat }{pi _ _ } (eqS ()) _
  73. _<<_ {nat }{sigma _ _ } (eqS ()) _
  74. _<<_ {pi _ _ }{nat } (eqS ()) _
  75. _<<_ {pi _ _ }{sigma _ _ } (eqS ()) _
  76. _<<_ {sigma _ _ }{nat } (eqS ()) _
  77. _<<_ {sigma _ _ }{pi _ _ } (eqS ()) _
  78. _<<_ {nat }{nat } p x = x
  79. _<<_ {pi A F }{pi B G } (eqS < B=A , F=G >) (el < g , pg >) =
  80. el < f , (\{x}{y} -> pf x y) >
  81. where
  82. f : (x : El A) -> El (F ! x)
  83. f x = F=G x << g (B=A << x)
  84. pf : (x y : El A)(x=y : x == y) -> f x == pFam F x=y << f y
  85. pf x y x=y =
  86. chain> F=G x << g (B=A << x)
  87. === F=G x << _ << g (B=A << y) by p<< _ (pg (p<< B=A x=y))
  88. === pFam F x=y << F=G y << g (B=A << y) by pfi2 _ _ _ _ _
  89. where
  90. open module C = Chain _==_ (ref {_}) (trans {_})
  91. _<<_ {sigma A F}{sigma B G} (eqS < A=B , F=G >) (el < y , Gy >) =
  92. el < A=B << y , F=G y << Gy >
  93. p<< : {A B : S}(A=B : A =S B) -> Map _==_ _==_ (_<<_ A=B)
  94. p<< {nat}{nat} _ x=y = x=y
  95. p<< {pi A F} {pi B G} (eqS < B=A , F=G >)
  96. {el < f , pf >} {el < g , pg >} (eq f=g) = eq cf=cg
  97. where
  98. cf=cg : (x : El A) -> F=G x << f (B=A << x) == F=G x << g (B=A << x)
  99. cf=cg x = p<< (F=G x) (f=g (B=A << x))
  100. p<< {sigma A F}{sigma B G}(eqS < A=B , F=G >)
  101. {el < x , Gx >}{el < y , Gy >} (eq < x=y , Gx=Gy >) =
  102. eq < cx=cy , cGx=cGy >
  103. where
  104. cx=cy : A=B << x == A=B << y
  105. cx=cy = p<< A=B x=y
  106. cGx=cGy : F=G x << Gx == pFam F cx=cy << F=G y << Gy
  107. cGx=cGy =
  108. chain> F=G x << Gx
  109. === F=G x << pFam G x=y << Gy by p<< (F=G x) Gx=Gy
  110. === pFam F cx=cy << F=G y << Gy by pfi2 _ _ _ _ Gy
  111. where
  112. open module C = Chain _==_ (ref {_}) (trans {_})
  113. p<< {nat }{pi _ _ } (eqS ()) _
  114. p<< {nat }{sigma _ _} (eqS ()) _
  115. p<< {pi _ _ }{nat } (eqS ()) _
  116. p<< {pi _ _ }{sigma _ _} (eqS ()) _
  117. p<< {sigma _ _}{nat } (eqS ()) _
  118. p<< {sigma _ _}{pi _ _ } (eqS ()) _
  119. refS : Refl _=S_
  120. refS {nat} = eqS T
  121. refS {pi A F} = eqS < refS , (\x -> symS (pFam F (ref<< x))) >
  122. refS {sigma A F} = eqS < refS , (\x -> pFam F (ref<< x)) >
  123. transS : Trans _=S_
  124. transS {nat }{nat }{pi _ _ } _ (eqS ())
  125. transS {nat }{nat }{sigma _ _ } _ (eqS ())
  126. transS {nat }{pi _ _ } (eqS ()) _
  127. transS {nat }{sigma _ _ } (eqS ()) _
  128. transS {pi _ _ }{nat } (eqS ()) _
  129. transS {pi _ _ }{pi _ _ }{nat } _ (eqS ())
  130. transS {pi _ _ }{pi _ _ }{sigma _ _ } _ (eqS ())
  131. transS {pi _ _ }{sigma _ _ } (eqS ()) _
  132. transS {sigma _ _ }{nat } (eqS ()) _
  133. transS {sigma _ _ }{pi _ _ } (eqS ()) _
  134. transS {sigma _ _ }{sigma _ _ }{nat } _ (eqS ())
  135. transS {sigma _ _ }{sigma _ _ }{pi _ _ } _ (eqS ())
  136. transS {nat}{nat}{nat} p q = p
  137. transS {pi A F}{pi B G}{pi C H}
  138. (eqS < B=A , F=G >) (eqS < C=B , G=H >) = eqS < C=A , F=H >
  139. where
  140. open module C = Chain _=S_ refS transS
  141. C=A = transS C=B B=A
  142. F=H : F =Fam H >> C=A
  143. F=H x =
  144. chain> F ! x
  145. === G ! (B=A << x) by F=G x
  146. === H ! (C=B << B=A << x) by G=H (B=A << x)
  147. === H ! (C=A << x) by pFam H (sym (trans<< C=B B=A x))
  148. transS {sigma A F}{sigma B G}{sigma C H}
  149. (eqS < A=B , F=G >)(eqS < B=C , G=H >) = eqS < A=C , F=H >
  150. where
  151. open module C = Chain _=S_ refS transS
  152. A=C = transS A=B B=C
  153. F=H : F >> A=C =Fam H
  154. F=H x =
  155. chain> F ! (A=C << x)
  156. === F ! (A=B << B=C << x) by pFam F (trans<< A=B B=C x)
  157. === G ! (B=C << x) by F=G (B=C << x)
  158. === H ! x by G=H x
  159. symS : Sym _=S_
  160. symS {nat }{pi _ _ } (eqS ())
  161. symS {nat }{sigma _ _ } (eqS ())
  162. symS {pi _ _ }{nat } (eqS ())
  163. symS {pi _ _ }{sigma _ _ } (eqS ())
  164. symS {sigma _ _ }{nat } (eqS ())
  165. symS {sigma _ _ }{pi _ _ } (eqS ())
  166. symS {nat}{nat} p = p
  167. symS {pi A F}{pi B G} (eqS < B=A , F=G >) = eqS < A=B , G=F >
  168. where
  169. open module C = Chain _=S_ refS transS
  170. A=B = symS B=A
  171. G=F : G =Fam F >> A=B
  172. G=F x = symS (
  173. chain> F ! (A=B << x)
  174. === G ! (B=A << A=B << x) by F=G (A=B << x)
  175. === G ! (refS << x) by pFam G (casttrans B=A A=B refS x)
  176. === G ! x by pFam G (ref<< x)
  177. )
  178. symS {sigma A F}{sigma B G}(eqS < A=B , F=G >) = eqS < B=A , G=F >
  179. where
  180. open module C = Chain _=S_ refS transS
  181. B=A = symS A=B
  182. G=F : G >> B=A =Fam F
  183. G=F x =
  184. chain> G ! (B=A << x)
  185. === F ! (A=B << B=A << x) by symS (F=G _)
  186. === F ! (refS << x) by pFam F (casttrans _ _ _ x)
  187. === F ! x by pFam F (castref _ x)
  188. pfi : {A B : S}(p q : A =S B)(x : El B) -> p << x == q << x
  189. pfi {nat }{pi _ _ } (eqS ()) _ _
  190. pfi {nat }{sigma _ _ } (eqS ()) _ _
  191. pfi {pi _ _ }{nat } (eqS ()) _ _
  192. pfi {pi _ _ }{sigma _ _ } (eqS ()) _ _
  193. pfi {sigma _ _ }{nat } (eqS ()) _ _
  194. pfi {sigma _ _ }{pi _ _ } (eqS ()) _ _
  195. pfi {nat}{nat} _ _ x = ref
  196. pfi {pi A F}{pi B G} (eqS < B=A1 , F=G1 >) (eqS < B=A2 , F=G2 >)
  197. (el < g , pg >) = eq g1=g2
  198. where
  199. g1=g2 : (x : El A) -> F=G1 x << g (B=A1 << x)
  200. == F=G2 x << g (B=A2 << x)
  201. g1=g2 x =
  202. chain> F=G1 x << g (B=A1 << x)
  203. === F=G1 x << _ << g (B=A2 << x) by p<< _ (pg (pfi B=A1 B=A2 x))
  204. === F=G2 x << g (B=A2 << x) by casttrans _ _ _ _
  205. where
  206. open module C = Chain _==_ (ref {_}) (trans {_})
  207. pfi {sigma A F}{sigma B G} (eqS < A=B1 , F=G1 >) (eqS < A=B2 , F=G2 >)
  208. (el < y , Gy >) = eq < x1=x2 , Fx1=Fx2 >
  209. where
  210. x1=x2 : A=B1 << y == A=B2 << y
  211. x1=x2 = pfi A=B1 A=B2 y
  212. Fx1=Fx2 : F=G1 y << Gy == pFam F x1=x2 << F=G2 y << Gy
  213. Fx1=Fx2 = sym (casttrans _ _ _ _)
  214. ref<< : {A : S}(x : El A) -> refS << x == x
  215. ref<< {nat} x = ref
  216. ref<< {sigma A F} (el < x , Fx >) = eq < ref<< x , pfi _ _ Fx >
  217. ref<< {pi A F } (el < f , pf >) = eq rf=f
  218. where
  219. rf=f : (x : El A) -> _ << f (refS << x) == f x
  220. rf=f x =
  221. chain> _ << f (refS << x)
  222. === _ << pFam F (ref<< x) << f x by p<< _ (pf (ref<< x))
  223. === _ << f x by sym (trans<< _ _ (f x))
  224. === f x by castref _ _
  225. where open module C = Chain _==_ (ref {_}) (trans {_})
  226. trans<< : {A B C : S}(A=B : A =S B)(B=C : B =S C)(x : El C) ->
  227. transS A=B B=C << x == A=B << B=C << x
  228. trans<< {nat }{nat }{pi _ _ } _ (eqS ()) _
  229. trans<< {nat }{nat }{sigma _ _ } _ (eqS ()) _
  230. trans<< {nat }{pi _ _ } (eqS ()) _ _
  231. trans<< {nat }{sigma _ _ } (eqS ()) _ _
  232. trans<< {pi _ _ }{nat } (eqS ()) _ _
  233. trans<< {pi _ _ }{pi _ _ }{nat } _ (eqS ()) _
  234. trans<< {pi _ _ }{pi _ _ }{sigma _ _ } _ (eqS ()) _
  235. trans<< {pi _ _ }{sigma _ _ } (eqS ()) _ _
  236. trans<< {sigma _ _ }{nat } (eqS ()) _ _
  237. trans<< {sigma _ _ }{pi _ _ } (eqS ()) _ _
  238. trans<< {sigma _ _ }{sigma _ _ }{nat } _ (eqS ()) _
  239. trans<< {sigma _ _ }{sigma _ _ }{pi _ _ } _ (eqS ()) _
  240. trans<< {nat}{nat}{nat} _ _ _ = ref
  241. trans<< {pi A F}{pi B G}{pi C H}
  242. (eqS < B=A , F=G >)(eqS < C=B , G=H >)
  243. (el < h , ph >) = eq prf
  244. where
  245. C=A = transS C=B B=A
  246. prf : (x : El A) -> _
  247. prf x =
  248. chain> _ << h (C=A << x)
  249. === _ << _ << h (C=B << B=A << x) by p<< _ (ph (trans<< _ _ x))
  250. === F=G x << G=H _ << h (_ << _ << x) by pfi2 _ _ _ _ _
  251. where open module C' = Chain _==_ (ref {_}) (trans {_})
  252. trans<< {sigma A F}{sigma B G}{sigma C H}
  253. (eqS < A=B , F=G >)(eqS < B=C , G=H >)
  254. (el < z , Hz >) = eq < trans<< A=B B=C z , prf >
  255. where
  256. prf =
  257. chain> _ << Hz
  258. === _ << Hz by pfi _ _ _
  259. === _ << _ << Hz by trans<< _ _ _
  260. === _ << F=G _ << G=H z << Hz by trans<< _ _ _
  261. where open module C' = Chain _==_ (ref {_}) (trans {_})
  262. -- we never need this one, but it feels like it should be here...
  263. sym<< : {A B : S}(A=B : A =S B)(x : El B) ->
  264. symS A=B << A=B << x == x
  265. sym<< A=B x =
  266. chain> symS A=B << A=B << x
  267. === refS << x by casttrans _ _ _ x
  268. === x by ref<< x
  269. where open module C' = Chain _==_ (ref {_}) (trans {_})
  270. castref : {A : S}(p : A =S A)(x : El A) -> p << x == x
  271. castref A=A x =
  272. chain> A=A << x
  273. === refS << x by pfi A=A refS x
  274. === x by ref<< x
  275. where open module C = Chain _==_ (ref {_}) (trans {_})
  276. casttrans : {A B C : S}(A=B : A =S B)(B=C : B =S C)(A=C : A =S C)(x : El C) ->
  277. A=B << B=C << x == A=C << x
  278. casttrans A=B B=C A=C x =
  279. chain> A=B << B=C << x
  280. === _ << x by sym (trans<< _ _ _)
  281. === A=C << x by pfi _ _ _
  282. where open module C' = Chain _==_ (ref {_}) (trans {_})
  283. pfi2 : {A B1 B2 C : S}
  284. (A=B1 : A =S B1)(A=B2 : A =S B2)(B1=C : B1 =S C)(B2=C : B2 =S C)
  285. (x : El C) -> A=B1 << B1=C << x == A=B2 << B2=C << x
  286. pfi2 A=B1 A=B2 B1=C B2=C x =
  287. chain> A=B1 << B1=C << x
  288. === _ << x by casttrans _ _ _ x
  289. === A=B2 << B2=C << x by trans<< _ _ x
  290. where
  291. open module C = Chain _==_ (ref {_}) (trans {_})
  292. ref : {A : S} -> Refl {El A} _==_
  293. ref {nat} {el n} = eq (refN {n})
  294. ref {pi A F} {el < f , pf >} = eq \x -> ref
  295. ref {sigma A F} {el < x , Fx >} = eq < ref , sym (castref _ _) >
  296. trans : {A : S} -> Trans {El A} _==_
  297. trans {nat}{el x}{el y}{el z} (eq p) (eq q) = eq (transN {x}{y}{z} p q)
  298. trans {pi A F}{el < f , pf >}{el < g , pg >}{el < h , ph >}
  299. (eq f=g)(eq g=h) = eq \x -> trans (f=g x) (g=h x)
  300. trans {sigma A F}{el < x , Fx >}{el < y , Fy >}{el < z , Fz >}
  301. (eq < x=y , Fx=Fy >)(eq < y=z , Fy=Fz >) =
  302. eq < x=z , Fx=Fz >
  303. where
  304. x=z = trans x=y y=z
  305. Fx=Fz =
  306. chain> Fx
  307. === pFam F x=y << Fy by Fx=Fy
  308. === pFam F x=y << pFam F y=z << Fz by p<< _ Fy=Fz
  309. === pFam F x=z << Fz by casttrans _ _ _ _
  310. where open module C = Chain _==_ (ref {_}) (trans {_})
  311. sym : {A : S} -> Sym {El A} _==_
  312. sym {nat}{el x}{el y} (eq p) = eq (symN {x}{y} p)
  313. sym {pi A F}{el < f , pf >}{el < g , pg >}
  314. (eq f=g) = eq \x -> sym (f=g x)
  315. sym {sigma A F}{el < x , Fx >}{el < y , Fy >}
  316. (eq < x=y , Fx=Fy >) = eq < y=x , Fy=Fx >
  317. where
  318. y=x = sym x=y
  319. Fy=Fx = sym (
  320. chain> pFam F y=x << Fx
  321. === pFam F y=x << pFam F x=y << Fy by p<< (pFam F y=x) Fx=Fy
  322. === refS << Fy by casttrans _ _ _ _
  323. === Fy by castref _ _
  324. )
  325. where open module C = Chain _==_ (ref {_}) (trans {_})
  326. refFam : {A : S} -> Refl (_=Fam_ {A})
  327. refFam x = refS
  328. transFam : {A : S} -> Trans (_=Fam_ {A})
  329. transFam F=G G=H x = transS (F=G x) (G=H x)
  330. symFam : {A : S} -> Sym (_=Fam_ {A})
  331. symFam F=G x = symS (F=G x)
  332. castref2 : {A B : S}(A=B : A =S B)(B=A : B =S A)(x : El A) ->
  333. A=B << B=A << x == x
  334. castref2 A=B B=A x = trans (casttrans A=B B=A refS x) (ref<< x)