Skip to content

Commit 52e361d

Browse files
committed
Compatibility with Agda PR #7322: add quantity to reflected syntax of constructor
1 parent 8a623ed commit 52e361d

File tree

2 files changed

+26
-18
lines changed

2 files changed

+26
-18
lines changed

src/Reflection/AST/Definition.agda

+25-17
Original file line numberDiff line numberDiff line change
@@ -15,9 +15,10 @@ open import Relation.Nullary.Decidable.Core using (map′; _×-dec_;
1515
open import Relation.Binary.Definitions using (DecidableEquality)
1616
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong; cong₂)
1717

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
2122

2223
------------------------------------------------------------------------
2324
-- Re-exporting type publically
@@ -56,8 +57,14 @@ record′-injective₂ refl = refl
5657
record′-injective : {c c′ fs fs′} record′ c fs ≡ record′ c′ fs′ c ≡ c′ × fs ≡ fs′
5758
record′-injective = < record′-injective₁ , record′-injective₂ >
5859

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₂ >
6168

6269
infix 4 _≟_
6370

@@ -70,39 +77,40 @@ data-type pars cs ≟ data-type pars′ cs′ =
7077
record′ c fs ≟ record′ c′ fs′ =
7178
map′ (uncurry (cong₂ record′)) record′-injective
7279
(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′)
7583
axiom ≟ axiom = yes refl
7684
primitive′ ≟ primitive′ = yes refl
7785

7886
-- antidiagonal
7987
function cs ≟ data-type pars cs₁ = no (λ ())
8088
function cs ≟ record′ c fs = no (λ ())
81-
function cs ≟ constructor′ d = no (λ ())
89+
function cs ≟ constructor′ d q = no (λ ())
8290
function cs ≟ axiom = no (λ ())
8391
function cs ≟ primitive′ = no (λ ())
8492
data-type pars cs ≟ function cs₁ = no (λ ())
8593
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 (λ ())
8795
data-type pars cs ≟ axiom = no (λ ())
8896
data-type pars cs ≟ primitive′ = no (λ ())
8997
record′ c fs ≟ function cs = no (λ ())
9098
record′ c fs ≟ data-type pars cs = no (λ ())
91-
record′ c fs ≟ constructor′ d = no (λ ())
99+
record′ c fs ≟ constructor′ d q = no (λ ())
92100
record′ c fs ≟ axiom = no (λ ())
93101
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 (λ ())
99107
axiom ≟ function cs = no (λ ())
100108
axiom ≟ data-type pars cs = no (λ ())
101109
axiom ≟ record′ c fs = no (λ ())
102-
axiom ≟ constructor′ d = no (λ ())
110+
axiom ≟ constructor′ d q = no (λ ())
103111
axiom ≟ primitive′ = no (λ ())
104112
primitive′ ≟ function cs = no (λ ())
105113
primitive′ ≟ data-type pars cs = no (λ ())
106114
primitive′ ≟ record′ c fs = no (λ ())
107-
primitive′ ≟ constructor′ d = no (λ ())
115+
primitive′ ≟ constructor′ d q = no (λ ())
108116
primitive′ ≟ axiom = no (λ ())

src/Reflection/AST/Show.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -149,6 +149,6 @@ showDefinition (data-type pars cs) =
149149
"datatype" <+> ℕ.show pars <+> braces (intersperse ", " (map showName cs))
150150
showDefinition (record′ c fs) =
151151
"record" <+> showName c <+> braces (intersperse ", " (map (showName ∘′ unArg) fs))
152-
showDefinition (constructor′ d) = "constructor" <+> showName d
152+
showDefinition (constructor′ d q) = "constructor" <+> showName d <+> showQuantity q
153153
showDefinition axiom = "axiom"
154154
showDefinition primitive′ = "primitive"

0 commit comments

Comments
 (0)