Congruence.agda 1.3 KB

1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253
  1. module Logic.Congruence where
  2. import Prelude
  3. import Logic.Relations
  4. import Logic.Equivalence
  5. open Prelude
  6. open Logic.Relations
  7. open Logic.Equivalence using (Equivalence)
  8. renaming (module Equivalence to Proj)
  9. data Congruence (A : Set) : Set1 where
  10. congruence :
  11. (Eq : Equivalence A) ->
  12. Congruent (Proj._==_ Eq) ->
  13. Congruence A
  14. module Projections where
  15. eq : {A : Set} -> Congruence A -> Rel A
  16. eq (congruence Eq _) = Proj._==_ Eq
  17. refl : {A : Set}(Cong : Congruence A) -> Reflexive (eq Cong)
  18. refl (congruence Eq _) = Proj.refl Eq
  19. sym : {A : Set}(Cong : Congruence A) -> Symmetric (eq Cong)
  20. sym (congruence Eq _) = Proj.sym Eq
  21. trans : {A : Set}(Cong : Congruence A) -> Transitive (eq Cong)
  22. trans (congruence Eq _) = Proj.trans Eq
  23. cong : {A : Set}(Cong : Congruence A) -> Congruent (eq Cong)
  24. cong (congruence _ c) = c
  25. module CongruenceM {A : Set}(Cong : Congruence A) where
  26. _==_ = Projections.eq Cong
  27. refl = Projections.refl Cong
  28. sym = Projections.sym Cong
  29. trans = Projections.trans Cong
  30. cong = Projections.cong Cong
  31. cong2 : (f : A -> A -> A)(a b c d : A) -> a == c -> b == d -> f a b == f c d
  32. cong2 f a b c d ac bd = trans _ _ _ rem1 rem2
  33. where
  34. rem1 : f a b == f a d
  35. rem1 = cong (f a) _ _ bd
  36. rem2 : f a d == f c d
  37. rem2 = cong (flip f d) _ _ ac