@@ -15,9 +15,10 @@ open import Relation.Nullary.Decidable.Core using (map′; _×-dec_;
15
15
open import Relation.Binary.Definitions using (DecidableEquality)
16
16
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂)
17
17
18
- import Reflection.AST.Argument as Arg
19
- import Reflection.AST.Name as Name
20
- import Reflection.AST.Term as Term
18
+ import Reflection.AST.Argument as Arg
19
+ import Reflection.AST.Argument.Quantity as Quantity
20
+ import Reflection.AST.Name as Name
21
+ import Reflection.AST.Term as Term
21
22
22
23
------------------------------------------------------------------------
23
24
-- Re-exporting type publically
@@ -56,8 +57,14 @@ record′-injective₂ refl = refl
56
57
record′-injective : ∀ {c c′ fs fs′} → record′ c fs ≡ record′ c′ fs′ → c ≡ c′ × fs ≡ fs′
57
58
record′-injective = < record′-injective₁ , record′-injective₂ >
58
59
59
- constructor′-injective : ∀ {c c′} → constructor′ c ≡ constructor′ c′ → c ≡ c′
60
- constructor′-injective refl = refl
60
+ constructor′-injective₁ : ∀ {c c′ q q′} → constructor′ c q ≡ constructor′ c′ q′ → c ≡ c′
61
+ constructor′-injective₁ refl = refl
62
+
63
+ constructor′-injective₂ : ∀ {c c′ q q′} → constructor′ c q ≡ constructor′ c′ q′ → q ≡ q′
64
+ constructor′-injective₂ refl = refl
65
+
66
+ constructor′-injective : ∀ {c c′ q q′} → constructor′ c q ≡ constructor′ c′ q′ → c ≡ c′ × q ≡ q′
67
+ constructor′-injective = < constructor′-injective₁ , constructor′-injective₂ >
61
68
62
69
infix 4 _≟_
63
70
@@ -70,39 +77,40 @@ data-type pars cs ≟ data-type pars′ cs′ =
70
77
record′ c fs ≟ record′ c′ fs′ =
71
78
map′ (uncurry (cong₂ record′)) record′-injective
72
79
(c Name.≟ c′ ×-dec List.≡-dec (Arg.≡-dec Name._≟_) fs fs′)
73
- constructor′ d ≟ constructor′ d′ =
74
- map′ (cong constructor′) constructor′-injective (d Name.≟ d′)
80
+ constructor′ d q ≟ constructor′ d′ q′ =
81
+ map′ (uncurry (cong₂ constructor′)) constructor′-injective
82
+ (d Name.≟ d′ ×-dec q Quantity.≟ q′)
75
83
axiom ≟ axiom = yes refl
76
84
primitive′ ≟ primitive′ = yes refl
77
85
78
86
-- antidiagonal
79
87
function cs ≟ data-type pars cs₁ = no (λ ())
80
88
function cs ≟ record′ c fs = no (λ ())
81
- function cs ≟ constructor′ d = no (λ ())
89
+ function cs ≟ constructor′ d q = no (λ ())
82
90
function cs ≟ axiom = no (λ ())
83
91
function cs ≟ primitive′ = no (λ ())
84
92
data-type pars cs ≟ function cs₁ = no (λ ())
85
93
data-type pars cs ≟ record′ c fs = no (λ ())
86
- data-type pars cs ≟ constructor′ d = no (λ ())
94
+ data-type pars cs ≟ constructor′ d q = no (λ ())
87
95
data-type pars cs ≟ axiom = no (λ ())
88
96
data-type pars cs ≟ primitive′ = no (λ ())
89
97
record′ c fs ≟ function cs = no (λ ())
90
98
record′ c fs ≟ data-type pars cs = no (λ ())
91
- record′ c fs ≟ constructor′ d = no (λ ())
99
+ record′ c fs ≟ constructor′ d q = no (λ ())
92
100
record′ c fs ≟ axiom = no (λ ())
93
101
record′ c fs ≟ primitive′ = no (λ ())
94
- constructor′ d ≟ function cs = no (λ ())
95
- constructor′ d ≟ data-type pars cs = no (λ ())
96
- constructor′ d ≟ record′ c fs = no (λ ())
97
- constructor′ d ≟ axiom = no (λ ())
98
- constructor′ d ≟ primitive′ = no (λ ())
102
+ constructor′ d q ≟ function cs = no (λ ())
103
+ constructor′ d q ≟ data-type pars cs = no (λ ())
104
+ constructor′ d q ≟ record′ c fs = no (λ ())
105
+ constructor′ d q ≟ axiom = no (λ ())
106
+ constructor′ d q ≟ primitive′ = no (λ ())
99
107
axiom ≟ function cs = no (λ ())
100
108
axiom ≟ data-type pars cs = no (λ ())
101
109
axiom ≟ record′ c fs = no (λ ())
102
- axiom ≟ constructor′ d = no (λ ())
110
+ axiom ≟ constructor′ d q = no (λ ())
103
111
axiom ≟ primitive′ = no (λ ())
104
112
primitive′ ≟ function cs = no (λ ())
105
113
primitive′ ≟ data-type pars cs = no (λ ())
106
114
primitive′ ≟ record′ c fs = no (λ ())
107
- primitive′ ≟ constructor′ d = no (λ ())
115
+ primitive′ ≟ constructor′ d q = no (λ ())
108
116
primitive′ ≟ axiom = no (λ ())
0 commit comments