Skip to content

Commit 7042909

Browse files
committed
derive.eqbOK.register_axiom
1 parent e8a6ab8 commit 7042909

File tree

8 files changed

+66
-16
lines changed

8 files changed

+66
-16
lines changed

apps/derive/elpi/eqType.elpi

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -64,8 +64,8 @@ pred self o:eqb.trm.
6464

6565
pred valid i:eqb.trm, o:diagnostic.
6666
valid (eqb.global X) ok :- global X = {{ PrimInt63.int }}, !.
67-
valid (eqb.global (indt I)) ok :- eqType I _, !.
68-
valid (eqb.app (indt I) A Args) D :- eqType I EQT, !, valid-eqType EQT [A|Args] D.
67+
valid (eqb.global GR) ok :- eqType GR _, !.
68+
valid (eqb.app GR A Args) D :- eqType GR EQT, !, valid-eqType EQT [A|Args] D.
6969
valid T (error S) :- S is "not an eqType: " ^ {std.any->string T}.
7070

7171
pred valid-eqType i:eqb.eqType, i:list eqb.trm, o:diagnostic.
@@ -79,10 +79,10 @@ valid-eqType (eqb.value-param _ F) [_|TS] D :- std.do-ok! D [
7979
].
8080

8181
pred irrelevant? i:term, o:eqb.trm, o:diagnostic.
82-
irrelevant? (app [{{ @eq }}, global (indt EqType), A, B]) (eqb.app EQ EQTYPE [A1,B1]) D :- std.do-ok! D [
82+
irrelevant? (app [{{ @eq }}, global EqType, A, B]) (eqb.app EQ EQTYPE [A1,B1]) D :- std.do-ok! D [
8383
std.lift-ok (eqType EqType _) "Not an eqType", %eqb-for Bool Bool _,
8484
std.lift-ok ({{ @eq }} = global EQ) "",
85-
term->trm (global (indt EqType)) EQTYPE,
85+
term->trm (global EqType) EQTYPE,
8686
term->trm A A1,
8787
term->trm B B1,
8888
].
@@ -138,7 +138,7 @@ pred main i:inductive, o:list prop.
138138
main I [C] :-
139139
std.assert-ok! (translate-indt I EQT) "derive.eqType.ast: translate",
140140
std.assert-ok! (validate-eqType EQT) "derive.eqType.ast: validate",
141-
C = (eqType I EQT),
141+
C = (eqType (indt I) EQT),
142142
coq.elpi.accumulate _ "derive.eqType.db" (clause _ _ C).
143143

144144
}

apps/derive/elpi/eqb.elpi

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ macro @pi-trm N T F :-
99

1010
pred derive.eqb.main i:gref, i:string, o:list prop.
1111
derive.eqb.main (indt I) Prefix CL :- std.do! [
12-
std.assert! (eqType I FI) "this inductive is not supported",
12+
std.assert! (eqType (indt I) FI) "this inductive is not supported",
1313

1414
derive.eqb.eqbf.main FI FI [] [] R,
1515
std.assert-ok! (coq.typecheck R Rty) "derive.eqbf generates illtyped term",

apps/derive/elpi/eqbOK.elpi

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,7 @@ add-reflect (eqb.inductive _ _) Correct Refl {{iffP2 lp:Correct lp:Refl}}.
3434

3535
pred main i:gref, i:string, o:list prop.
3636
main (indt I) Prefix [CL] :- std.do! [
37-
std.assert! (eqType I FI) "this inductive is not supported",
37+
std.assert! (eqType (indt I) FI) "this inductive is not supported",
3838
std.assert! (eqcorrect-for (indt I) Correct Refl) "run eqbcorrect before",
3939

4040
add-reflect FI (global (const Correct)) (global (const Refl)) Breflect,

apps/derive/elpi/eqbcorrect.elpi

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -23,7 +23,7 @@ pred config o:term, o:term, o:term, o:
2323
pred main i:gref, i:string, o:list prop.
2424
main (indt I) Prefix CLs :- std.do! [
2525

26-
std.assert! (eqType I FI) "this inductive is not supported",
26+
std.assert! (eqType (indt I) FI) "this inductive is not supported",
2727

2828
std.assert! (induction-db I Indu) "call derive.induction before",
2929

@@ -102,11 +102,11 @@ main (indc _) _ _ :- stop "derive.eqbcorrect does not work on a constructor".
102102

103103
pred search-eqcorrect-for i:term, o:term, o:term.
104104
search-eqcorrect-for (global (indt I)) (global (const C)) (global (const R)) :-
105-
std.assert! (eqType I (eqb.inductive _ _)) "unknown or not applied enough type",
105+
std.assert! (eqType (indt I) (eqb.inductive _ _)) "unknown or not applied enough type",
106106
eqcorrect-for (indt I) C R.
107107

108108
search-eqcorrect-for (app[global (indt I)|Args]) CArgs RArgs :-
109-
std.assert! (eqType I F) "unknown",
109+
std.assert! (eqType (indt I) F) "unknown",
110110
eqcorrect-for (indt I) C R,
111111
search-eqcorrect-apply F Args (global (const C)) (global (const R)) CArgs RArgs.
112112

apps/derive/elpi/fields.elpi

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ main I Prefix AllCL :- std.do! [
2121

2222
std.assert! (tag-for I Tag) "no tag for this inductive, run that derivation first",
2323

24-
std.assert! (eqType I FI) "this inductive is not supported",
24+
std.assert! (eqType(indt I) FI) "this inductive is not supported",
2525
coq.env.indt I _ _ _ Arity KS _,
2626

2727
box (global (indt I)) KS FI CLB,

apps/derive/tests-stdlib/test_derive.v

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -170,3 +170,20 @@ End derive_container.
170170
About wimpls.wimpls.
171171
About wimpls.Kwi.
172172
Redirect "tmp" Check Kwi _ (refl_equal 3).
173+
174+
Section TestRegister.
175+
Variable T : Type.
176+
Definition is_T : T -> Type := fun x => True.
177+
Definition is_T_inhab : forall x : T, is_T x := fun x => I.
178+
Variable eqb : T -> T -> bool.
179+
Variable eqb_correct : forall x y,eqb x y = true -> x = y.
180+
Variable eqb_refl : forall x, eqb x x = true.
181+
182+
derive.eqbOK.register_axiom T is_T is_T_inhab eqb eqb_correct eqb_refl.
183+
184+
Inductive foo := X : T -> foo.
185+
186+
#[only(eqbOK)] derive foo.
187+
188+
Redirect "tmp" Print is_foo_inhab.
189+
End TestRegister.

apps/derive/theories/derive/eqType_ast.v

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -24,10 +24,11 @@ type eqb.stop eqb.trm -> eqb.arguments.
2424
type eqb.type-param (eqb.trm -> eqb.eqType) -> eqb.eqType.
2525
type eqb.value-param eqb.trm -> (eqb.trm -> eqb.eqType) -> eqb.eqType.
2626
type eqb.inductive inductive -> (eqb.trm -> list eqb.constructor) -> eqb.eqType.
27+
type eqb.axiom eqb.eqType.
2728

2829
type eqb.constructor constructor -> eqb.arguments -> eqb.constructor.
2930

30-
pred eqType i:inductive, o:eqb.eqType.
31+
pred eqType i:gref, o:eqb.eqType.
3132

3233
}}.
3334

@@ -60,6 +61,6 @@ Elpi Accumulate derive File eqType.
6061

6162
Elpi Accumulate derive lp:{{
6263

63-
derivation (indt T) _ ff (derive "eqType_ast" (derive.eqType.ast.main T) (eqType T _)).
64+
derivation (indt T) _ ff (derive "eqType_ast" (derive.eqType.ast.main T) (eqType (indt T) _)).
6465

6566
}}.

apps/derive/theories/derive/eqbOK.v

Lines changed: 35 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -2,9 +2,9 @@
22
33
license: GNU Lesser General Public License Version 2.1 or later
44
------------------------------------------------------------------------- *)
5-
6-
Require Import eqb_core_defs.
7-
Require Import tag eqType_ast fields eqb eqbcorrect derive.
5+
6+
From elpi.apps.derive Require Import eqb_core_defs.
7+
From elpi.apps.derive Require Import tag eqType_ast fields eqb eqbcorrect derive.
88

99
From elpi.apps.derive.elpi Extra Dependency "eqbOK.elpi" as eqbOK.
1010
From elpi.apps.derive.elpi Extra Dependency "eqType.elpi" as eqType.
@@ -61,3 +61,35 @@ derivation (indt T) Prefix ff (derive "eqbOK" (derive.eqbOK.main (indt T) Prefix
6161
derivation (const T) Prefix ff (derive "eqbOK-alias" (derive.eqbOK.main (const T) Prefix) (eqbok-for (const T) _)).
6262

6363
}}.
64+
65+
66+
Elpi Command derive.eqbOK.register_axiom.
67+
Elpi Accumulate Db derive.eqb.db.
68+
Elpi Accumulate Db derive.eqbcorrect.db.
69+
Elpi Accumulate Db derive.param1.db.
70+
Elpi Accumulate Db derive.param1.trivial.db.
71+
Elpi Accumulate Db derive.eqType.db.
72+
Elpi Accumulate File eqType.
73+
Elpi Accumulate lp:{{
74+
main [str Type, str IsT, str IsTinhab, str Eqb, str Correct, str Refl] :- !,
75+
coq.locate Type GrType,
76+
coq.locate IsT GRisT,
77+
coq.locate IsTinhab GRisTinhab,
78+
coq.locate Eqb GrEqb,
79+
coq.locate Correct GrCorrect,
80+
coq.locate Refl GrRefl,
81+
GrRefl = const R,
82+
GrCorrect = const C,
83+
coq.elpi.accumulate _ "derive.eqb.db" (clause _ _ (eqb-done GrType)),
84+
coq.elpi.accumulate _ "derive.eqb.db" (clause _ _ (eqb-for (global GrType) (global GrType) (global GrEqb))),
85+
coq.elpi.accumulate _ "derive.eqbcorrect.db" (clause _ _ (eqcorrect-for GrType C R)),
86+
coq.elpi.accumulate _ "derive.eqbcorrect.db" (clause _ _ (correct-lemma-for (global GrType) (global GrCorrect))),
87+
coq.elpi.accumulate _ "derive.eqbcorrect.db" (clause _ _ (refl-lemma-for (global GrType) (global GrRefl))),
88+
coq.elpi.accumulate _ "derive.eqType.db" (clause _ _ (eqType GrType eqb.axiom)),
89+
coq.elpi.accumulate _ "derive.param1.db" (clause _ _ (reali-done GrType)),
90+
coq.elpi.accumulate _ "derive.param1.db" (clause _ (before "reali:fail") (reali (global GrType) (global GRisT) :- !)),
91+
coq.elpi.accumulate _ "derive.param1.db" (clause _ (before "realiR:fail") (realiR (global GrType) (global GRisT) :- !)),
92+
coq.elpi.accumulate _ "derive.param1.trivial.db" (clause _ _ (param1-inhab-db (global GRisT) (global GRisTinhab))).
93+
main _ :- coq.error "usage: derive.eqbOK.register_axiom T is_T is_T_inhab eqb eqb_correct eqb_refl.".
94+
}}.
95+
Elpi Export derive.eqbOK.register_axiom.

0 commit comments

Comments
 (0)