LostConstraint.agda 923 B

1234567891011121314151617181920212223242526272829303132333435363738
  1. -- The bug appeared since I forgot to keep the constraint when trying to unify
  2. -- a blocked term (I just didn't instantiate).
  3. module LostConstraint where
  4. infixr 10 _=>_
  5. data Setoid : Set1 where
  6. setoid : Set -> Setoid
  7. El : Setoid -> Set
  8. El (setoid A) = A
  9. eq : (A : Setoid) -> El A -> El A -> Set
  10. eq (setoid A) = e
  11. where
  12. postulate e : (x, y : A) -> Set
  13. data _=>_ (A B : Setoid) : Set where
  14. lam : (f : El A -> El B)
  15. -> ((x : El A) -> eq A x x
  16. -> eq B (f x) (f x)
  17. )
  18. -> A => B
  19. app : {A B : Setoid} -> (A => B) -> El A -> El B
  20. app (lam f _) = f
  21. postulate EqFun : {A B : Setoid}(f, g : A => B) -> Set
  22. lam2 : {A B C : Setoid} ->
  23. (f : El A -> El B -> El C) ->
  24. (x : El A) -> B => C
  25. lam2 {A}{B}{C} f x = lam (f x) (lem _)
  26. where
  27. postulate
  28. lem : (x : El A)(y : El B) -> eq B y y -> eq C (f x y) (f x y)