SET.agda 20 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602
  1. module SET where
  2. ----------------------------------------------------------------------------
  3. -- Auxiliary.
  4. ----------------------------------------------------------------------------
  5. data Fun (X Y : Set) : Set where
  6. fun : (X -> Y) -> Fun X Y
  7. {-
  8. Unop : Set -> Set
  9. Unop X = Fun X X
  10. Binop : Set -> Set
  11. Binop X = Fun X (Fun X X)
  12. -}
  13. -- We need to replace Pred X by its RHS (less readable!)
  14. -- Pred : Set -> Set1
  15. -- Pred X = X -> Set
  16. -- Pow = Pred
  17. -- Rel : Set -> Set1
  18. -- Rel X = X -> X -> Set
  19. data Reflexive {X : Set} (R : X -> X -> Set) : Set where
  20. reflexive : ((x : X) -> R x x) -> Reflexive R
  21. data Symmetrical {X : Set} (R : X -> X -> Set) : Set where
  22. symmetrical : ( {x1 x2 : X} -> R x1 x2 -> R x2 x1) -> Symmetrical R
  23. {-
  24. Transitive {X : Set}(R : X -> X -> Set) : Set
  25. = (x1 x2 x3 : X) |-> R x1 x2 -> R x2 x3 -> R x1 x3
  26. Compositional {X : Set}(R : X -> X -> Set) : Set
  27. = (x1 : X) |-> (x2 : X) |-> (x3 : X) |-> R x2 x3 -> R x1 x2 -> R x1 x3
  28. -}
  29. data Substitutive {X : Set} (R : X -> X -> Set) : Set1 where
  30. substitutive : ( (P : X -> Set) -> {x1 x2 : X} -> R x1 x2 -> P x1 -> P x2)
  31. -> Substitutive R
  32. {-
  33. Collapsed (X : Set) : Set1
  34. = (P : X -> Set) -> (x1 x2 : X) |-> P x1 -> P x2
  35. id {X : Set} : X -> X
  36. = \x -> x
  37. cmp (|X |Y |Z : Set) : (Y -> Z) -> (X -> Y) -> X -> Z
  38. = \f -> \g -> \x -> f (g x)
  39. seq (|X |Y |Z : Set)(f : X -> Y)(g : Y -> Z) : X -> Z
  40. = cmp g f
  41. const (|X |Y : Set)(x : X)(y : Y) : X
  42. = x
  43. proj {X : Set}(Y : X -> Set)(x : X)(f : (x : X) -> Y x) : Y x
  44. = f x
  45. flip {X : Set}{Y : Set}{Z : Set}(f : X -> Y -> Z)(y : Y)(x : X) : Z
  46. = f x y
  47. FlipRel {X : Set}(R : X -> X -> Set)(x1 : X)(x2 : X) : Set
  48. = R x2 x1
  49. ----------------------------------------------------------------------------
  50. -- Product sets.
  51. ----------------------------------------------------------------------------
  52. Prod (X : Set)(Y : X -> Set) : Set
  53. = (x : X) -> Y x
  54. mapProd {X : Set}
  55. {Y1 : X -> Set}
  56. {Y2 : X -> Set}
  57. (f : (x : X) -> Y1 x -> Y2 x)
  58. : Prod X Y1 -> Prod X Y2
  59. = \g -> \x -> f x (g x)
  60. -- Fun(X : Set)(Y : Set) = X -> Y
  61. mapFun (|X1 |X2 |Y1 |Y2 : Set)
  62. : (X2 -> X1) -> (Y1 -> Y2) -> (X1 -> Y1) -> X2 -> Y2
  63. = \f -> \g -> \h -> \x ->
  64. g (h (f x))
  65. ----------------------------------------------------------------------------
  66. -- Identity proof sets.
  67. ----------------------------------------------------------------------------
  68. Id {X : Set} : X -> X -> Set
  69. = idata ref (x : X) : _ x x
  70. refId {X : Set} : Reflexive Id
  71. = \(x : X) -> ref@_ x
  72. elimId (|X : Set)
  73. (C : (x1 x2 : X) |-> Id x1 x2 -> Set)
  74. (refC : (x : X) -> C (refId x))
  75. (|x1 |x2 : X)
  76. (u : Id x1 x2) :
  77. C u
  78. = case u of { (ref x) -> refC x;}
  79. abstract whenId {X : Set}(C : X -> X -> Set)(c : (x : X) -> C x x)
  80. : (x1 x2 : X) |-> Id x1 x2 -> C x1 x2
  81. = elimId (\x1 x2 |-> \(u : Id x1 x2) -> C x1 x2) c
  82. abstract substId {X : Set} : Substitutive Id
  83. = \(C : X -> Set) ->
  84. whenId (\x1 x2 -> C x1 -> C x2) (\x -> id)
  85. abstract mapId {X : Set}{Y : Set}(f : X -> Y)
  86. : (x1 x2 : X) |-> Id x1 x2 -> Id (f x1) (f x2)
  87. = whenId (\x1 x2 -> Id (f x1) (f x2)) (\(x : X) -> refId (f x))
  88. abstract symId {X : Set} : Symmetrical Id
  89. = whenId (\(x1 x2 : X) -> Id x2 x1) refId
  90. abstract cmpId {X : Set} : Compositional Id
  91. = let lem : (x y : X) |-> Id x y -> (z : X) |-> Id z x -> Id z y
  92. = whenId ( \(x y : _) -> (z : X) |-> Id z x -> Id z y)
  93. ( \x -> \z |-> id)
  94. in \(x1 x2 x3 : _) |->
  95. \(u : Id x2 x3) ->
  96. \(v : Id x1 x2) ->
  97. lem u v
  98. abstract tranId {X : Set} : Transitive Id
  99. = \(x1 x2 x3 : X) |->
  100. \(u : Id x1 x2) ->
  101. \(v : Id x2 x3) ->
  102. cmpId v u
  103. ----------------------------------------------------------------------------
  104. -- The empty set.
  105. ----------------------------------------------------------------------------
  106. -}
  107. data Zero : Set where
  108. {-
  109. abstract whenZero (X : Set)(z : Zero) : X
  110. = case z of { }
  111. elimZero (C : Zero -> Set)(z : Zero) : C z
  112. = case z of { }
  113. abstract collZero : Collapsed Zero
  114. = \(C : Zero -> Set) ->
  115. \(z1 z2 : Zero) |->
  116. \(c : C z1) ->
  117. case z1 of { }
  118. ----------------------------------------------------------------------------
  119. -- The singleton set.
  120. ----------------------------------------------------------------------------
  121. -}
  122. data Unit : Set where
  123. unit : Unit
  124. {-
  125. elUnit = tt
  126. elimUnit (C : Unit -> Set)(c_tt : C tt@_)(u : Unit) : C u
  127. = case u of { (tt) -> c_tt;}
  128. abstract collUnit : Collapsed Unit
  129. = \(C : Unit -> Set) ->
  130. \(u1 u2 : Unit) |->
  131. \(c : C u1) ->
  132. case u1 of { (tt) -> case u2 of { (tt) -> c;};}
  133. ----------------------------------------------------------------------------
  134. -- The successor set adds a new element.
  135. ----------------------------------------------------------------------------
  136. Succ (X : Set) : Set
  137. = data zer | suc (x : X)
  138. zerSucc {X : Set} : Succ X
  139. = zer@_
  140. sucSucc {X : Set}(x : X) : Succ X
  141. = suc@_ x
  142. elimSucc {X : Set}
  143. (C : Succ X -> Set)
  144. (c_z : C zer@_)
  145. (c_s : (x : X) -> C (suc@_ x))
  146. (x' : Succ X)
  147. : C x'
  148. = case x' of {
  149. (zer) -> c_z;
  150. (suc x) -> c_s x;}
  151. whenSucc (|X |Y : Set)(y_z : Y)(y_s : X -> Y)(x' : Succ X) : Y
  152. = case x' of {
  153. (zer) -> y_z;
  154. (suc x) -> y_s x;}
  155. mapSucc (|X |Y : Set)(f : X -> Y) : Succ X -> Succ Y
  156. = whenSucc zer@(Succ Y) (\(x : X) -> suc@_ (f x)) -- (Succ Y)
  157. ----------------------------------------------------------------------------
  158. -- The (binary) disjoint union.
  159. ----------------------------------------------------------------------------
  160. data Plus (X Y : Set) = inl (x : X) | inr (y : Y)
  161. elimPlus (|X |Y : Set)
  162. (C : Plus X Y -> Set)
  163. (c_lft : (x : X) -> C (inl@_ x))
  164. (c_rgt : (y : Y) -> C (inr@_ y))
  165. (xy : Plus X Y)
  166. : C xy
  167. = case xy of {
  168. (inl x) -> c_lft x;
  169. (inr y) -> c_rgt y;}
  170. when (|X |Y |Z : Set)(f : X -> Z)(g : Y -> Z) : Plus X Y -> Z
  171. = \xy -> case xy of {
  172. (inl x) -> f x;
  173. (inr y) -> g y;}
  174. whenPlus = when
  175. mapPlus (|X1 |X2 |Y1 |Y2 : Set)(f : X1 -> X2)(g : Y1 -> Y2)
  176. : Plus X1 Y1 -> Plus X2 Y2
  177. = when (\x1 -> inl (f x1)) (\y1 -> inr (g y1))
  178. swapPlus (|X |Y : Set) : Plus X Y -> Plus Y X
  179. = when inr inl
  180. ----------------------------------------------------------------------------
  181. -- Dependent pairs.
  182. ----------------------------------------------------------------------------
  183. Sum (X : Set)(Y : X -> Set) : Set
  184. = sig{fst : X;
  185. snd : Y fst;}
  186. dep_pair {X : Set}{Y : X -> Set}(x : X)(y : Y x) : Sum X Y
  187. = struct {fst = x; snd = y;}
  188. dep_fst {X : Set}{Y : X -> Set}(xy : Sum X Y) : X
  189. = xy.fst
  190. dep_snd {X : Set}{Y : X -> Set}(xy : Sum X Y) : Y (dep_fst xy)
  191. = xy.snd
  192. dep_cur {X : Set}{Y : X -> Set}{Z : Set}(f : Sum X Y -> Z)
  193. : (x : X) |-> Y x -> Z
  194. = \x |-> \y -> f (dep_pair x y)
  195. dep_uncur {X : Set}{Y : X -> Set}{Z : Set}
  196. : ((x : X) -> Y x -> Z) -> Sum X Y -> Z
  197. = \(f : (x : X) -> (x' : Y x) -> Z) -> \(xy : Sum X Y) -> f xy.fst xy.snd
  198. dep_curry {X : Set}
  199. {Y : X -> Set}
  200. (Z : Sum X Y -> Set)
  201. (f : (xy : Sum X Y) -> Z xy)
  202. : (x : X) -> (y : Y x) -> Z (dep_pair x y)
  203. = \(x : X) -> \(y : Y x) -> f (dep_pair x y)
  204. dep_uncurry {X : Set}
  205. {Y : X -> Set}
  206. (Z : Sum X Y -> Set)
  207. (f : (x : X) ->
  208. (y : Y x) ->
  209. Z (dep_pair x y))
  210. (xy : Sum X Y)
  211. : Z xy
  212. = f xy.fst xy.snd
  213. mapSum {X : Set}{Y1 : X -> Set}{Y2 : X -> Set}(f : (x : X) -> Y1 x -> Y2 x)
  214. : Sum X Y1 -> Sum X Y2
  215. = \(p : Sum X Y1) -> dep_pair p.fst (f p.fst p.snd)
  216. elimSum = dep_uncurry
  217. ----------------------------------------------------------------------------
  218. -- Nondependent pairs (binary) cartesian product.
  219. ----------------------------------------------------------------------------
  220. Times (X : Set)(Y : Set) : Set
  221. = Sum X (\(x : X) -> Y)
  222. pair {X : Set}{Y : Set} : X -> Y -> Times X Y
  223. = \(x : X) ->
  224. \(y : Y) ->
  225. struct {
  226. fst = x;
  227. snd = y;}
  228. fst {X : Set}{Y : Set} : Times X Y -> X
  229. = \(xy : Times X Y) -> xy.fst
  230. snd {X : Set}{Y : Set} : Times X Y -> Y
  231. = \(xy : Times X Y) -> xy.snd
  232. pairfun {X : Set}{Y : Set}{Z : Set}(f : X -> Y)(g : X -> Z)(x : X)
  233. : Times Y Z
  234. = pair (f x) (g x)
  235. mapTimes {X1 : Set}{X2 : Set}{Y1 : Set}{Y2 : Set}
  236. : (f : X1 -> X2) -> (g : Y1 -> Y2) -> Times X1 Y1 -> Times X2 Y2
  237. = \(f : (x : X1) -> X2) ->
  238. \(g : (x : Y1) -> Y2) ->
  239. \(xy : Times X1 Y1) ->
  240. pair (f xy.fst) (g xy.snd)
  241. swapTimes {X : Set}{Y : Set} : Times X Y -> Times Y X
  242. = pairfun snd fst
  243. cur {X : Set}{Y : Set}{Z : Set}(f : Times X Y -> Z) : X -> Y -> Z
  244. = \(x : X) -> \(y : Y) -> f (pair |_ |_ x y)
  245. uncur {X : Set}{Y : Set}{Z : Set}(f : X -> Y -> Z) : Times X Y -> Z
  246. = \(xy : Times X Y) -> f xy.fst xy.snd
  247. curry {X : Set}
  248. {Y : Set}
  249. {Z : Times X Y -> Set}
  250. (f : (xy : Times X Y) -> Z xy)
  251. : (x : X) ->
  252. (y : Y) ->
  253. Z (pair |_ |_ x y)
  254. = \(x : X) ->
  255. \(y : Y) ->
  256. f (pair |_ |_ x y)
  257. uncurry {X : Set}
  258. {Y : Set}
  259. {Z : Times X Y -> Set}
  260. (f : (x : X) ->
  261. (y : Y) ->
  262. Z (pair |_ |_ x y))
  263. : (xy : Times X Y) -> Z xy
  264. = \(xy : Times X Y) -> f xy.fst xy.snd
  265. elimTimes = uncurry
  266. ----------------------------------------------------------------------------
  267. -- Natural numbers.
  268. ----------------------------------------------------------------------------
  269. Nat : Set
  270. = data zer | suc (m : Nat)
  271. zero : Nat
  272. = zer@_
  273. succ (x : Nat) : Nat
  274. = suc@_ x
  275. elimNat (C : Nat -> Set)
  276. : (c_z : C zer@_) ->
  277. (c_s : (x : Nat) -> C x -> C (suc@_ x)) ->
  278. (m : Nat) ->
  279. C m
  280. = \(c_z : C zer@_) ->
  281. \(c_s : (x : Nat) -> (x' : C x) -> C (suc@_ x)) ->
  282. \(m : Nat) ->
  283. case m of {
  284. (zer) -> c_z;
  285. (suc m') -> c_s m' (elimNat C c_z c_s m');}
  286. ----------------------------------------------------------------------------
  287. -- Linear universe of finite sets.
  288. ----------------------------------------------------------------------------
  289. Fin (m : Nat) : Set
  290. = case m of {
  291. (zer) -> Zero;
  292. (suc m') -> Succ (Fin m');}
  293. valFin (m : Nat) : Fin m -> Nat
  294. = \(n : Fin m) ->
  295. case m of {
  296. (zer) -> case n of { };
  297. (suc m') ->
  298. case n of {
  299. (zer) -> zer@_;
  300. (suc n') -> suc@_ (valFin m' n');};}
  301. zeroFin (m : Nat) : Fin (succ m)
  302. = zer@_
  303. succFin (m : Nat)(n : Fin m) : Fin (succ m)
  304. = suc@_ n
  305. ----------------------------------------------------------------------------
  306. -- Do these really belong here?
  307. ----------------------------------------------------------------------------
  308. HEAD (X : Set1)(m : Nat)(f : Fin (succ m) -> X) : X
  309. = f (zeroFin m)
  310. TAIL (X : Set1)(m : Nat)(f : Fin (succ m) -> X) : Fin m -> X
  311. = \(n : Fin m) -> f (succFin m n)
  312. ----------------------------------------------------------------------------
  313. -- Lists.
  314. ----------------------------------------------------------------------------
  315. List (X : Set) : Set
  316. = data nil | con (x : X) (xs : List X)
  317. nil {X : Set} : List X
  318. = nil@_
  319. con {X : Set}(x : X)(xs : List X) : List X
  320. = con@_ x xs
  321. elimList {X : Set}
  322. (C : List X -> Set)
  323. (c_nil : C (nil |_))
  324. (c_con : (x : X) -> (xs : List X) -> C xs -> C (con@_ x xs))
  325. (xs : List X)
  326. : C xs
  327. = case xs of {
  328. (nil) -> c_nil;
  329. (con x xs') -> c_con x xs' (elimList |_ C c_nil c_con xs');}
  330. ----------------------------------------------------------------------------
  331. -- Tuples are "dependently typed vectors".
  332. ----------------------------------------------------------------------------
  333. Nil : Set
  334. = data nil
  335. Con (X0 : Set)(X' : Set) : Set
  336. = data con (x : X0) (xs : X')
  337. Tuple (m : Nat)(X : Fin m -> Set) : Set
  338. = case m of {
  339. (zer) -> Nil;
  340. (suc m') -> Con (X zer@_) (Tuple m' (\(n : Fin m') -> X (suc@_ n)));}
  341. ----------------------------------------------------------------------------
  342. -- Vectors homogeneously typed tuples.
  343. ----------------------------------------------------------------------------
  344. Vec (X : Set)(m : Nat) : Set
  345. = Tuple m (\(n : Fin m) -> X)
  346. ----------------------------------------------------------------------------
  347. -- Monoidal expressions.
  348. ----------------------------------------------------------------------------
  349. Mon (X : Set) : Set
  350. = data unit | at (x : X) | mul (xs1 : Mon X) (xs2 : Mon X)
  351. ----------------------------------------------------------------------------
  352. -- Propositions.
  353. ----------------------------------------------------------------------------
  354. Imply (X : Set)(Y : Set) : Set
  355. = X -> Y
  356. -}
  357. Absurd : Set
  358. Absurd = Zero
  359. Taut : Set
  360. Taut = Unit
  361. {-
  362. Not (X : Set) : Set
  363. = X -> Absurd
  364. Exist {X : Set}(P : X -> Set) : Set
  365. = Sum X P
  366. Forall (X : Set)(P : X -> Set) : Set
  367. = (x : X) -> P x
  368. And (X : Set)(Y : Set) : Set
  369. = Times X Y
  370. Iff (X : Set)(Y : Set) : Set
  371. = And (Imply X Y) (Imply Y X)
  372. Or (X : Set)(Y : Set) : Set
  373. = Plus X Y
  374. Decidable (X : Set) : Set
  375. = Or X (Imply X Absurd)
  376. DecidablePred {X : Set}(P : X -> Set) : Set
  377. = (x : X) -> Decidable (P x)
  378. DecidableRel {X : Set}(R : X -> X -> Set) : Set
  379. = (x1 : X) -> (x2 : X) -> Decidable (R x1 x2)
  380. Least {X : Set}((<=) : X -> X -> Set)(P : X -> Set) : X -> Set
  381. = \(x : X) -> And (P x) ((x' : X) -> P x' -> (x <= x'))
  382. Greatest {X : Set}((<=) : X -> X -> Set)(P : X -> Set) : X -> Set
  383. = \(x : X) -> And (P x) ((x' : X) -> P x' -> (x' <= x))
  384. ----------------------------------------------------------------------------
  385. -- Booleans.
  386. ----------------------------------------------------------------------------
  387. -}
  388. data Bool : Set where
  389. true : Bool
  390. false : Bool
  391. {-
  392. elimBool (C : Bool -> Set)(c_t : C true@_)(c_f : C false@_)(p : Bool)
  393. : C p
  394. = case p of {
  395. (true) -> c_t;
  396. (false) -> c_f;}
  397. whenBool (C : Set)(c_t : C)(c_f : C) : Bool -> C
  398. = elimBool (\(x : Bool) -> C) c_t c_f
  399. pred (X : Set) : Set
  400. = X -> Bool
  401. -}
  402. -- rel (X : Set) : Set
  403. -- = X -> X -> Bool
  404. True : Bool -> Set
  405. True (true) = Taut
  406. True (false) = Absurd
  407. {-
  408. bool2set = True
  409. pred2Pred {X : Set} : pred X -> X -> Set
  410. = \(p : pred X) -> \(x : X) -> True (p x)
  411. rel2Rel {X : Set} : (X -> X -> Bool) -> X -> X -> Set
  412. = \(r : (X -> X -> Bool)) -> \(x : X) -> \(y : X) -> True (r x y)
  413. decTrue (p : Bool) : Decidable (True p)
  414. = case p of {
  415. (true) -> inl@_ tt;
  416. (false) -> inr@_ (id |_);}
  417. abstract dec_lem {P : Set}(decP : Decidable P)
  418. : Exist |_ (\(b : Bool) -> Iff (True b) P)
  419. = case decP of {
  420. (inl trueP) ->
  421. struct {
  422. fst = true@_;
  423. snd =
  424. struct {
  425. fst = const |_ |_ trueP; -- (True true@_)
  426. snd = const |_ |_ tt;};};
  427. (inr notP) ->
  428. struct {
  429. fst = false@_;
  430. snd =
  431. struct {
  432. fst = whenZero P;
  433. snd = notP;};};}
  434. dec2bool : (P : Set) |-> (decP : Decidable P) -> Bool
  435. = \(P : Set) |-> \(decP : Decidable P) -> (dec_lem |_ decP).fst
  436. dec2bool_spec {P : Set}(decP : Decidable P)
  437. : Iff (True (dec2bool |_ decP)) P
  438. = (dec_lem |_ decP).snd
  439. abstract collTrue : (b : Bool) -> Collapsed (True b)
  440. = let aux (X : Set)(C : X -> Set)
  441. : (b : Bool) ->
  442. (f : True b -> X) ->
  443. (t1 : True b) |->
  444. (t2 : True b) |->
  445. C (f t1) -> C (f t2)
  446. = \(b : Bool) ->
  447. case b of {
  448. (true) ->
  449. \(f : (x : True true@_) -> X) ->
  450. \(t1 t2 : True true@_) |->
  451. \(c : C (f t1)) ->
  452. case t1 of { (tt) -> case t2 of { (tt) -> c;};};
  453. (false) ->
  454. \(f : (x : True false@_) -> X) ->
  455. \(t1 t2 : True false@_) |->
  456. \(c : C (f t1)) ->
  457. case t1 of { };}
  458. in \(b : Bool) -> \(P : True b -> Set) -> aux (True b) P b id
  459. bool2nat (p : Bool) : Nat
  460. = case p of {
  461. (true) -> succ zero;
  462. (false) -> zero;}
  463. ----------------------------------------------------------------------------
  464. -- Decidable subsets.
  465. ----------------------------------------------------------------------------
  466. Filter {X : Set}(p : pred X) : Set
  467. = Sum X (pred2Pred |_ p)
  468. ----------------------------------------------------------------------------
  469. -- Equality.
  470. ----------------------------------------------------------------------------
  471. -- "Deq" stands for "datoid equality" and represents exactly the data
  472. -- that has to be added to turn a set into a datoid.
  473. Deq (X : Set) : Set1
  474. = sig{eq : X -> X -> Bool;
  475. ref : (x : X) -> True (eq x x);
  476. subst :
  477. (C : X -> Set) -> (x1 x2 : X)|-> True (eq x1 x2) -> C x1 -> C x2;}
  478. -- The "Equality" type represents the data that has to be added to turna
  479. -- set into a setoid.
  480. Equality (X : Set) : Set1
  481. = sig{Equal : X -> X -> Set;
  482. ref : Reflexive |_ Equal;
  483. sym : Symmetrical |_ Equal;
  484. tran : Transitive |_ Equal;}
  485. -}
  486. data Datoid : Set1 where
  487. datoid : (Elem : Set) ->
  488. (eq : Elem -> Elem -> Bool) ->
  489. (ref : (x : Elem) -> True (eq x x)) ->
  490. (subst : Substitutive (\x1 -> \x2 -> True (eq x1 x2))) ->
  491. Datoid
  492. pElem : Datoid -> Set
  493. pElem (datoid Elem _ _ _) = Elem
  494. {-
  495. ElD (X : Datoid) : Set
  496. = X.Elem
  497. eqD {X : Datoid} : ElD X -> ElD X -> Bool
  498. = X.eq
  499. EqD {X : Datoid}(x1 x2 : ElD X) : Set
  500. = True (X.eq x1 x2)
  501. Setoid : Set1
  502. = sig{Elem : Set;
  503. Equal : Elem -> Elem -> Set;
  504. ref : (x : Elem) -> Equal x x;
  505. sym : (x1 : Elem) |-> (x2 : Elem) |-> Equal x1 x2 -> Equal x2 x1;
  506. tran :
  507. (x1 : Elem) |->
  508. (x2 : Elem) |->
  509. (x3 : Elem) |->
  510. Equal x1 x2 -> Equal x2 x3 -> Equal x1 x3;}
  511. El (X : Setoid) : Set
  512. = X.Elem
  513. Eq {X : Setoid} : Rel (El X)
  514. = X.Equal
  515. NotEq {X : Setoid} : Rel (El X)
  516. = \x1-> \x2-> Not (Eq |X x1 x2)
  517. Respectable {X : Setoid}(P : El X -> Set) : Set
  518. = (x1 x2 : El X) |-> Eq |X x1 x2 -> P x1 -> P x2
  519. RspEq {X Y : Setoid}(f : El X -> El Y) : Set
  520. = (x1 x2 : El X) |-> Eq |X x1 x2 -> Eq |Y (f x1) (f x2)
  521. RspEq2 (|X |Y |Z : Setoid)(f : El X -> El Y -> El Z)
  522. : Set
  523. = (x1 x2 : X.Elem) |-> (y1 y2 : Y.Elem) ->
  524. Eq |X x1 x2 ->
  525. Eq |Y y1 y2 ->
  526. Eq |Z (f x1 y1) (f x2 y2)
  527. D2S (Y : Datoid) : Setoid
  528. = struct {
  529. Elem = Y.Elem;
  530. Equal = \(x1 x2 : Elem) -> True (Y.eq x1 x2);
  531. ref = Y.ref;
  532. sym =
  533. \(x1 x2 : Elem) |->
  534. \(u : Equal x1 x2) ->
  535. Y.subst (\(x : Y.Elem) -> Equal x x1) |_ |_ u (ref x1);
  536. tran =
  537. \(x1 x2 x3 : Elem) |->
  538. \(u : Equal x1 x2) ->
  539. \(v : Equal x2 x3) ->
  540. Y.subst (Equal x1) |_ |_ v u;}
  541. {-# Alfa unfoldgoals off
  542. brief on
  543. hidetypeannots off
  544. wide
  545. nd
  546. hiding on
  547. con "nil" as "[]" with symbolfont
  548. con "con" infix as " : " with symbolfont
  549. var "Forall" as "\"" with symbolfont
  550. var "Exist" as "$" with symbolfont
  551. var "And" infix as "&" with symbolfont
  552. var "Or" infix as "Ú" with symbolfont
  553. var "Iff" infix as "«" with symbolfont
  554. var "Not" as "Ø" with symbolfont
  555. var "Imply" infix as "É" with symbolfont
  556. var "Taut" as "T" with symbolfont
  557. var "Absurd" as "^" with symbolfont
  558. var "El" mixfix as "|_|" with symbolfont
  559. var "Eq" distfix3 as "==" with symbolfont
  560. var "NotEq" distfix3 as "=|=" with symbolfont
  561. var "True" mixfix as "|_|" with symbolfont
  562. var "ElD" mixfix as "|_|" with symbolfont
  563. var "EqD" distfix3 as "==" with symbolfont
  564. var "Id" distfix3 as "=" with symbolfont
  565. #-}
  566. -}