07-subclasses.agda 4.8 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208
  1. -- {-# OPTIONS --verbose tc.records.ifs:15 #-}
  2. -- {-# OPTIONS --verbose tc.constr.findInScope:15 #-}
  3. -- {-# OPTIONS --verbose tc.term.args.ifs:15 #-}
  4. -- {-# OPTIONS --verbose cta.record.ifs:15 #-}
  5. -- {-# OPTIONS --verbose tc.section.apply:25 #-}
  6. -- {-# OPTIONS --verbose tc.mod.apply:100 #-}
  7. -- {-# OPTIONS --verbose scope.rec:15 #-}
  8. -- {-# OPTIONS --verbose tc.rec.def:15 #-}
  9. module InstanceArguments.07-subclasses where
  10. module Imports where
  11. module L where
  12. open import Agda.Primitive public
  13. using (Level; _⊔_) renaming (lzero to zero; lsuc to suc)
  14. -- extract from Function
  15. id : ∀ {a} {A : Set a} → A → A
  16. id x = x
  17. _$_ : ∀ {a b} {A : Set a} {B : A → Set b} →
  18. ((x : A) → B x) → ((x : A) → B x)
  19. f $ x = f x
  20. _∘_ : ∀ {a b c}
  21. {A : Set a} {B : A → Set b} {C : {x : A} → B x → Set c} →
  22. (∀ {x} (y : B x) → C y) → (g : (x : A) → B x) →
  23. ((x : A) → C (g x))
  24. f ∘ g = λ x → f (g x)
  25. -- extract from Data.Bool
  26. infixr 5 _∨_
  27. data Bool : Set where
  28. true : Bool
  29. false : Bool
  30. not : Bool → Bool
  31. not true = false
  32. not false = true
  33. _∨_ : Bool → Bool → Bool
  34. true ∨ b = true
  35. false ∨ b = b
  36. -- extract from Relation.Nullary.Decidable and friends
  37. infix 3 ¬_
  38. data ⊥ : Set where
  39. ¬_ : ∀ {ℓ} → Set ℓ → Set ℓ
  40. ¬ P = P → ⊥
  41. data Dec {p} (P : Set p) : Set p where
  42. yes : ( p : P) → Dec P
  43. no : (¬p : ¬ P) → Dec P
  44. ⌊_⌋ : ∀ {p} {P : Set p} → Dec P → Bool
  45. ⌊ yes _ ⌋ = true
  46. ⌊ no _ ⌋ = false
  47. -- extract from Relation.Binary.PropositionalEquality
  48. data _≡_ {a} {A : Set a} (x : A) : A → Set a where
  49. refl : x ≡ x
  50. cong : ∀ {a b} {A : Set a} {B : Set b}
  51. (f : A → B) {x y} → x ≡ y → f x ≡ f y
  52. cong f refl = refl
  53. -- extract from Data.Nat
  54. data ℕ : Set where
  55. zero : ℕ
  56. suc : (n : ℕ) → ℕ
  57. {-# BUILTIN NATURAL ℕ #-}
  58. pred : ℕ → ℕ
  59. pred zero = zero
  60. pred (suc n) = n
  61. _≟_ : (x y : ℕ) → Dec (x ≡ y)
  62. zero ≟ zero = yes refl
  63. suc m ≟ suc n with m ≟ n
  64. suc m ≟ suc .m | yes refl = yes refl
  65. suc m ≟ suc n | no prf = no (prf ∘ cong pred)
  66. zero ≟ suc n = no λ()
  67. suc m ≟ zero = no λ()
  68. open Imports
  69. -- Begin of actual example!
  70. record Eq (A : Set) : Set where
  71. field eq : A → A → Bool
  72. primEqBool : Bool → Bool → Bool
  73. primEqBool true = id
  74. primEqBool false = not
  75. instance
  76. eqBool : Eq Bool
  77. eqBool = record { eq = primEqBool }
  78. primEqNat : ℕ → ℕ → Bool
  79. primEqNat a b = ⌊ a ≟ b ⌋
  80. primLtNat : ℕ → ℕ → Bool
  81. primLtNat 0 _ = true
  82. primLtNat (suc a) (suc b) = primLtNat a b
  83. primLtNat _ _ = false
  84. neq : {t : Set} → {{eqT : Eq t}} → t → t → Bool
  85. neq a b = not $ eq a b
  86. where open Eq {{...}}
  87. record Ord₁ (A : Set) : Set where
  88. infix 6 _<_
  89. field _<_ : A → A → Bool
  90. eqA : Eq A
  91. instance
  92. ord₁Nat : Ord₁ ℕ
  93. ord₁Nat = record { _<_ = primLtNat; eqA = eqNat }
  94. where eqNat : Eq ℕ
  95. eqNat = record { eq = primEqNat }
  96. record Ord₂ {A : Set} (eqA : Eq A) : Set where
  97. infix 6 _<_
  98. field _<_ : A → A → Bool
  99. instance
  100. ord₂Nat : Ord₂ (record { eq = primEqNat })
  101. ord₂Nat = record { _<_ = primLtNat }
  102. record Ord₃ (A : Set) : Set where
  103. infix 6 _<_
  104. field _<_ : A → A → Bool
  105. eqA : Eq A
  106. open Eq eqA public
  107. instance
  108. ord₃Nat : Ord₃ ℕ
  109. ord₃Nat = record { _<_ = primLtNat; eqA = eqNat }
  110. where eqNat : Eq ℕ
  111. eqNat = record { eq = primEqNat }
  112. record Ord₄ {A : Set} {{eqA : Eq A}} : Set where
  113. infix 6 _<_
  114. field _<_ : A → A → Bool
  115. open Eq eqA public
  116. instance
  117. ord₄Nat : Ord₄ {{record { eq = primEqNat }}}
  118. ord₄Nat = record { _<_ = primLtNat }
  119. module test₁ where
  120. open Ord₁ {{...}}
  121. open Eq {{...}}
  122. instance
  123. eqNat : Eq ℕ
  124. eqNat = eqA
  125. test₁ = 5 < 3
  126. test₂ = eq 5 3
  127. test₃ = eq true false
  128. test₄ : {A : Set} → {{ ordA : Ord₁ A }} → A → A → Bool
  129. test₄ a b = a < b ∨ eq a b
  130. where
  131. instance
  132. eqA' : Eq _
  133. eqA' = eqA
  134. module test₃ where
  135. open Ord₃ {{...}}
  136. open Eq {{...}} renaming (eq to eq')
  137. test₁ = 5 < 3
  138. test₂ = eq 5 3
  139. test₃ = eq' true false
  140. test₄ : {A : Set} → {{ ordA : Ord₃ A }} → A → A → Bool
  141. test₄ a b = a < b ∨ eq a b
  142. module test₄ where
  143. open Ord₄ {{...}}
  144. open Eq {{...}} renaming (eq to eq')
  145. test₁ = 5 < 3
  146. test₂ = eq 5 3
  147. test₃ = eq' true false
  148. test₄ : {A : Set} → {{eqA : Eq A}} → {{ ordA : Ord₄ }} → A → A → Bool
  149. test₄ a b = a < b ∨ eq a b
  150. module test₄′ where
  151. open Ord₄ {{...}} hiding (eq)
  152. open Eq {{...}}
  153. eqNat : Eq ℕ
  154. eqNat = record { eq = primEqNat }
  155. test₁ = 5 < 3
  156. test₂ = eq 5 3
  157. test₃ = eq true false
  158. test₄ : {A : Set} → {{ eqA : Eq A }} → {{ ordA : Ord₄ }} → A → A → Bool
  159. test₄ a b = a < b ∨ eq a b