diff --git a/Makefile b/Makefile index 8c91fd0ab..005a37477 100644 --- a/Makefile +++ b/Makefile @@ -91,13 +91,19 @@ doc: $(DOCDEP) .PHONY: force build all test doc -Makefile.coq Makefile.coq.conf: src/coq_elpi_builtins_HOAS.ml src/coq_elpi_config.ml _CoqProject +Makefile.coq Makefile.coq.conf: src/coq_elpi_builtins_HOAS.ml src/coq_elpi_builtins_arg_HOAS.ml src/coq_elpi_config.ml _CoqProject @$(COQBIN)/coq_makefile -f _CoqProject -o Makefile.coq @$(MAKE) --no-print-directory -f Makefile.coq .merlin Makefile.test.coq Makefile.test.coq.conf: _CoqProject.test @$(COQBIN)/coq_makefile -f _CoqProject.test -o Makefile.test.coq Makefile.examples.coq Makefile.examples.coq.conf: _CoqProject.examples @$(COQBIN)/coq_makefile -f _CoqProject.examples -o Makefile.examples.coq +src/coq_elpi_builtins_arg_HOAS.ml: elpi/coq-arg-HOAS.elpi Makefile.coq.local + echo "(* Automatically generated from $<, don't edit *)" > $@ + echo "(* Regenerate via 'make $@' *)" >> $@ + echo "let code = {|" >> $@ + cat $< >> $@ + echo "|}" >> $@ src/coq_elpi_builtins_HOAS.ml: elpi/coq-HOAS.elpi Makefile.coq.local echo "(* Automatically generated from $<, don't edit *)" > $@ echo "(* Regenerate via 'make $@' *)" >> $@ diff --git a/Makefile.coq.local b/Makefile.coq.local index ae2d441fc..2be5c1670 100644 --- a/Makefile.coq.local +++ b/Makefile.coq.local @@ -17,6 +17,7 @@ install-extra:: df="`$(COQMKFILE) -destination-of theories/elpi.vo $(COQLIBS)`";\ install -m 0644 elpi-builtin.elpi "$(COQLIBINSTALL)/$$df";\ install -m 0644 coq-builtin.elpi "$(COQLIBINSTALL)/$$df";\ + install -m 0644 coq-builtin-synterp.elpi "$(COQLIBINSTALL)/$$df";\ install -m 0644 elpi/coq-lib.elpi "$(COQLIBINSTALL)/$$df";\ install -m 0644 elpi/elpi_elaborator.elpi "$(COQLIBINSTALL)/$$df" diff --git a/_CoqProject b/_CoqProject index 563bd9ebf..345c05f09 100644 --- a/_CoqProject +++ b/_CoqProject @@ -15,6 +15,7 @@ # NES -R apps/NES/theories elpi.apps +-R apps/NES/elpi elpi.apps.NES -R apps/NES/tests elpi.apps.NES.tests -R apps/NES/examples elpi.apps.NES.examples @@ -59,7 +60,10 @@ src/coq_elpi_glob_quotation.mli src/coq_elpi_arg_HOAS.ml src/coq_elpi_arg_HOAS.mli src/coq_elpi_arg_syntax.mlg +src/coq_elpi_builtins_arg_HOAS.ml src/coq_elpi_builtins_HOAS.ml +src/coq_elpi_builtins_synterp.ml +src/coq_elpi_builtins_synterp.mli src/coq_elpi_builtins.ml src/coq_elpi_builtins.mli src/coq_elpi_config.ml diff --git a/_CoqProject.test b/_CoqProject.test index 4093a2ebc..71e43c409 100644 --- a/_CoqProject.test +++ b/_CoqProject.test @@ -54,3 +54,4 @@ tests/test_link_order_import2.v tests/test_link_order_import3.v tests/test_query_extra_dep.v tests/test_toposort.v +tests/test_synterp.v diff --git a/apps/NES/elpi/nes_interp.elpi b/apps/NES/elpi/nes_interp.elpi new file mode 100644 index 000000000..ae85a0475 --- /dev/null +++ b/apps/NES/elpi/nes_interp.elpi @@ -0,0 +1,85 @@ +namespace nes { + +% Print a namespace +pred print-path i:list string, i:(gref -> coq.pp -> prop). +print-path Path PP :- std.do! [ + std.map {std.findall (ns Path _)} (p\ mp\ p = ns _ mp) MPs, + print.pp-list MPs (print.pp-module Path PP) Out, + coq.say {coq.pp->string Out}, +]. + +namespace print { + + pred pp-list i:list A, i:(A -> coq.pp -> prop), o:coq.pp. + pp-list L F Out :- std.do! [ + std.map-filter L F PPs, + Out = coq.pp.box (coq.pp.v 0) {std.intersperse (coq.pp.brk 0 0) PPs}, + ]. + + kind context type. + type context + list string -> % readable path + int -> % length of full path + (gref -> coq.pp -> prop) -> + context. + + % Hides `aux` modules + pred readable-path i:context, i:list string, o:list string. + readable-path (context Prefix N _) FullPath Path :- std.do! [ + std.drop N FullPath RelPath, + std.append Prefix RelPath Path, + ]. + + pred module-context i:list string, i:modpath, i:(gref -> coq.pp -> prop), o:context. + module-context Prefix MP PP Ctx :- std.do! [ + coq.modpath->path MP FullPath, + Ctx = context Prefix {std.length FullPath} PP, + ]. + + pred submodule-context i:context, i:modpath, o:context. + submodule-context (context _ _ PP as Ctx) MP Ctx' :- std.do! [ + coq.modpath->path MP FullPath, + readable-path Ctx FullPath Path, + Ctx' = context Path {std.length FullPath} PP, + ]. + + pred pp-module i:list string, i:(gref -> coq.pp -> prop), i:modpath, o:coq.pp. + pp-module Prefix PP MP Out :- std.do! [ + pp-module-items {module-context Prefix MP PP} {coq.env.module MP} Out, + ]. + + pred pp-module-items i:context i:list module-item, o:coq.pp. + pp-module-items Ctx Items Out :- + pp-list Items (pp-module-item Ctx) Out. + + pred pp-module-item i:context, i:module-item, o:coq.pp. + pp-module-item (context _ _ PP) (gref GR) Out :- PP GR Out, !. + pp-module-item Ctx (submodule MP Items) Out :- std.do! [ + pp-module-items {submodule-context Ctx MP} Items Out, + ]. + pp-module-item Ctx (module-type MTP) Out :- pp-modtypath Ctx MTP Out. + pp-module-item Ctx (module-type-functor MTP _) Out :- pp-modtypath Ctx MTP Out. + pp-module-item Ctx (module-functor MP _) Out :- pp-modpath Ctx MP Out. + + pred pp-path i:context, i:string, i:list string, o:coq.pp. + pp-path Ctx What FullPath Out :- std.do! [ + readable-path Ctx FullPath Path, + Out = coq.pp.box coq.pp.h [ + coq.pp.str What, coq.pp.spc, + coq.pp.str {std.string.concat "." Path}, + ], + ]. + + pred pp-modtypath i:context, i:modtypath, o:coq.pp. + pp-modtypath Ctx MTP Out :- std.do! [ + pp-path Ctx "Module Type" {coq.modtypath->path MTP} Out, + ]. + + pred pp-modpath i:context, i:modpath, o:coq.pp. + pp-modpath Ctx MP Out :- std.do! [ + pp-path Ctx "Module" {coq.modpath->path MP} Out, + ]. +} + + +} \ No newline at end of file diff --git a/apps/NES/elpi/nes.elpi b/apps/NES/elpi/nes_synterp.elpi similarity index 62% rename from apps/NES/elpi/nes.elpi rename to apps/NES/elpi/nes_synterp.elpi index 525d29f65..70f9db68d 100644 --- a/apps/NES/elpi/nes.elpi +++ b/apps/NES/elpi/nes_synterp.elpi @@ -86,13 +86,7 @@ string->ns "" [] :- !. string->ns S L :- string->non-empty-ns S L. pred ns->string i:list string, o:string. -ns->string NS S :- nes.join "." NS S. - -:index (_ 1) -pred join i:string, i:list string, o:string. -join _ [] "". -join _ [X] X :- !. -join Sep [X|XS] S :- join Sep XS S0, S is X ^ Sep ^ S0. +ns->string NS S :- std.string.concat "." NS S. pred begin-path i:list string, o:list prop. begin-path [] []. @@ -101,7 +95,7 @@ begin-path ([_|_] as Path) Out :- std.do! [ if (open-ns _ NSCP) (std.assert! (NSCP = CP) "NS: cannot begin a namespace inside a module that is inside a namespace") true, std.map {std.findall (open-ns Y_ P_)} open-ns->string Stack, if (Stack = []) true (std.do! [ - coq.locate-all {join "." Path} L, + coq.locate-all {std.string.concat "." Path} L, if (std.do! [ std.mem L (loc-modpath M), coq.modpath->path M MP, @@ -173,93 +167,5 @@ resolve.walk S Ctx SP Path :- std.do! [ pred resolve.err i:string. resolve.err S :- coq.error "NES: Namespace not found:" S. -% The (closed) namespace [NS] containing global [GR], or []. -pred gref->path i:gref, o:list string. -gref->path GR NS :- std.do! [ - coq.gref->path GR MP, - if (gref->path.aux MP NS) true (NS = []), -]. -pred gref->path.aux i:list string, o:list string. -gref->path.aux MP NS :- ns NS M, coq.modpath->path M MP, !. - -% Print a namespace -pred print-path i:list string, i:(gref -> coq.pp -> prop). -print-path Path PP :- std.do! [ - std.map {std.findall (ns Path _)} (p\ mp\ p = ns _ mp) MPs, - print.pp-list MPs (print.pp-module Path PP) Out, - coq.say {coq.pp->string Out}, -]. -namespace print { - - pred pp-list i:list A, i:(A -> coq.pp -> prop), o:coq.pp. - pp-list L F Out :- std.do! [ - std.map-filter L F PPs, - Out = coq.pp.box (coq.pp.v 0) {std.intersperse (coq.pp.brk 0 0) PPs}, - ]. - - kind context type. - type context - list string -> % readable path - int -> % length of full path - (gref -> coq.pp -> prop) -> - context. - - % Hides `aux` modules - pred readable-path i:context, i:list string, o:list string. - readable-path (context Prefix N _) FullPath Path :- std.do! [ - std.drop N FullPath RelPath, - std.append Prefix RelPath Path, - ]. - - pred module-context i:list string, i:modpath, i:(gref -> coq.pp -> prop), o:context. - module-context Prefix MP PP Ctx :- std.do! [ - coq.modpath->path MP FullPath, - Ctx = context Prefix {std.length FullPath} PP, - ]. - - pred submodule-context i:context, i:modpath, o:context. - submodule-context (context _ _ PP as Ctx) MP Ctx' :- std.do! [ - coq.modpath->path MP FullPath, - readable-path Ctx FullPath Path, - Ctx' = context Path {std.length FullPath} PP, - ]. - - pred pp-module i:list string, i:(gref -> coq.pp -> prop), i:modpath, o:coq.pp. - pp-module Prefix PP MP Out :- std.do! [ - pp-module-items {module-context Prefix MP PP} {coq.env.module MP} Out, - ]. - - pred pp-module-items i:context i:list module-item, o:coq.pp. - pp-module-items Ctx Items Out :- - pp-list Items (pp-module-item Ctx) Out. - - pred pp-module-item i:context, i:module-item, o:coq.pp. - pp-module-item (context _ _ PP) (gref GR) Out :- PP GR Out, !. - pp-module-item Ctx (submodule MP Items) Out :- std.do! [ - pp-module-items {submodule-context Ctx MP} Items Out, - ]. - pp-module-item Ctx (module-type MTP) Out :- pp-modtypath Ctx MTP Out. - pp-module-item Ctx (module-type-functor MTP _) Out :- pp-modtypath Ctx MTP Out. - pp-module-item Ctx (module-functor MP _) Out :- pp-modpath Ctx MP Out. - - pred pp-path i:context, i:string, i:list string, o:coq.pp. - pp-path Ctx What FullPath Out :- std.do! [ - readable-path Ctx FullPath Path, - Out = coq.pp.box coq.pp.h [ - coq.pp.str What, coq.pp.spc, - coq.pp.str {nes.join "." Path}, - ], - ]. - - pred pp-modtypath i:context, i:modtypath, o:coq.pp. - pp-modtypath Ctx MTP Out :- std.do! [ - pp-path Ctx "Module Type" {coq.modtypath->path MTP} Out, - ]. - - pred pp-modpath i:context, i:modpath, o:coq.pp. - pp-modpath Ctx MP Out :- std.do! [ - pp-path Ctx "Module" {coq.modpath->path MP} Out, - ]. -} } diff --git a/apps/NES/tests/test_NES_lib.v b/apps/NES/tests/test_NES_lib.v index b3ada6aea..fdf64df58 100644 --- a/apps/NES/tests/test_NES_lib.v +++ b/apps/NES/tests/test_NES_lib.v @@ -1,23 +1,29 @@ -From elpi.apps.NES Extra Dependency "nes.elpi" as nes. +From elpi.apps.NES Extra Dependency "nes_synterp.elpi" as nes_synterp. From elpi.apps Require Import NES. Elpi Command Make. -Elpi Accumulate Db NES.db. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate File nes_synterp. +#[synterp] Elpi Accumulate lp:{{ - main [str Path] :- std.do! [ + main-synterp [str Path] ActionsToOpen :- std.do! [ nes.string->ns Path NS, nes.begin-path NS OpenNS, - OpenNS => std.do! [ - coq.env.add-const "x" {{ 42 }} _ @transparent! _C, - nes.end-path NS _NewNS, - ], + coq.synterp-actions ActionsToOpen, + OpenNS => nes.end-path NS _NewNS, ]. main _ :- coq.error "usage: Make ". }}. +#[interp] Elpi Accumulate lp:{{ + main-interp [str _] ActionsBefore :- std.do! [ + std.forall ActionsBefore coq.replay-synterp-action, + coq.env.add-const "x" {{ 42 }} _ @transparent! _C, + coq.replay-all-missing-synterp-actions, + ]. +}}. Elpi Typecheck. Elpi Export Make. Make Cats.And.Dogs. +Print Cats.And.Dogs.x. diff --git a/apps/NES/theories/NES.v b/apps/NES/theories/NES.v index 4114f2c89..7758e6b9f 100644 --- a/apps/NES/theories/NES.v +++ b/apps/NES/theories/NES.v @@ -1,8 +1,9 @@ -From elpi.apps.NES Extra Dependency "nes.elpi" as nes. +From elpi.apps.NES Extra Dependency "nes_synterp.elpi" as nes_synterp. +From elpi.apps.NES Extra Dependency "nes_interp.elpi" as nes_interp. From elpi Require Import elpi. -Elpi Db NES.db lp:{{ +#[synterp] Elpi Db NES.db lp:{{ pred open-ns o:string, o:list string. :name "open-ns:begin" @@ -16,9 +17,9 @@ pred ns o:path, o:modpath. }}. Elpi Command NES.Status. -Elpi Accumulate Db NES.db. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate File nes_synterp. +#[synterp] Elpi Accumulate lp:{{ main _ :- coq.say "NES: current namespace" {nes.current-path}, @@ -30,52 +31,61 @@ Elpi Typecheck. Elpi Export NES.Status. Elpi Command NES.Begin. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate File nes_synterp. +#[synterp] Elpi Accumulate lp:{{ - main [str NS] :- nes.begin-path {nes.string->non-empty-ns NS} _. + main [str NS] :- !, nes.begin-path {nes.string->non-empty-ns NS} _. main _ :- coq.error "usage: NES.Begin ". }}. -Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate Db NES.db. +#[interp] Elpi Accumulate lp:{{ main _ :- coq.replay-all-missing-synterp-actions. }}. Elpi Typecheck. Elpi Export NES.Begin. Elpi Command NES.End. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate File nes_synterp. +#[synterp] Elpi Accumulate lp:{{ main [str NS] :- nes.end-path {nes.string->non-empty-ns NS} _. main _ :- coq.error "usage: NES.End ". }}. -Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate Db NES.db. +#[interp] Elpi Accumulate lp:{{ main _ :- coq.replay-all-missing-synterp-actions. }}. Elpi Typecheck. Elpi Export NES.End. Elpi Command NES.Open. -Elpi Accumulate Db NES.db. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate File nes_synterp. +#[synterp] Elpi Accumulate lp:{{ main [str NS] :- nes.open-path {nes.resolve NS}. main _ :- coq.error "usage: NES.Open ". }}. +#[interp] Elpi Accumulate lp:{{ main _ :- coq.replay-all-missing-synterp-actions. }}. Elpi Typecheck. Elpi Export NES.Open. (* List the contents a namespace *) Elpi Command NES.List. -Elpi Accumulate Db NES.db. -Elpi Accumulate File nes. -Elpi Accumulate lp:{{ +#[synterp] Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate File nes_synterp. +#[interp] Elpi Accumulate File nes_interp. +#[synterp] Elpi Accumulate lp:{{ + main-synterp [str NS] (pr DB Path) :- nes.resolve NS Path, std.findall (ns O_ P_) DB. +}}. +#[interp] Elpi Accumulate lp:{{ + typeabbrev path (list string). + pred ns o:path, o:modpath. pred pp-gref i:gref, o:coq.pp. pp-gref GR PP :- coq.term->pp (global GR) PP. - main [str NS] :- nes.print-path {nes.resolve NS} pp-gref. + main-interp [str _] (pr DB Path) :- DB => nes.print-path Path pp-gref. main _ :- coq.error "usage: NES.List ". }}. @@ -84,9 +94,15 @@ Elpi Export NES.List. (* NES.List with types *) Elpi Command NES.Print. -Elpi Accumulate Db NES.db. -Elpi Accumulate File nes. +#[synterp] Elpi Accumulate Db NES.db. +#[synterp] Elpi Accumulate File nes_synterp. +#[interp] Elpi Accumulate File nes_interp. +#[synterp] Elpi Accumulate lp:{{ + main-synterp [str NS] (pr DB Path) :- nes.resolve NS Path, std.findall (ns O_ P_) DB. +}}. Elpi Accumulate lp:{{ + typeabbrev path (list string). + pred ns o:path, o:modpath. pred pp-gref i:gref, o:coq.pp. pp-gref GR PP :- std.do! [ @@ -97,7 +113,7 @@ Elpi Accumulate lp:{{ ], ]. - main [str NS] :- nes.print-path {nes.resolve NS} pp-gref. + main-interp [str _] (pr DB Path) :- DB => nes.print-path Path pp-gref. main _ :- coq.error "usage: NES.Print ". }}. diff --git a/apps/coercion/src/coq_elpi_coercion_hook.mlg b/apps/coercion/src/coq_elpi_coercion_hook.mlg index 89b8d0d45..532c28abf 100644 --- a/apps/coercion/src/coq_elpi_coercion_hook.mlg +++ b/apps/coercion/src/coq_elpi_coercion_hook.mlg @@ -21,15 +21,17 @@ let elpi_coercion_hook program env sigma ~flags v ~inferred ~expected = let state = API.State.set Coq_elpi_builtins.tactic_mode state true in state, (loc, qatts), gls in - let cprogram, _ = get_and_compile program in - match run ~static_check:false cprogram (`Fun query) with - | API.Execute.Success solution -> - let gls = Evar.Set.singleton goal_evar in - let sigma, _, _ = Coq_elpi_HOAS.solution2evd sigma solution gls in - if Evd.is_defined sigma goal_evar then Some (sigma, goal) else None - | API.Execute.NoMoreSteps - | API.Execute.Failure -> None - | exception (Coq_elpi_utils.LtacFail (level, msg)) -> None + match Interp.get_and_compile program with + | None -> None + | Some (cprogram, _) -> + match Interp.run ~static_check:false cprogram (`Fun query) with + | API.Execute.Success solution -> + let gls = Evar.Set.singleton goal_evar in + let sigma, _, _ = Coq_elpi_HOAS.solution2evd sigma solution gls in + if Evd.is_defined sigma goal_evar then Some (sigma, goal) else None + | API.Execute.NoMoreSteps + | API.Execute.Failure -> None + | exception (Coq_elpi_utils.LtacFail (level, msg)) -> None let add_coercion_hook = let coercion_hook_program = Summary.ref ~name:"elpi-coercion" None in diff --git a/apps/derive/theories/derive.v b/apps/derive/theories/derive.v index 32d7cb2f4..5d92cf97e 100644 --- a/apps/derive/theories/derive.v +++ b/apps/derive/theories/derive.v @@ -86,6 +86,20 @@ main _ :- usage. usage :- coq.error "Usage: derive \n\tderive Inductive name Params : Arity := Constructors.". +}}. +#[synterp] Elpi Accumulate lp:{{ + main [indt-decl D] :- !, + declare-module-for-ind D. + main _. + + pred declare-module-for-ind i:indt-decl. + declare-module-for-ind (parameter _ _ _ F) :- + pi p\ declare-module-for-ind (F p). + declare-module-for-ind (inductive N _ _ _) :- + coq.env.begin-module N none, coq.env.end-module _. + declare-module-for-ind (record N _ _ _) :- + coq.env.begin-module N none, coq.env.end-module _. + }}. Elpi Typecheck. Elpi Export derive. diff --git a/apps/locker/theories/locker.v b/apps/locker/theories/locker.v index 0e433d1b7..f0112ee04 100644 --- a/apps/locker/theories/locker.v +++ b/apps/locker/theories/locker.v @@ -72,6 +72,17 @@ mlock Definition foo : T := bo. Elpi Command mlock. Elpi Accumulate File locker. +#[synterp] Elpi Accumulate lp:{{ + pred synterp-action i:id. + synterp-action ID :- + Module is ID ^ "_Locked", + coq.env.begin-module-type Module, + coq.env.end-module-type TY, + coq.env.begin-module ID (some TY), + coq.env.end-module _. + main [const-decl ID _ _] :- synterp-action ID. + main [upoly-const-decl ID _ _ _] :- synterp-action ID. +}}. Elpi Accumulate lp:{{ main [const-decl ID (some Bo) Ty] :- !, locker.module-lock ID Bo Ty none. main [upoly-const-decl ID (some Bo) Ty UD] :- !, locker.module-lock ID Bo Ty (some UD). diff --git a/apps/tc/src/coq_elpi_class_tactics_takeover.ml b/apps/tc/src/coq_elpi_class_tactics_takeover.ml index 24462edcf..e27e0c758 100644 --- a/apps/tc/src/coq_elpi_class_tactics_takeover.ml +++ b/apps/tc/src/coq_elpi_class_tactics_takeover.ml @@ -115,14 +115,16 @@ let solve_TC program env sigma depth unique ~best_effort filter = let state = API.State.set Coq_elpi_builtins.tactic_mode state true in state, (loc, qatts), gls in - let cprogram, _ = Coq_elpi_vernacular.get_and_compile program in - match Coq_elpi_vernacular.run ~static_check:false cprogram (`Fun query) with - | API.Execute.Success solution -> - let sigma, _, _ = Coq_elpi_HOAS.solution2evd sigma solution glss in - Some(false,sigma) - | API.Execute.NoMoreSteps -> CErrors.user_err Pp.(str "elpi run out of steps") - | API.Execute.Failure -> elpi_fails program - | exception (Coq_elpi_utils.LtacFail (level, msg)) -> elpi_fails program + match Coq_elpi_vernacular.Interp.get_and_compile program with + | None -> assert false + | Some (cprogram,_) -> + match Coq_elpi_vernacular.Interp.run ~static_check:false cprogram (`Fun query) with + | API.Execute.Success solution -> + let sigma, _, _ = Coq_elpi_HOAS.solution2evd sigma solution glss in + Some(false,sigma) + | API.Execute.NoMoreSteps -> CErrors.user_err Pp.(str "elpi run out of steps") + | API.Execute.Failure -> elpi_fails program + | exception (Coq_elpi_utils.LtacFail (level, msg)) -> elpi_fails program let handle_takeover coq_solver env sigma (cl: Intpart.set) = let t = Unix.gettimeofday () in diff --git a/apps/tc/src/coq_elpi_tc_register.ml b/apps/tc/src/coq_elpi_tc_register.ml index 6ec035370..1097ca12c 100644 --- a/apps/tc/src/coq_elpi_tc_register.ml +++ b/apps/tc/src/coq_elpi_tc_register.ml @@ -63,7 +63,7 @@ let inObservation1 = let observer_evt ((loc, name, atts) : loc_name_atts) (x : Event.t) = let open Coq_elpi_vernacular in - let run_program e = run_program loc name ~atts e in + let run_program e = Interp.run_program loc name ~syndata:None ~atts e in match x with | Event.NewClass cl -> Lib.add_leaf (inObservation (run_program,cl)) | Event.NewInstance inst -> Lib.add_leaf (inObservation1 (run_program,inst)) diff --git a/coq-builtin-synterp.elpi b/coq-builtin-synterp.elpi new file mode 100644 index 000000000..1572a503d --- /dev/null +++ b/coq-builtin-synterp.elpi @@ -0,0 +1,366 @@ + + +% -- Misc --------------------------------------------------------- + +% [coq.info ...] Prints an info message +external type coq.info variadic any prop. + +% [coq.notice ...] Prints a notice message +external type coq.notice variadic any prop. + +% [coq.say ...] Prints a notice message +external type coq.say variadic any prop. + +% [coq.warn ...] Prints a generic warning message +external type coq.warn variadic any prop. + +% [coq.warning Category Name ...] +% Prints a warning message with a Name and Category which can be used +% to silence this warning or turn it into an error. See coqc -w command +% line option +external type coq.warning string -> string -> variadic any prop. + +% [coq.error ...] Prints and *aborts* the program. It is a fatal error for +% Elpi and Ltac +external type coq.error variadic any prop. + +% [coq.version VersionString Major Minor Patch] Fetches the version of Coq, +% as a string and as 3 numbers +external pred coq.version o:string, o:int, o:int, o:int. + + +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-arg-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% This section contains the low level data types linking Coq and elpi. +% In particular the entry points for commands + + +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +% Entry points +% +% Command and tactic invocation +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% Entry point for commands. Eg. "#[att=true] Elpi mycommand foo 3 (f x)." becomes +% main [str "foo", int 3, trm (app[f,x])] +% in a context where +% attributes [attribute "att" (leaf "true")] +% holds. The encoding of terms is described below. +% See also the coq.parse-attributes utility. +pred main i:list argument. +pred main-interp i:list argument, i:any. +pred main-synterp i:list argument, o:any. +pred usage. +pred attributes o:list attribute. + +% see coq-lib.elpi for coq.parse-attributes generating the options below +type get-option string -> A -> prop. + +% The data type of arguments (for commands or tactics) +kind argument type. +type int int -> argument. % Eg. 1 -2. +type str string -> argument. % Eg. x "y" z.w. or any Coq keyword/symbol +type trm term -> argument. % Eg. (t). + +% Extra arguments for commands. [Definition], [Axiom], [Record] and [Context] +% take precedence over the [str] argument above (when not "quoted"). +% +% Eg. Record or Inductive +type indt-decl indt-decl -> argument. +% Eg. #[universes(polymorphic,...)] Record or Inductive +type upoly-indt-decl indt-decl -> upoly-decl -> argument. +type upoly-indt-decl indt-decl -> upoly-decl-cumul -> argument. +% Eg. Definition or Axiom (when the body is none) +type const-decl id -> option term -> arity -> argument. +% Eg. #[universes(polymorphic,...)] Definition or Axiom +type upoly-const-decl id -> option term -> arity -> upoly-decl -> argument. +% Eg. Context A (b : A). +type ctx-decl context-decl -> argument. + +% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +kind indt-decl type. +kind indc-decl type. +kind record-decl type. + +% An arity is written, in Coq syntax, as: +% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U +% This syntax is used, for example, in the type of an inductive type or +% in the type of constructors. We call the abstractions on the left of ":" +% "parameters" while we call the type following the ":" (proper) arity. + +% Note: in some contexts, like the type of an inductive type constructor, +% Coq makes no distinction between these two writings +% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... +% while Elpi is a bit more restrictive, since it understands user directives +% such as the implicit status of an arguments (eg, using {} instead of () around +% the binder), only on parameters. +% Moreover parameters carry the name given by the user as an "id", while binders +% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see +% also the HOAS of terms). A user command can hence only use the names of +% parameters, and not the names of "forall" quantified variables in the arity. +% +% See also the arity->term predicate in coq-lib.elpi + +type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. +type arity term -> arity. + +type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. +type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive +type record id -> term -> id -> record-decl -> indt-decl. + +type constructor id -> arity -> indc-decl. + +type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. +type end-record record-decl. + +% Example. +% Remark that A is a regular parameter; y is a non-uniform parameter and t +% also features an index of type bool. +% +% Inductive t (A : Type) | (y : nat) : bool -> Type := +% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true +% | K2 : t A y false +% +% is written +% +% (parameter "A" explicit {{ Type }} a\ +% inductive "t" tt (parameter "y" explicit {{ nat }} _\ +% arity {{ bool -> Type }}) +% t\ +% [ constructor "K1" +% (parameter "y" explicit {{ nat }} y\ +% (parameter "x" explicit a x\ +% (parameter "n" maximal {{ nat }} n\ +% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) +% , constructor "K2" +% (parameter "y" explicit {{ nat }} y\ +% arity {{ lp:t lp:y false }}) ]) +% +% Remark that the uniform parameters are not passed to occurrences of t, since +% they never change, while non-uniform parameters are both abstracted +% in each constructor type and passed as arguments to t. +% +% The coq.typecheck-indt-decl API can be used to fill in implicit arguments +% an infer universe constraints in the declaration above (e.g. the hidden +% argument of "=" in the arity of K1). +% +% Note: when and inductive type declaration is passed as an argument to an +% Elpi command non uniform parameters must be separated from the uniform ones +% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since +% version 1.4, in Coq this separator is optional, but not in Elpi). + +% Context declaration (used as an argument to Elpi commands) +kind context-decl type. +% Eg. (x : T) or (x := B), body is optional, type may be a variable +type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. +type context-end context-decl. + +typeabbrev field-attributes (list field-attribute). + +macro @global! :- get-option "coq:locality" "global". +macro @local! :- get-option "coq:locality" "local". + + +% Coq terms are not visible at synterp time, they are always holes + +kind term type. + +% -- Parsing time APIs +% ---------------------------------------------------- + +% [id] is a name that matters, we piggy back on Elpi's strings. +% Note: [name] is a name that does not matter. +typeabbrev id string. + + +% Name of a module /*E*/ +typeabbrev modpath (ctype "modpath"). + + +% Name of a module type /*E*/ +typeabbrev modtypath (ctype "modtypath"). + + +% [coq.locate-module ModName ModPath] locates a module. It's a fatal error +% if ModName cannot be located. *E* +external pred coq.locate-module i:id, o:modpath. + +% [coq.locate-module-type ModName ModPath] locates a module. It's a fatal +% error if ModName cannot be located. *E* +external pred coq.locate-module-type i:id, o:modtypath. + + +kind located type. +type loc-modpath modpath -> located. +type loc-modtypath modtypath -> located. + + +% [coq.locate-all Name Located] finds all possible meanings of a string. +% Does not fail. +external pred coq.locate-all i:id, o:list located. + +% Coq Module inline directive +kind coq.inline type. +type coq.inline.no coq.inline. % Coq's [no inline] (aka !) +type coq.inline.default coq.inline. % The default, can be omitted +type coq.inline.at int -> coq.inline. % Coq's [inline at ] + +external pred coq.env.begin-module-functor % Starts a functor *E* + i:id, % The name of the functor + i:option modtypath, % Its module type + i:list (pair id modtypath). % Parameters of the functor + + +pred coq.env.begin-module i:id, i:option modtypath. +coq.env.begin-module Name MP :- + coq.env.begin-module-functor Name MP []. + + +% [coq.env.end-module ModPath] end the current module that becomes known as +% ModPath *E* +external pred coq.env.end-module o:modpath. + +external pred coq.env.begin-module-type-functor % Starts a module type functor *E* + i:id, % The name of the functor + i:list (pair id modtypath). % The parameters of the functor + + +pred coq.env.begin-module-type i:id. +coq.env.begin-module-type Name :- + coq.env.begin-module-type-functor Name []. + + +% [coq.env.end-module-type ModTyPath] end the current module type that +% becomes known as ModPath *E* +external pred coq.env.end-module-type o:modtypath. + +external pred coq.env.apply-module-functor % Applies a functor *E* + i:id, % The name of the new module + i:option modtypath, % Its module type + i:modpath, % The functor being applied + i:list modpath, % Its arguments + i:coq.inline, % Arguments inlining + o:modpath. % The modpath of the new module + +external pred coq.env.apply-module-type-functor % Applies a type functor *E* + i:id, % The name of the new module type + i:modtypath, % The functor + i:list modpath, % Its arguments + i:coq.inline, % Arguments inlining + o:modtypath. % The modtypath of the new module type + +% [coq.env.include-module ModPath Inline] is like the vernacular Include, +% Inline can be omitted *E* +external pred coq.env.include-module i:modpath, i:coq.inline. + +% [coq.env.include-module-type ModTyPath Inline] is like the vernacular +% Include Type, Inline can be omitted *E* +external pred coq.env.include-module-type i:modtypath, i:coq.inline. + +% [coq.env.import-module ModPath] is like the vernacular Import *E* +external pred coq.env.import-module i:modpath. + +% [coq.env.export-module ModPath] is like the vernacular Export *E* +external pred coq.env.export-module i:modpath. + +% [coq.env.begin-section Name] starts a section named Name *E* +external pred coq.env.begin-section i:id. + +% [coq.env.end-section] end the current section *E* +external pred coq.env.end-section . + +% [coq.modpath->path MP FullPath] extract the full kernel name, each +% component is a separate list item +external pred coq.modpath->path i:modpath, o:list string. + +% [coq.modtypath->path MTP FullPath] extract the full kernel name, each +% component is a separate list item +external pred coq.modtypath->path i:modtypath, o:list string. + +% [coq.modpath->library MP LibraryPath] extract the enclosing module which +% can be Required +external pred coq.modpath->library i:modpath, o:modpath. + +% [coq.modtypath->library MTP LibraryPath] extract the enclosing module +% which can be Required +external pred coq.modtypath->library i:modtypath, o:modpath. + +% [coq.env.current-path Path] lists the current module path +external pred coq.env.current-path o:list string. + +% [coq.env.current-section-path Path] lists the current section path +external pred coq.env.current-section-path o:list string. + +% clauses +% +% A clause like +% :name "foo" :before "bar" foo X Y :- bar X Z, baz Z Y +% is represented as +% clause "foo" (before "bar") (pi x y z\ foo x y :- bar x z, baz z y) +% that is exactly what one would load in the context using =>. +% +% The name and the grafting specification can be left unspecified. +kind clause type. +type clause id -> grafting -> prop -> clause. + +% Specify if the clause has to be grafted before or after a named clause +kind grafting type. +type before id -> grafting. +type after id -> grafting. + +% Specify to which module the clause should be attached to +kind scope type. +type execution-site scope. % The module inside which the Elpi program is run +type current scope. % The module being defined (see begin/end-module) +type library scope. % The outermost module (carrying the file name) + + +% see coq.elpi.accumulate-clauses +pred coq.elpi.accumulate i:scope, i:id, i:clause. +coq.elpi.accumulate S N C :- coq.elpi.accumulate-clauses S N [C]. + + +% [coq.elpi.accumulate-clauses Scope DbName Clauses] +% Declare that, once the program is over, the given clauses has to be +% added to the given db (see Elpi Db). +% Clauses usually belong to Coq modules: the Scope argument lets one +% select which module: +% - execution site (default) is the module in which the pogram is +% invoked +% - current is the module currently being constructed (see +% begin/end-module) +% - library is the current file (the module that is named after the file) +% The clauses are visible as soon as the enclosing module is +% Imported. +% Clauses cannot be accumulated inside functors. +% Supported attributes: +% - @local! (default: false, discard at the end of section or module) +% - @global! (default: false, always active, only if Scope is +% execution-site, discouraged) +external pred coq.elpi.accumulate-clauses i:scope, i:id, i:list clause. + +% Action executed during the parsing phase (aka synterp) +kind synterp-action type. +type begin-module id -> synterp-action. +type begin-module-type id -> synterp-action. +type begin-section id -> synterp-action. +type end-module modpath -> synterp-action. +type end-module-type modtypath -> synterp-action. +type end-section synterp-action. +type apply-module-functor id -> synterp-action. +type apply-module-type-functor id -> synterp-action. +type include-module modpath -> synterp-action. +type include-module-type modtypath -> synterp-action. +type import-module modpath -> synterp-action. +type export-module modpath -> synterp-action. + +% [coq.synterp-actions A] Get the list of actions performed during the +% parsing phase (aka synterp) up to now. +external pred coq.synterp-actions o:list synterp-action. + + + + diff --git a/coq-builtin.elpi b/coq-builtin.elpi index 2e9741208..1f36482a8 100644 --- a/coq-builtin.elpi +++ b/coq-builtin.elpi @@ -14,13 +14,12 @@ %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% -%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-arg-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % This section contains the low level data types linking Coq and elpi. -% In particular: -% - the data type for terms and the evar_map entries (a sequent) -% - the entry points for commands and tactics (main and solve) +% In particular the entry points for commands + % %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % Entry points @@ -35,16 +34,13 @@ % holds. The encoding of terms is described below. % See also the coq.parse-attributes utility. pred main i:list argument. +pred main-interp i:list argument, i:any. +pred main-synterp i:list argument, o:any. pred usage. pred attributes o:list attribute. -% Entry point for tactics. Eg. "elpi mytactic foo 3 (f x)." becomes -% solve -% Where [str "foo", int 3, trm (app[f,x])] is part of . -% The encoding of goals is described below. -% msolve is for tactics that operate on multiple goals (called via all: ). -pred solve i:goal, o:list sealed-goal. -pred msolve i:list sealed-goal, o:list sealed-goal. +% see coq-lib.elpi for coq.parse-attributes generating the options below +type get-option string -> A -> prop. % The data type of arguments (for commands or tactics) kind argument type. @@ -52,9 +48,6 @@ type int int -> argument. % Eg. 1 -2. type str string -> argument. % Eg. x "y" z.w. or any Coq keyword/symbol type trm term -> argument. % Eg. (t). -% Extra arguments for tactics -type tac ltac1-tactic -> argument. - % Extra arguments for commands. [Definition], [Axiom], [Record] and [Context] % take precedence over the [str] argument above (when not "quoted"). % @@ -70,6 +63,111 @@ type upoly-const-decl id -> option term -> arity -> upoly-decl -> argument. % Eg. Context A (b : A). type ctx-decl context-decl -> argument. +% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +kind indt-decl type. +kind indc-decl type. +kind record-decl type. + +% An arity is written, in Coq syntax, as: +% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U +% This syntax is used, for example, in the type of an inductive type or +% in the type of constructors. We call the abstractions on the left of ":" +% "parameters" while we call the type following the ":" (proper) arity. + +% Note: in some contexts, like the type of an inductive type constructor, +% Coq makes no distinction between these two writings +% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... +% while Elpi is a bit more restrictive, since it understands user directives +% such as the implicit status of an arguments (eg, using {} instead of () around +% the binder), only on parameters. +% Moreover parameters carry the name given by the user as an "id", while binders +% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see +% also the HOAS of terms). A user command can hence only use the names of +% parameters, and not the names of "forall" quantified variables in the arity. +% +% See also the arity->term predicate in coq-lib.elpi + +type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. +type arity term -> arity. + +type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. +type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive +type record id -> term -> id -> record-decl -> indt-decl. + +type constructor id -> arity -> indc-decl. + +type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. +type end-record record-decl. + +% Example. +% Remark that A is a regular parameter; y is a non-uniform parameter and t +% also features an index of type bool. +% +% Inductive t (A : Type) | (y : nat) : bool -> Type := +% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true +% | K2 : t A y false +% +% is written +% +% (parameter "A" explicit {{ Type }} a\ +% inductive "t" tt (parameter "y" explicit {{ nat }} _\ +% arity {{ bool -> Type }}) +% t\ +% [ constructor "K1" +% (parameter "y" explicit {{ nat }} y\ +% (parameter "x" explicit a x\ +% (parameter "n" maximal {{ nat }} n\ +% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) +% , constructor "K2" +% (parameter "y" explicit {{ nat }} y\ +% arity {{ lp:t lp:y false }}) ]) +% +% Remark that the uniform parameters are not passed to occurrences of t, since +% they never change, while non-uniform parameters are both abstracted +% in each constructor type and passed as arguments to t. +% +% The coq.typecheck-indt-decl API can be used to fill in implicit arguments +% an infer universe constraints in the declaration above (e.g. the hidden +% argument of "=" in the arity of K1). +% +% Note: when and inductive type declaration is passed as an argument to an +% Elpi command non uniform parameters must be separated from the uniform ones +% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since +% version 1.4, in Coq this separator is optional, but not in Elpi). + +% Context declaration (used as an argument to Elpi commands) +kind context-decl type. +% Eg. (x : T) or (x := B), body is optional, type may be a variable +type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. +type context-end context-decl. + +typeabbrev field-attributes (list field-attribute). + +macro @global! :- get-option "coq:locality" "global". +macro @local! :- get-option "coq:locality" "local". + + + +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% This section contains the low level data types linking Coq and elpi. +% In particular the data type for terms and the evar_map entries (a sequent) +% and the entry points for tactics + +% Entry point for tactics. Eg. "elpi mytactic foo 3 (f x)." becomes +% solve +% Where [str "foo", int 3, trm (app[f,x])] is part of . +% The encoding of goals is described below. +% msolve is for tactics that operate on multiple goals (called via all: ). +pred solve i:goal, o:list sealed-goal. +pred msolve i:list sealed-goal, o:list sealed-goal. + +% Extra arguments for tactics +type tac ltac1-tactic -> argument. + % %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % Coq's terms % @@ -354,9 +452,6 @@ typeabbrev opaque? bool. macro @opaque! :- tt. macro @transparent! :- ff. %%%%%%% Attributes to be passed to APIs as in @local! => coq.something %%%%%%%% -macro @global! :- get-option "coq:locality" "global". -macro @local! :- get-option "coq:locality" "local". - macro @primitive! :- get-option "coq:primitive" tt. % primitive records macro @reversible! :- get-option "coq:reversible" tt. % coercions macro @no-tc! :- get-option "coq:no_tc" tt. % skip typeclass inference @@ -394,92 +489,9 @@ macro @redflags! F :- get-option "coq:redflags" F. % for whd & co % both arguments are strings eg "8.12.0" "use foo instead" macro @deprecated! Since Msg :- get-option "coq:deprecated" (pr Since Msg). - macro @ltacfail! N :- get-option "ltac:fail" N. - -% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -kind indt-decl type. -kind indc-decl type. -kind record-decl type. - -% An arity is written, in Coq syntax, as: -% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U -% This syntax is used, for example, in the type of an inductive type or -% in the type of constructors. We call the abstractions on the left of ":" -% "parameters" while we call the type following the ":" (proper) arity. - -% Note: in some contexts, like the type of an inductive type constructor, -% Coq makes no distinction between these two writings -% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... -% while Elpi is a bit more restrictive, since it understands user directives -% such as the implicit status of an arguments (eg, using {} instead of () around -% the binder), only on parameters. -% Moreover parameters carry the name given by the user as an "id", while binders -% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see -% also the HOAS of terms). A user command can hence only use the names of -% parameters, and not the names of "forall" quantified variables in the arity. -% -% See also the arity->term predicate in coq-lib.elpi - -type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. -type arity term -> arity. - -type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. -type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive -type record id -> term -> id -> record-decl -> indt-decl. - -type constructor id -> arity -> indc-decl. - -type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. -type end-record record-decl. - -% Example. -% Remark that A is a regular parameter; y is a non-uniform parameter and t -% also features an index of type bool. -% -% Inductive t (A : Type) | (y : nat) : bool -> Type := -% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true -% | K2 : t A y false -% -% is written -% -% (parameter "A" explicit {{ Type }} a\ -% inductive "t" tt (parameter "y" explicit {{ nat }} _\ -% arity {{ bool -> Type }}) -% t\ -% [ constructor "K1" -% (parameter "y" explicit {{ nat }} y\ -% (parameter "x" explicit a x\ -% (parameter "n" maximal {{ nat }} n\ -% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) -% , constructor "K2" -% (parameter "y" explicit {{ nat }} y\ -% arity {{ lp:t lp:y false }}) ]) -% -% Remark that the uniform parameters are not passed to occurrences of t, since -% they never change, while non-uniform parameters are both abstracted -% in each constructor type and passed as arguments to t. -% -% The coq.typecheck-indt-decl API can be used to fill in implicit arguments -% an infer universe constraints in the declaration above (e.g. the hidden -% argument of "=" in the arity of K1). -% -% Note: when and inductive type declaration is passed as an argument to an -% Elpi command non uniform parameters must be separated from the uniform ones -% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since -% version 1.4, in Coq this separator is optional, but not in Elpi). - -% Context declaration (used as an argument to Elpi commands) -kind context-decl type. -% Eg. (x : T) or (x := B), body is optional, type may be a variable -type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. -type context-end context-decl. - -typeabbrev field-attributes (list field-attribute). - % retrocompatibility macro for Coq v8.10 macro @coercion! :- [coercion reversible]. @@ -577,10 +589,6 @@ typeabbrev modpath (ctype "modpath"). typeabbrev modtypath (ctype "modtypath"). -% -- Environment: read ------------------------------------------------ - -% Note: The type [term] is defined in coq-HOAS.elpi - % Result of coq.locate-all kind located type. type loc-gref gref -> located. @@ -601,6 +609,10 @@ external pred coq.locate-all i:id, o:list located. % It's a fatal error if Name cannot be located. external pred coq.locate i:id, o:gref. +% -- Environment: read ------------------------------------------------ + +% Note: The type [term] is defined in coq-HOAS.elpi + % [coq.env.typeof GR Ty] reads the type Ty of a global reference. % Supported attributes: % - @uinstance! I (default: fresh instance I) @@ -813,24 +825,23 @@ type coq.inline.at int -> coq.inline. % Coq's [inline at ] % [coq.env.fresh-global-id X X] can be used to check if X is taken external pred coq.env.fresh-global-id i:id, o:id. -external pred coq.env.begin-module-functor % Starts a functor *E* +external pred coq.env.begin-module-functor % Starts a functor. bla bla i:id, % The name of the functor - i:option modtypath, % Its module type - i:list (pair id modtypath). % Parameters of the functor + i:option modtypath, % Its module type (optional) + i:list (pair id modtypath). % Parameters of the functor (optional) pred coq.env.begin-module i:id, i:option modtypath. -coq.env.begin-module Name MP :- - coq.env.begin-module-functor Name MP []. +coq.env.begin-module Name MP :- coq.env.begin-module-functor Name MP []. % [coq.env.end-module ModPath] end the current module that becomes known as -% ModPath *E* +% ModPath *E*. bla bla external pred coq.env.end-module o:modpath. -external pred coq.env.begin-module-type-functor % Starts a module type functor *E* +external pred coq.env.begin-module-type-functor % Starts a module type functor *E*. bla bla i:id, % The name of the functor - i:list (pair id modtypath). % The parameters of the functor + i:list (pair id modtypath). % The parameters of the functor (optional) pred coq.env.begin-module-type i:id. @@ -839,30 +850,30 @@ coq.env.begin-module-type Name :- % [coq.env.end-module-type ModTyPath] end the current module type that -% becomes known as ModPath *E* +% becomes known as ModPath *E*. bla bla external pred coq.env.end-module-type o:modtypath. -external pred coq.env.apply-module-functor % Applies a functor *E* +external pred coq.env.apply-module-functor % Applies a functor *E*. bla bla i:id, % The name of the new module - i:option modtypath, % Its module type - i:modpath, % The functor being applied - i:list modpath, % Its arguments - i:coq.inline, % Arguments inlining + i:option modtypath, % Its module type (optional) + i:modpath, % The functor being applied (optional) + i:list modpath, % Its arguments (optional) + i:coq.inline, % Arguments inlining (optional) o:modpath. % The modpath of the new module -external pred coq.env.apply-module-type-functor % Applies a type functor *E* +external pred coq.env.apply-module-type-functor % Applies a type functor *E*. bla bla i:id, % The name of the new module type - i:modtypath, % The functor - i:list modpath, % Its arguments - i:coq.inline, % Arguments inlining + i:modtypath, % The functor (optional) + i:list modpath, % Its arguments (optional) + i:coq.inline, % Arguments inlining (optional) o:modtypath. % The modtypath of the new module type -% [coq.env.include-module ModPath Inline] is like the vernacular Include, -% Inline can be omitted *E* +% [coq.env.include-module ModPath Inline (optional)] is like the vernacular +% Include, Inline can be omitted *E*. bla bla external pred coq.env.include-module i:modpath, i:coq.inline. -% [coq.env.include-module-type ModTyPath Inline] is like the vernacular -% Include Type, Inline can be omitted *E* +% [coq.env.include-module-type ModTyPath Inline (optional)] is like the +% vernacular Include Type, Inline can be omitted *E*. bla bla external pred coq.env.include-module-type i:modtypath, i:coq.inline. % [coq.env.import-module ModPath] is like the vernacular Import *E* @@ -1374,11 +1385,6 @@ type node list attribute -> attribute-value. kind attribute type. type attribute string -> attribute-value -> attribute. - -% see coq-lib.elpi for coq.parse-attributes generating the options below -type get-option string -> A -> prop. - - % -- Coq's pretyper --------------------------------------------------- % [coq.sigma.print] Prints Coq's Evarmap and the mapping to/from Elpi's @@ -1762,6 +1768,32 @@ external pred coq.elpi.add-predicate i:string, i:string, i:string, % PredName to Args external pred coq.elpi.predicate i:string, i:list any, o:prop. +% -- Synterp ---------------------------------------------------------- + +% Action executed during the parsing phase (aka synterp) +kind synterp-action type. +type begin-module id -> synterp-action. +type begin-module-type id -> synterp-action. +type begin-section id -> synterp-action. +type end-module modpath -> synterp-action. +type end-module-type modtypath -> synterp-action. +type end-section synterp-action. +type apply-module-functor id -> synterp-action. +type apply-module-type-functor id -> synterp-action. +type include-module modpath -> synterp-action. +type include-module-type modtypath -> synterp-action. +type import-module modpath -> synterp-action. +type export-module modpath -> synterp-action. + +% [coq.next-synterp-action A] Get the next action performed during parsing +% (aka synterp), that is also the next action to be performed during +% execution (aka interp). See also coq.replay-synterp-action +external pred coq.next-synterp-action o:synterp-action. + +% [coq.replay-all-missing-synterp-actions] Execute all actions needed in +% order to match the actions performed at parsing time (aka synterp) +external pred coq.replay-all-missing-synterp-actions . + % -- Utils ------------------------------------------------------------ kind coq.gref.set type. @@ -1802,6 +1834,14 @@ external pred coq.gref.set.elements i:coq.gref.set, o:list gref. % [coq.gref.set.cardinal M N] N is the number of elements of M external pred coq.gref.set.cardinal i:coq.gref.set, o:int. +% [coq.gref.set.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.gref.set.filter i:coq.gref.set, i:gref -> prop, + o:coq.gref.set. + +% [coq.gref.set.map M F M1] Map M w.r.t. the predicate F +external pred coq.gref.set.map i:coq.gref.set, i:gref -> gref -> prop, + o:coq.gref.set. + % CAVEAT: the type parameter of coq.gref.map must be a closed term kind coq.gref.map type -> type. @@ -1827,6 +1867,14 @@ external pred coq.gref.map.find i:gref, i:coq.gref.map A, o:A. external pred coq.gref.map.bindings i:coq.gref.map A, o:list (pair gref A). +% [coq.gref.map.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.gref.map.filter i:coq.gref.map A, i:gref -> A -> prop, + o:coq.gref.map A. + +% [coq.gref.map.map M F M1] Map M w.r.t. the predicate F +external pred coq.gref.map.map i:coq.gref.map A, + i:gref -> A -> B -> prop, o:coq.gref.map B. + kind coq.univ.set type. % [coq.univ.set.empty A] The empty set @@ -1865,6 +1913,14 @@ external pred coq.univ.set.elements i:coq.univ.set, o:list univ. % [coq.univ.set.cardinal M N] N is the number of elements of M external pred coq.univ.set.cardinal i:coq.univ.set, o:int. +% [coq.univ.set.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.univ.set.filter i:coq.univ.set, i:univ -> prop, + o:coq.univ.set. + +% [coq.univ.set.map M F M1] Map M w.r.t. the predicate F +external pred coq.univ.set.map i:coq.univ.set, i:univ -> univ -> prop, + o:coq.univ.set. + % CAVEAT: the type parameter of coq.univ.map must be a closed term kind coq.univ.map type -> type. @@ -1890,6 +1946,14 @@ external pred coq.univ.map.find i:univ, i:coq.univ.map A, o:A. external pred coq.univ.map.bindings i:coq.univ.map A, o:list (pair univ A). +% [coq.univ.map.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.univ.map.filter i:coq.univ.map A, i:univ -> A -> prop, + o:coq.univ.map A. + +% [coq.univ.map.map M F M1] Map M w.r.t. the predicate F +external pred coq.univ.map.map i:coq.univ.map A, + i:univ -> A -> B -> prop, o:coq.univ.map B. + kind coq.univ.variable.set type. % [coq.univ.variable.set.empty A] The empty set @@ -1940,6 +2004,16 @@ external pred coq.univ.variable.set.elements i:coq.univ.variable.set, external pred coq.univ.variable.set.cardinal i:coq.univ.variable.set, o:int. +% [coq.univ.variable.set.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.univ.variable.set.filter i:coq.univ.variable.set, + i:univ.variable -> prop, + o:coq.univ.variable.set. + +% [coq.univ.variable.set.map M F M1] Map M w.r.t. the predicate F +external pred coq.univ.variable.set.map i:coq.univ.variable.set, + i:univ.variable -> univ.variable -> prop, + o:coq.univ.variable.set. + % CAVEAT: the type parameter of coq.univ.variable.map must be a closed % term @@ -1971,6 +2045,16 @@ external pred coq.univ.variable.map.find i:univ.variable, external pred coq.univ.variable.map.bindings i:coq.univ.variable.map A, o:list (pair univ.variable A). +% [coq.univ.variable.map.filter M F M1] Filter M w.r.t. the predicate F +external pred coq.univ.variable.map.filter i:coq.univ.variable.map A, + i:univ.variable -> A -> prop, + o:coq.univ.variable.map A. + +% [coq.univ.variable.map.map M F M1] Map M w.r.t. the predicate F +external pred coq.univ.variable.map.map i:coq.univ.variable.map A, + i:univ.variable -> A -> B -> prop, + o:coq.univ.variable.map B. + % Coq box types for pretty printing: % - Vertical block: each break leads to a new line % - Horizontal block: no line breaking diff --git a/coq-elpi.opam b/coq-elpi.opam index 5fb2abfba..e79779681 100644 --- a/coq-elpi.opam +++ b/coq-elpi.opam @@ -14,7 +14,7 @@ build: [ [ make "build" "COQBIN=%{bin}%/" "ELPIDIR=%{prefix}%/lib/elpi" "OCAML install: [ make "install" "COQBIN=%{bin}%/" "ELPIDIR=%{prefix}%/lib/elpi" ] depends: [ "stdlib-shims" - "elpi" {>= "1.16.5" & < "1.18.0~"} + "elpi" {>= "1.18.1" & < "1.19.0~"} "coq" {>= "8.18" & < "8.19~" } "dot-merlin-reader" {with-dev} "ocaml-lsp-server" {with-dev} diff --git a/elpi-builtin.elpi b/elpi-builtin.elpi index 0c1dd7f4d..32349ec6b 100644 --- a/elpi-builtin.elpi +++ b/elpi-builtin.elpi @@ -68,13 +68,15 @@ stop :- halt. % -- Evaluation -- +pred (is) o:A, i:A. + +X is Y :- calc Y X. + % [calc Expr Out] unifies Out with the value of Expr. It can be used in % tandem with spilling, eg [f {calc (N + 1)}] external pred calc i:A, o:A. -pred (is) o:A, i:A. - -X is Y :- calc Y X. +% --- Operators --- type (-) A -> A -> A. @@ -83,7 +85,6 @@ type (i-) int -> int -> int. type (r-) float -> float -> float. type (+) int -> int -> int. - type (+) float -> float -> float. type (i+) int -> int -> int. @@ -91,7 +92,6 @@ type (i+) int -> int -> int. type (r+) float -> float -> float. type (*) int -> int -> int. - type (*) float -> float -> float. type (/) float -> float -> float. @@ -103,7 +103,6 @@ type (div) int -> int -> int. type (^) string -> string -> string. type (~) int -> int. - type (~) float -> float. type (i~) int -> int. @@ -111,7 +110,6 @@ type (i~) int -> int. type (r~) float -> float. type abs int -> int. - type abs float -> float. type iabs int -> int. @@ -119,11 +117,9 @@ type iabs int -> int. type rabs float -> float. type max int -> int -> int. - type max float -> float -> float. type min int -> int -> int. - type min float -> float -> float. type sqrt float -> float. @@ -741,6 +737,16 @@ external pred std.string.map.find i:string, i:std.string.map A, o:A. external pred std.string.map.bindings i:std.string.map A, o:list (pair string A). +% [std.string.map.filter M F M1] Filter M w.r.t. the predicate F +external pred std.string.map.filter i:std.string.map A, + i:string -> A -> prop, + o:std.string.map A. + +% [std.string.map.map M F M1] Map M w.r.t. the predicate F +external pred std.string.map.map i:std.string.map A, + i:string -> A -> B -> prop, + o:std.string.map B. + % CAVEAT: the type parameter of std.int.map must be a closed term kind std.int.map type -> type. @@ -764,6 +770,14 @@ external pred std.int.map.find i:int, i:std.int.map A, o:A. % [std.int.map.bindings M L] L is M transformed into an associative list external pred std.int.map.bindings i:std.int.map A, o:list (pair int A). +% [std.int.map.filter M F M1] Filter M w.r.t. the predicate F +external pred std.int.map.filter i:std.int.map A, i:int -> A -> prop, + o:std.int.map A. + +% [std.int.map.map M F M1] Map M w.r.t. the predicate F +external pred std.int.map.map i:std.int.map A, i:int -> A -> B -> prop, + o:std.int.map B. + % CAVEAT: the type parameter of std.loc.map must be a closed term kind std.loc.map type -> type. @@ -787,6 +801,14 @@ external pred std.loc.map.find i:loc, i:std.loc.map A, o:A. % [std.loc.map.bindings M L] L is M transformed into an associative list external pred std.loc.map.bindings i:std.loc.map A, o:list (pair loc A). +% [std.loc.map.filter M F M1] Filter M w.r.t. the predicate F +external pred std.loc.map.filter i:std.loc.map A, i:loc -> A -> prop, + o:std.loc.map A. + +% [std.loc.map.map M F M1] Map M w.r.t. the predicate F +external pred std.loc.map.map i:std.loc.map A, i:loc -> A -> B -> prop, + o:std.loc.map B. + kind std.string.set type. % [std.string.set.empty A] The empty set @@ -827,6 +849,15 @@ external pred std.string.set.elements i:std.string.set, o:list string. % [std.string.set.cardinal M N] N is the number of elements of M external pred std.string.set.cardinal i:std.string.set, o:int. +% [std.string.set.filter M F M1] Filter M w.r.t. the predicate F +external pred std.string.set.filter i:std.string.set, i:string -> prop, + o:std.string.set. + +% [std.string.set.map M F M1] Map M w.r.t. the predicate F +external pred std.string.set.map i:std.string.set, + i:string -> string -> prop, + o:std.string.set. + kind std.int.set type. % [std.int.set.empty A] The empty set @@ -865,6 +896,14 @@ external pred std.int.set.elements i:std.int.set, o:list int. % [std.int.set.cardinal M N] N is the number of elements of M external pred std.int.set.cardinal i:std.int.set, o:int. +% [std.int.set.filter M F M1] Filter M w.r.t. the predicate F +external pred std.int.set.filter i:std.int.set, i:int -> prop, + o:std.int.set. + +% [std.int.set.map M F M1] Map M w.r.t. the predicate F +external pred std.int.set.map i:std.int.set, i:int -> int -> prop, + o:std.int.set. + kind std.loc.set type. % [std.loc.set.empty A] The empty set @@ -903,6 +942,14 @@ external pred std.loc.set.elements i:std.loc.set, o:list loc. % [std.loc.set.cardinal M N] N is the number of elements of M external pred std.loc.set.cardinal i:std.loc.set, o:int. +% [std.loc.set.filter M F M1] Filter M w.r.t. the predicate F +external pred std.loc.set.filter i:std.loc.set, i:loc -> prop, + o:std.loc.set. + +% [std.loc.set.map M F M1] Map M w.r.t. the predicate F +external pred std.loc.set.map i:std.loc.set, i:loc -> loc -> prop, + o:std.loc.set. + #line 0 "builtin_map.elpi" kind std.map type -> type -> type. type std.map std.map.private.map K V -> (K -> K -> cmp -> prop) -> std.map K V. diff --git a/elpi/coq-HOAS.elpi b/elpi/coq-HOAS.elpi index 599ff4b16..bd7d8ebb0 100644 --- a/elpi/coq-HOAS.elpi +++ b/elpi/coq-HOAS.elpi @@ -3,25 +3,8 @@ %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % This section contains the low level data types linking Coq and elpi. -% In particular: -% - the data type for terms and the evar_map entries (a sequent) -% - the entry points for commands and tactics (main and solve) - -% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% -% Entry points -% -% Command and tactic invocation -% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -% Entry point for commands. Eg. "#[att=true] Elpi mycommand foo 3 (f x)." becomes -% main [str "foo", int 3, trm (app[f,x])] -% in a context where -% attributes [attribute "att" (leaf "true")] -% holds. The encoding of terms is described below. -% See also the coq.parse-attributes utility. -pred main i:list argument. -pred usage. -pred attributes o:list attribute. +% In particular the data type for terms and the evar_map entries (a sequent) +% and the entry points for tactics % Entry point for tactics. Eg. "elpi mytactic foo 3 (f x)." becomes % solve @@ -31,30 +14,9 @@ pred attributes o:list attribute. pred solve i:goal, o:list sealed-goal. pred msolve i:list sealed-goal, o:list sealed-goal. -% The data type of arguments (for commands or tactics) -kind argument type. -type int int -> argument. % Eg. 1 -2. -type str string -> argument. % Eg. x "y" z.w. or any Coq keyword/symbol -type trm term -> argument. % Eg. (t). - % Extra arguments for tactics type tac ltac1-tactic -> argument. -% Extra arguments for commands. [Definition], [Axiom], [Record] and [Context] -% take precedence over the [str] argument above (when not "quoted"). -% -% Eg. Record or Inductive -type indt-decl indt-decl -> argument. -% Eg. #[universes(polymorphic,...)] Record or Inductive -type upoly-indt-decl indt-decl -> upoly-decl -> argument. -type upoly-indt-decl indt-decl -> upoly-decl-cumul -> argument. -% Eg. Definition or Axiom (when the body is none) -type const-decl id -> option term -> arity -> argument. -% Eg. #[universes(polymorphic,...)] Definition or Axiom -type upoly-const-decl id -> option term -> arity -> upoly-decl -> argument. -% Eg. Context A (b : A). -type ctx-decl context-decl -> argument. - % %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % Coq's terms % @@ -339,9 +301,6 @@ typeabbrev opaque? bool. macro @opaque! :- tt. macro @transparent! :- ff. %%%%%%% Attributes to be passed to APIs as in @local! => coq.something %%%%%%%% -macro @global! :- get-option "coq:locality" "global". -macro @local! :- get-option "coq:locality" "local". - macro @primitive! :- get-option "coq:primitive" tt. % primitive records macro @reversible! :- get-option "coq:reversible" tt. % coercions macro @no-tc! :- get-option "coq:no_tc" tt. % skip typeclass inference @@ -379,91 +338,8 @@ macro @redflags! F :- get-option "coq:redflags" F. % for whd & co % both arguments are strings eg "8.12.0" "use foo instead" macro @deprecated! Since Msg :- get-option "coq:deprecated" (pr Since Msg). - macro @ltacfail! N :- get-option "ltac:fail" N. - -% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% - -kind indt-decl type. -kind indc-decl type. -kind record-decl type. - -% An arity is written, in Coq syntax, as: -% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U -% This syntax is used, for example, in the type of an inductive type or -% in the type of constructors. We call the abstractions on the left of ":" -% "parameters" while we call the type following the ":" (proper) arity. - -% Note: in some contexts, like the type of an inductive type constructor, -% Coq makes no distinction between these two writings -% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... -% while Elpi is a bit more restrictive, since it understands user directives -% such as the implicit status of an arguments (eg, using {} instead of () around -% the binder), only on parameters. -% Moreover parameters carry the name given by the user as an "id", while binders -% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see -% also the HOAS of terms). A user command can hence only use the names of -% parameters, and not the names of "forall" quantified variables in the arity. -% -% See also the arity->term predicate in coq-lib.elpi - -type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. -type arity term -> arity. - -type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. -type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive -type record id -> term -> id -> record-decl -> indt-decl. - -type constructor id -> arity -> indc-decl. - -type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. -type end-record record-decl. - -% Example. -% Remark that A is a regular parameter; y is a non-uniform parameter and t -% also features an index of type bool. -% -% Inductive t (A : Type) | (y : nat) : bool -> Type := -% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true -% | K2 : t A y false -% -% is written -% -% (parameter "A" explicit {{ Type }} a\ -% inductive "t" tt (parameter "y" explicit {{ nat }} _\ -% arity {{ bool -> Type }}) -% t\ -% [ constructor "K1" -% (parameter "y" explicit {{ nat }} y\ -% (parameter "x" explicit a x\ -% (parameter "n" maximal {{ nat }} n\ -% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) -% , constructor "K2" -% (parameter "y" explicit {{ nat }} y\ -% arity {{ lp:t lp:y false }}) ]) -% -% Remark that the uniform parameters are not passed to occurrences of t, since -% they never change, while non-uniform parameters are both abstracted -% in each constructor type and passed as arguments to t. -% -% The coq.typecheck-indt-decl API can be used to fill in implicit arguments -% an infer universe constraints in the declaration above (e.g. the hidden -% argument of "=" in the arity of K1). -% -% Note: when and inductive type declaration is passed as an argument to an -% Elpi command non uniform parameters must be separated from the uniform ones -% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since -% version 1.4, in Coq this separator is optional, but not in Elpi). - -% Context declaration (used as an argument to Elpi commands) -kind context-decl type. -% Eg. (x : T) or (x := B), body is optional, type may be a variable -type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. -type context-end context-decl. - -typeabbrev field-attributes (list field-attribute). - % retrocompatibility macro for Coq v8.10 macro @coercion! :- [coercion reversible]. diff --git a/elpi/coq-arg-HOAS.elpi b/elpi/coq-arg-HOAS.elpi new file mode 100644 index 000000000..6e11a7fe4 --- /dev/null +++ b/elpi/coq-arg-HOAS.elpi @@ -0,0 +1,133 @@ +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-arg-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% This section contains the low level data types linking Coq and elpi. +% In particular the entry points for commands + + +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +% Entry points +% +% Command and tactic invocation +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% Entry point for commands. Eg. "#[att=true] Elpi mycommand foo 3 (f x)." becomes +% main [str "foo", int 3, trm (app[f,x])] +% in a context where +% attributes [attribute "att" (leaf "true")] +% holds. The encoding of terms is described below. +% See also the coq.parse-attributes utility. +pred main i:list argument. +pred main-interp i:list argument, i:any. +pred main-synterp i:list argument, o:any. +pred usage. +pred attributes o:list attribute. + +% see coq-lib.elpi for coq.parse-attributes generating the options below +type get-option string -> A -> prop. + +% The data type of arguments (for commands or tactics) +kind argument type. +type int int -> argument. % Eg. 1 -2. +type str string -> argument. % Eg. x "y" z.w. or any Coq keyword/symbol +type trm term -> argument. % Eg. (t). + +% Extra arguments for commands. [Definition], [Axiom], [Record] and [Context] +% take precedence over the [str] argument above (when not "quoted"). +% +% Eg. Record or Inductive +type indt-decl indt-decl -> argument. +% Eg. #[universes(polymorphic,...)] Record or Inductive +type upoly-indt-decl indt-decl -> upoly-decl -> argument. +type upoly-indt-decl indt-decl -> upoly-decl-cumul -> argument. +% Eg. Definition or Axiom (when the body is none) +type const-decl id -> option term -> arity -> argument. +% Eg. #[universes(polymorphic,...)] Definition or Axiom +type upoly-const-decl id -> option term -> arity -> upoly-decl -> argument. +% Eg. Context A (b : A). +type ctx-decl context-decl -> argument. + +% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +kind indt-decl type. +kind indc-decl type. +kind record-decl type. + +% An arity is written, in Coq syntax, as: +% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U +% This syntax is used, for example, in the type of an inductive type or +% in the type of constructors. We call the abstractions on the left of ":" +% "parameters" while we call the type following the ":" (proper) arity. + +% Note: in some contexts, like the type of an inductive type constructor, +% Coq makes no distinction between these two writings +% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... +% while Elpi is a bit more restrictive, since it understands user directives +% such as the implicit status of an arguments (eg, using {} instead of () around +% the binder), only on parameters. +% Moreover parameters carry the name given by the user as an "id", while binders +% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see +% also the HOAS of terms). A user command can hence only use the names of +% parameters, and not the names of "forall" quantified variables in the arity. +% +% See also the arity->term predicate in coq-lib.elpi + +type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. +type arity term -> arity. + +type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. +type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive +type record id -> term -> id -> record-decl -> indt-decl. + +type constructor id -> arity -> indc-decl. + +type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. +type end-record record-decl. + +% Example. +% Remark that A is a regular parameter; y is a non-uniform parameter and t +% also features an index of type bool. +% +% Inductive t (A : Type) | (y : nat) : bool -> Type := +% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true +% | K2 : t A y false +% +% is written +% +% (parameter "A" explicit {{ Type }} a\ +% inductive "t" tt (parameter "y" explicit {{ nat }} _\ +% arity {{ bool -> Type }}) +% t\ +% [ constructor "K1" +% (parameter "y" explicit {{ nat }} y\ +% (parameter "x" explicit a x\ +% (parameter "n" maximal {{ nat }} n\ +% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) +% , constructor "K2" +% (parameter "y" explicit {{ nat }} y\ +% arity {{ lp:t lp:y false }}) ]) +% +% Remark that the uniform parameters are not passed to occurrences of t, since +% they never change, while non-uniform parameters are both abstracted +% in each constructor type and passed as arguments to t. +% +% The coq.typecheck-indt-decl API can be used to fill in implicit arguments +% an infer universe constraints in the declaration above (e.g. the hidden +% argument of "=" in the arity of K1). +% +% Note: when and inductive type declaration is passed as an argument to an +% Elpi command non uniform parameters must be separated from the uniform ones +% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since +% version 1.4, in Coq this separator is optional, but not in Elpi). + +% Context declaration (used as an argument to Elpi commands) +kind context-decl type. +% Eg. (x : T) or (x := B), body is optional, type may be a variable +type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. +type context-end context-decl. + +typeabbrev field-attributes (list field-attribute). + +macro @global! :- get-option "coq:locality" "global". +macro @local! :- get-option "coq:locality" "local". diff --git a/elpi/coq-lib.elpi b/elpi/coq-lib.elpi index 84962ebb2..9005218da 100644 --- a/elpi/coq-lib.elpi +++ b/elpi/coq-lib.elpi @@ -749,3 +749,19 @@ parse-attributes.aux [attribute S (leaf-loc V)|AS] Prefix CLS :- !, if (Diag = error Msg) (coq.error Msg) true, CLS = [get-option PS V|R], parse-attributes.aux AS Prefix R. + +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +pred coq.replay-synterp-action i:synterp-action. +coq.replay-synterp-action (begin-module ID) :- coq.env.begin-module ID _. +coq.replay-synterp-action (end-module _) :- coq.env.end-module _. +coq.replay-synterp-action (begin-module-type ID) :- coq.env.begin-module-type ID. +coq.replay-synterp-action (end-module-type _) :- coq.env.end-module-type _. +coq.replay-synterp-action (apply-module-functor ID) :- coq.env.apply-module-functor ID _ _ _ _ _. +coq.replay-synterp-action (apply-module-type-functor ID) :- coq.env.apply-module-type-functor ID _ _ _ _. +coq.replay-synterp-action (include-module MP) :- coq.env.include-module MP _. +coq.replay-synterp-action (include-module-type MP) :- coq.env.include-module-type MP _. +coq.replay-synterp-action (import-module MP) :- coq.env.import-module MP. +coq.replay-synterp-action (export-module MP) :- coq.env.export-module MP. +coq.replay-synterp-action (begin-section ID) :- coq.env.begin-section ID. +coq.replay-synterp-action (end-section) :- coq.env.end-section. diff --git a/elpi/elpi-command-template-synterp.elpi b/elpi/elpi-command-template-synterp.elpi new file mode 100644 index 000000000..6cef3c560 --- /dev/null +++ b/elpi/elpi-command-template-synterp.elpi @@ -0,0 +1,3 @@ +/* Loaded when Elpi Command has a code accumulated at #[synterp] time */ +/* license: GNU Lesser General Public License Version 2.1 or later */ +/* ------------------------------------------------------------------------- */ diff --git a/examples/tutorial_coq_elpi_command.v b/examples/tutorial_coq_elpi_command.v index 085f4e522..f6cd76f7c 100644 --- a/examples/tutorial_coq_elpi_command.v +++ b/examples/tutorial_coq_elpi_command.v @@ -727,6 +727,12 @@ Warnings can be reported using the :builtin:`coq.warning` which lets you pick a name and category. In turn these can be used to disable or make fatal your warnings (as any other Coq warning). +===================== +Parsing and Execution +===================== + +TODO + This is really the end, unless you want to learn more about writing `tactics `_ in Elpi, in that case look at that tutorial ;-) diff --git a/src/coq_elpi_HOAS.ml b/src/coq_elpi_HOAS.ml index 25a59b156..fa034230c 100644 --- a/src/coq_elpi_HOAS.ml +++ b/src/coq_elpi_HOAS.ml @@ -89,8 +89,8 @@ module UM = F.Map(struct let pp fmt x = Format.fprintf fmt "%a" Pp.pp_with (Univ.Universe.pr UnivNames.pr_with_global_universes x) end) -let um = S.declare ~name:"coq-elpi:evar-univ-map" - ~pp:UM.pp ~init:(fun () -> UM.empty) ~start:(fun x -> x) +let um = S.declare_component ~name:"coq-elpi:evar-univ-map" ~descriptor:interp_state + ~pp:UM.pp ~init:(fun () -> UM.empty) ~start:(fun x -> x) () let constraint_leq u1 u2 = @@ -793,11 +793,12 @@ let from_env_keep_univ_of_sigma ~env0 ~env sigma0 = let uctx = UState.demote_global_univs env uctx in from_env_sigma env (Evd.from_ctx uctx) - let init () = from_env (Global.env ()) + let init () = + from_env (Global.env ()) let engine : coq_engine S.component = - S.declare ~name:"coq-elpi:evmap-constraint-type" - ~pp:pp_coq_engine ~init ~start:(fun _ -> init()) + S.declare_component ~name:"coq-elpi:evmap-constraint-type" ~descriptor:interp_state + ~pp:pp_coq_engine ~init ~start:(fun _ -> init()) () end let () = pre_engine := Some CoqEngine_HOAS.engine @@ -1474,7 +1475,7 @@ let prepend_removals l = let removals, rest = List.partition (function RmEvar _ -> true | _ -> false) l in removals @ rest -let () = E.set_extra_goals_postprocessing (fun l state -> +let () = E.set_extra_goals_postprocessing ~descriptor:interp_hoas (fun l state -> generate_actual_goals state (prepend_removals (cancel_opposites Evar.Set.empty (removals_of Evar.Set.empty l) l))) @@ -1675,8 +1676,8 @@ module UIM = F.Map(struct let pp fmt x = Format.fprintf fmt "%a" Pp.pp_with (Univ.Instance.pr UnivNames.pr_with_global_universes x) end) -let uim = S.declare ~name:"coq-elpi:evar-univ-instance-map" - ~pp:UIM.pp ~init:(fun () -> UIM.empty) ~start:(fun x -> x) +let uim = S.declare_component ~name:"coq-elpi:evar-univ-instance-map" ~descriptor:interp_state + ~pp:UIM.pp ~init:(fun () -> UIM.empty) ~start:(fun x -> x) () let in_coq_poly_gref ~depth ~origin ~failsafe s t i = let open API.Conversion in diff --git a/src/coq_elpi_arg_HOAS.ml b/src/coq_elpi_arg_HOAS.ml index 4f122d67b..4d86d03f9 100644 --- a/src/coq_elpi_arg_HOAS.ml +++ b/src/coq_elpi_arg_HOAS.ml @@ -10,6 +10,8 @@ open Coq_elpi_utils open Coq_elpi_HOAS open Names +type phase = Interp | Synterp | Both + let push_name x = function | Names.Name.Name id -> let decl = Context.Named.Declaration.LocalAssum (Context.make_annot id Sorts.Relevant, Constr.mkProp) in @@ -25,7 +27,7 @@ let push_inductive_in_intern_env intern_env name params arity user_impls = let env = Global.env () in let sigma = Evd.from_env env in let sigma, ty = Pretyping.understand_tcc env sigma ~expected_type:Pretyping.IsType (Coq_elpi_utils.mk_gforall arity params) in - Constrintern.compute_internalization_env env sigma ~impls:intern_env + ty, Constrintern.compute_internalization_env env sigma ~impls:intern_env Constrintern.Inductive [name] [ty] [user_impls] let intern_tactic_constr = Ltac_plugin.Tacintern.intern_constr @@ -34,12 +36,22 @@ let intern_global_constr { Ltac_plugin.Tacintern.genv = env } ~intern_env t = let sigma = Evd.from_env env in Constrintern.intern_gen Pretyping.WithoutTypeConstraint env sigma ~impls:intern_env ~pattern_mode:false ~ltacvars:Constrintern.empty_ltac_sign t -let intern_global_constr_ty { Ltac_plugin.Tacintern.genv = env } ~intern_env t = +let intern_global_constr_ty { Ltac_plugin.Tacintern.genv = env } ~intern_env ?(expty=Pretyping.IsType) t = let sigma = Evd.from_env env in - Constrintern.intern_gen Pretyping.IsType env sigma ~impls:intern_env ~pattern_mode:false ~ltacvars:Constrintern.empty_ltac_sign t + Constrintern.intern_gen expty env sigma ~impls:intern_env ~pattern_mode:false ~ltacvars:Constrintern.empty_ltac_sign t let intern_global_context { Ltac_plugin.Tacintern.genv = env } ~intern_env ctx = Constrintern.intern_context env ~bound_univs:UnivNames.empty_binders intern_env ctx +let intern_global_context_synterp (ctx : Constrexpr.local_binder_expr list) : Glob_term.glob_decl list = + let open Constrexpr in + let intern_one h = + match h with + | CLocalAssum(nl,Default bk,_) -> nl |> List.map (fun n -> n.CAst.v,bk,None,mkGHole) + | CLocalAssum(nl,Generalized(bk,_),_) -> nl |> List.map (fun n -> n.CAst.v,bk,None,mkGHole) + | CLocalDef (n,_,None) -> [n.CAst.v,Glob_term.Explicit,None,mkGHole] + | CLocalDef (n,_,Some _) -> [n.CAst.v,Glob_term.Explicit,Some mkGHole,mkGHole] + | CLocalPattern _ -> nYI "irrefutable pattern in synterp" in + List.concat_map intern_one ctx |> List.rev module Cmd = struct @@ -246,6 +258,27 @@ let univpoly_of ~poly ~cumulative = let intern_record_decl glob_sign (it : raw_record_decl) = glob_sign, it +let raw_record_decl_to_glob_synterp ({ name; sort; parameters; constructor; fields; univpoly } : raw_record_decl_elpi) : glob_record_decl_elpi = + let name, space = sep_last_qualid name in + let params = intern_global_context_synterp parameters in + let params = List.rev params in + let arity = mkGHole in + let fields = + List.fold_left (fun acc -> function + | Vernacexpr.AssumExpr ({ CAst.v = name } as fn,bl,x), { Vernacexpr.rf_coercion = inst; rf_priority = pr; rf_notation = nots; rf_canonical = canon } -> + if nots <> [] then Coq_elpi_utils.nYI "notation in record fields"; + if pr <> None then Coq_elpi_utils.nYI "priority in record fields"; + let atts = { Coq_elpi_HOAS.is_canonical = canon; is_coercion = if inst = Vernacexpr.AddCoercion then Reversible else Off; name } in + let x = if bl = [] then x else Constrexpr_ops.mkCProdN bl x in + let entry = intern_global_context_synterp [Constrexpr.CLocalAssum ([fn],Constrexpr.Default Glob_term.Explicit,x)] in + let x = match entry with + | [_,_,_,x] -> x + | _ -> assert false in + (x, atts) :: acc + | Vernacexpr.DefExpr _, _ -> Coq_elpi_utils.nYI "DefExpr") + [] fields in + { name = (space, Names.Id.of_string name); arity; params; constructorname = constructor; fields = List.rev fields; univpoly } + let raw_record_decl_to_glob glob_sign ({ name; sort; parameters; constructor; fields; univpoly } : raw_record_decl_elpi) : glob_record_decl_elpi = let name, space = sep_last_qualid name in let sort = match sort with @@ -272,6 +305,21 @@ let raw_record_decl_to_glob glob_sign ({ name; sort; parameters; constructor; fi (glob_sign_params,intern_env,[]) fields in { name = (space, Names.Id.of_string name); arity; params; constructorname = constructor; fields = List.rev fields; univpoly } +let raw_indt_decl_to_glob_synterp ({ finiteness; name; parameters; non_uniform_parameters; arity; constructors; univpoly } : raw_indt_decl_elpi) : glob_indt_decl_elpi = + let name, space = sep_last_qualid name in + let name = Names.Id.of_string name in + let params = intern_global_context_synterp parameters in + let nuparams_given, nuparams = + match non_uniform_parameters with + | Some l -> true, l + | None -> false, [] in + let nuparams = intern_global_context_synterp nuparams in + let params = List.rev params in + let nuparams = List.rev nuparams in + let arity = mkGHole in + let constructors = List.map (fun (id,ty) -> id.CAst.v, mkGHole) constructors in + { finiteness; name = (space, name); arity; params; nuparams; nuparams_given; constructors; univpoly } + let raw_indt_decl_to_glob glob_sign ({ finiteness; name; parameters; non_uniform_parameters; arity; constructors; univpoly } : raw_indt_decl_elpi) : glob_indt_decl_elpi = let name, space = sep_last_qualid name in let name = Names.Id.of_string name in @@ -293,15 +341,19 @@ let raw_indt_decl_to_glob glob_sign ({ finiteness; name; parameters; non_uniform let glob_sign_params = push_glob_ctx allparams glob_sign in let arity = intern_global_constr_ty ~intern_env glob_sign_params indexes in let glob_sign_params_self = push_name glob_sign_params (Names.Name name) in - let intern_env = push_inductive_in_intern_env intern_env name allparams arity user_impls in + let indty, intern_env = push_inductive_in_intern_env intern_env name allparams arity user_impls in let constructors = List.map (fun (id,ty) -> id.CAst.v, - intern_global_constr_ty glob_sign_params_self ~intern_env ty) constructors in + intern_global_constr_ty ~expty:(Pretyping.OfType indty) glob_sign_params_self ~intern_env ty) constructors in { finiteness; name = (space, name); arity; params; nuparams; nuparams_given; constructors; univpoly } let intern_indt_decl glob_sign (it : raw_indt_decl) = glob_sign, it let expr_hole = CAst.make @@ Constrexpr.CHole(None,Namegen.IntroAnonymous) +let raw_context_decl_to_glob_synterp fields = + let fields = intern_global_context_synterp fields in + List.rev fields + let raw_context_decl_to_glob glob_sign fields = let _intern_env, fields = intern_global_context ~intern_env:Constrintern.empty_internalization_env glob_sign fields in List.rev fields @@ -349,6 +401,19 @@ let raw_constant_decl_to_constr ~depth coq_ctx state { name; typ = (bl,typ); bod | _ -> assert false +let raw_constant_decl_to_glob_synterp ({ name; atts; udecl; typ = (params,typ); body } : raw_constant_decl) state = + let params = intern_global_context_synterp params in + let params = List.rev params in + let typ = mkGHole in + let body = Option.map (fun _ -> mkGHole) body in + let poly = + let open Attributes in + parse polymorphic atts in + let udecl = + if poly then NonCumulative (([],true),(Univ.Constraints.empty,true)) + else NotUniversePolymorphic in + state, { name = raw_decl_name_to_glob name; params; typ; udecl; body } + let raw_constant_decl_to_glob glob_sign ({ name; atts; udecl; typ = (params,typ); body } : raw_constant_decl) state = let intern_env, params = intern_global_context glob_sign ~intern_env:Constrintern.empty_internalization_env params in let glob_sign_params = push_glob_ctx params glob_sign in @@ -528,6 +593,39 @@ let mk_indt_decl state univpoly r = assert(gls=[]); state, E.mkApp uideclc r [up] +let rec do_params_synterp ~depth params k state = + match params with + | [] -> k state + | (name,imp,ob,src) :: params -> + if ob <> None then Coq_elpi_utils.nYI "defined parameters in a record/inductive declaration"; + let src = E.mkDiscard in + let state, tgt = do_params_synterp ~depth params k state in + let state, imp = in_elpi_imp ~depth state imp in + state, in_elpi_parameter name ~imp src tgt + + +let rec do_fields_synterp ~depth fields state = + match fields with + | [] -> state, in_elpi_indtdecl_endrecord () + | (f,({ name; is_coercion; is_canonical } as att)) :: fields -> + let f = E.mkDiscard in + let state, fields = do_fields_synterp ~depth fields state in + in_elpi_indtdecl_field ~depth state att f fields + +let do_record_synterp ~depth ~name ~constructorname arity fields state = + let space, record_name = name in + let qrecord_name = Id.of_string_soft @@ String.concat "." (space @ [Id.to_string record_name]) in + let arity = E.mkDiscard in + let state, fields = do_fields_synterp ~depth fields state in + let constructor = match constructorname with + | None -> Name.Name (Id.of_string ("Build_" ^ Id.to_string record_name)) + | Some x -> Name.Name x in + state, in_elpi_indtdecl_record (Name.Name qrecord_name) arity constructor fields + +let grecord2lp_synterp ~depth state { Cmd.name; arity; params; constructorname; fields; univpoly } = + let state, r = do_params_synterp ~depth params (do_record_synterp ~depth ~name ~constructorname arity fields) state in + mk_indt_decl state univpoly r + let grecord2lp ~depth state { Cmd.name; arity; params; constructorname; fields; univpoly } = let open Coq_elpi_glob_quotation in let state, r = do_params params (do_record ~name ~constructorname arity fields) ~depth state in @@ -573,7 +671,21 @@ let drop_unit f ~depth state = let state, x, () = f ~depth state in state, x - +let ginductive2lp_synterp ~depth state { Cmd.finiteness; name; arity; params; nuparams; nuparams_given; constructors; univpoly } = + let space, indt_name = name in + let do_constructor ~depth state (name, ty) = + let state, ty = do_params_synterp nuparams (fun state -> state, in_elpi_arity E.mkDiscard) ~depth state in + state, in_elpi_indtdecl_constructor (Name.Name name) ty + in + let do_inductive_synterp ~depth state = + let qindt_name = Id.of_string_soft @@ String.concat "." (space @ [Id.to_string indt_name]) in + let state, arity = do_params_synterp nuparams (fun state -> state, in_elpi_arity E.mkDiscard) ~depth state in + let state, constructors = Coq_elpi_utils.list_map_acc (do_constructor ~depth ) state constructors in + state, in_elpi_indtdecl_inductive state finiteness (Name.Name qindt_name) arity constructors + in + let state, r = do_params_synterp params (do_inductive_synterp ~depth) ~depth state in + mk_indt_decl state univpoly r + let ginductive2lp ~depth state { Cmd.finiteness; name; arity; params; nuparams; nuparams_given; constructors; univpoly } = let open Coq_elpi_glob_quotation in let space, indt_name = name in @@ -606,6 +718,18 @@ let decl_name2lp name = let cdeclc = E.Constants.declare_global_symbol "const-decl" let ucdeclc = E.Constants.declare_global_symbol "upoly-const-decl" +let gdecl2lp_synterp ~depth state { Cmd.name; params; typ : _; body; udecl } = + let state, typ = do_params_synterp ~depth params (fun state -> state, in_elpi_arity E.mkDiscard) state in + let body = Option.map (fun _ -> E.mkDiscard) body in + let name = decl_name2lp name in + let state, body, gls = in_option ~depth state body in + match udecl with + | NotUniversePolymorphic -> state, E.mkApp cdeclc name [body;typ], gls + | Cumulative _ -> assert false + | NonCumulative ud -> + let state, ud, gls1 = universe_decl.API.Conversion.embed ~depth state ud in + state, E.mkApp ucdeclc name [body;typ;ud], gls @ gls1 + let cdecl2lp ~depth state { Cmd.name; params; typ; body; udecl } = let open Coq_elpi_glob_quotation in let state, typ = do_params params (do_arity typ) ~depth state in @@ -622,6 +746,17 @@ let cdecl2lp ~depth state { Cmd.name; params; typ; body; udecl } = let ctxitemc = E.Constants.declare_global_symbol "context-item" let ctxendc = E.Constants.declare_global_symbol "context-end" +let rec do_context_glob_synterp fields ~depth state = + match fields with + | [] -> state, E.mkGlobal ctxendc + | (name,bk,bo,ty) :: fields -> + let ty = E.mkDiscard in + let bo = Option.map (fun _ -> E.mkDiscard) bo in + let state, fields = do_context_glob_synterp fields ~depth state in + let state, bo, _ = in_option ~depth state bo in + let state, imp = in_elpi_imp ~depth state bk in + state, E.mkApp ctxitemc (in_elpi_id name) [imp;ty;bo;E.mkLam fields] + let rec do_context_glob fields ~depth state = match fields with | [] -> state, E.mkGlobal ctxendc @@ -786,6 +921,31 @@ let handle_template_polymorphism = function | Some false -> Some false | Some true -> err Pp.(str "#[universes(template)] is not supported") +let in_elpi_cmd_synterp ~depth ?calldepth state (x : Cmd.raw) = + let open Cmd in + match x with + | RecordDecl raw_rdecl -> + let raw_rdecl = of_coq_record_definition raw_rdecl in + let glob_rdecl = raw_record_decl_to_glob_synterp raw_rdecl in + let state, t = grecord2lp_synterp ~depth state glob_rdecl in + state, t, [] + | IndtDecl raw_indt -> + let raw_indt = of_coq_inductive_definition raw_indt in + let glob_indt = raw_indt_decl_to_glob_synterp raw_indt in + let state, t = ginductive2lp_synterp ~depth state glob_indt in + state, t, [] + | ConstantDecl raw_cdecl -> + let state, glob_cdecl = raw_constant_decl_to_glob_synterp raw_cdecl state in + gdecl2lp_synterp ~depth state glob_cdecl + | Context raw_ctx -> + let glob_ctx = raw_context_decl_to_glob_synterp raw_ctx in + let state, t = do_context_glob_synterp glob_ctx ~depth state in + state, E.mkApp ctxc t [], [] + | Int x -> in_elpi_int_arg ~depth state x + | String x -> in_elpi_string_arg ~depth state x + | Term raw_term -> + state, E.mkApp trmc E.mkDiscard [], [] + let in_elpi_cmd ~depth ?calldepth coq_ctx state ~raw (x : Cmd.top) = let open Cmd in let hyps = [] in diff --git a/src/coq_elpi_arg_HOAS.mli b/src/coq_elpi_arg_HOAS.mli index 98540e482..c946e60fe 100644 --- a/src/coq_elpi_arg_HOAS.mli +++ b/src/coq_elpi_arg_HOAS.mli @@ -5,6 +5,8 @@ open Elpi.API.RawData open Coq_elpi_utils +type phase = Interp | Synterp | Both + module Cmd : sig type raw_term = Constrexpr.constr_expr @@ -122,6 +124,11 @@ val in_elpi_cmd : raw:bool -> Cmd.top -> Elpi.API.State.t * term * Elpi.API.Conversion.extra_goals +val in_elpi_cmd_synterp : + depth:int -> ?calldepth:int -> + Elpi.API.State.t -> + Cmd.raw -> + Elpi.API.State.t * term * Elpi.API.Conversion.extra_goals type coq_arg = Cint of int | Cstr of string | Ctrm of EConstr.t | CLtac1 of Geninterp.Val.t diff --git a/src/coq_elpi_arg_syntax.mlg b/src/coq_elpi_arg_syntax.mlg index acc8d20d0..c0b3f9af9 100644 --- a/src/coq_elpi_arg_syntax.mlg +++ b/src/coq_elpi_arg_syntax.mlg @@ -127,7 +127,34 @@ let skip_attribute : (Str.regexp list option * Str.regexp list option) Attribute | Attributes.VernacFlagLeaf (Attributes.FlagString rex) -> Str.regexp rex :: o2l old | _ -> CErrors.user_err ?loc (Pp.str "Syntax error, use #[only=\"rex\"]")] -let raw_args_attribute = Attributes.(qualify_attribute "arguments" (bool_attribute ~name:"raw")) +let synterp_attribute : EA.phase option Attributes.attribute = + let open EA in + Attributes.attribute_of_list + ["phase", + (fun ?loc old -> function + | Attributes.VernacFlagLeaf (Attributes.FlagString "parsing") -> Synterp + | Attributes.VernacFlagLeaf (Attributes.FlagString "execution") -> Interp + | Attributes.VernacFlagLeaf (Attributes.FlagString "both") -> Both + | _ -> CErrors.user_err ?loc (Pp.str "Syntax error, use #[phase=\"parsing\"] or #[phase=\"execution\"] or #[phase=\"both\"]")) + ;"synterp", + (fun ?loc old -> function + | Attributes.VernacFlagEmpty -> Synterp + | _ -> CErrors.user_err ?loc (Pp.str "Syntax error, use #[synterp]")) + ;"interp", + (fun ?loc old -> function + | Attributes.VernacFlagEmpty -> Interp + | _ -> CErrors.user_err ?loc (Pp.str "Syntax error, use #[interp]")) + ;"both", + (fun ?loc old -> function + | Attributes.VernacFlagEmpty -> Both + | _ -> CErrors.user_err ?loc (Pp.str "Syntax error, use #[both]")) + + ] + +let skip_and_synterp_attributes = Attributes.Notations.(skip_attribute ++ synterp_attribute) + +let raw_args_attributes = + Attributes.(qualify_attribute "arguments" (bool_attribute ~name:"raw")) let validate_attributes a flags = let extra, raw_args = Attributes.parse_with_extra a flags in diff --git a/src/coq_elpi_builtins.ml b/src/coq_elpi_builtins.ml index bfb59afa6..07ff313f8 100644 --- a/src/coq_elpi_builtins.ml +++ b/src/coq_elpi_builtins.ml @@ -112,15 +112,10 @@ let pr_econstr_env options env sigma t = if options.hoas_holes = Some Heuristic then aux () expr else expr in Ppconstr.pr_constr_expr_n env sigma options.pplevel expr) -let tactic_mode = State.declare ~name:"coq-elpi:tactic-mode" +let tactic_mode : bool State.component = State.declare_component ~name:"coq-elpi:tactic-mode" ~descriptor:interp_state ~pp:(fun fmt x -> Format.fprintf fmt "%b" x) ~init:(fun () -> false) - ~start:(fun x -> x) -let invocation_site_loc = State.declare ~name:"coq-elpi:invocation-site-loc" - ~pp:(fun fmt x -> Format.fprintf fmt "%a" API.Ast.Loc.pp x) - ~init:(fun () -> API.Ast.Loc.initial "(should-not-happen)") - ~start:(fun x -> x) - + ~start:(fun x -> x) () let abstract__grab_global_env_keep_sigma api thunk = (); (fun state -> let state, result, gls = thunk state in Coq_elpi_HOAS.grab_global_env state, result, gls) @@ -146,29 +141,6 @@ let grab_global_env_drop_sigma api thunk = (); (fun state -> let state, result, gls = thunk state in Coq_elpi_HOAS.grab_global_env_drop_sigma state, result, gls) -let err_if_contains_alg_univ ~depth t = - let global_univs = UGraph.domain (Environ.universes (Global.env ())) in - let is_global u = - match Univ.Universe.level u with - | None -> true - | Some l -> Univ.Level.Set.mem l global_univs in - let rec aux ~depth acc t = - match E.look ~depth t with - | E.CData c when isuniv c -> - let u = univout c in - if is_global u then acc - else - begin match Univ.Universe.level u with - | None -> - err Pp.(strbrk "The hypothetical clause contains terms of type univ which are not global, you should abstract them out or replace them by global ones: " ++ - Univ.Universe.pr UnivNames.pr_with_global_universes u) - | _ -> Univ.Universe.Set.add u acc - end - | x -> Coq_elpi_utils.fold_elpi_term aux acc ~depth x - in - let univs = aux ~depth Univ.Universe.Set.empty t in - univs -;; let bool = B.bool let int = B.int @@ -207,16 +179,15 @@ let constr2lp_closed ~depth hyps constraints state t = let constr2lp_closed_ground ~depth hyps constraints state t = constr2lp_closed_ground ~depth hyps constraints state t -let clauses_for_later = - State.declare ~name:"coq-elpi:clauses_for_later" +let clauses_for_later_interp : _ State.component = + State.declare_component ~name:"coq-elpi:clauses_for_later" ~descriptor:interp_state ~init:(fun () -> []) ~start:(fun x -> x) ~pp:(fun fmt l -> List.iter (fun (dbname, code,vars,scope) -> Format.fprintf fmt "db:%s code:%a scope:%a\n" (String.concat "." dbname) - Elpi.API.Pp.Ast.program code Coq_elpi_utils.pp_scope scope) l) -;; + Elpi.API.Pp.Ast.program code Coq_elpi_utils.pp_scope scope) l) () let term = { CConv.ty = Conv.TyName "term"; @@ -269,7 +240,6 @@ let term_skeleton = { embed = (fun ~depth _ _ _ _ -> assert false); } -let prop = { B.any with Conv.ty = Conv.TyName "prop" } let sealed_goal = { Conv.ty = Conv.TyName "sealed-goal"; pp_doc = (fun fmt () -> ()); @@ -297,14 +267,7 @@ let tactic_arg : (Coq_elpi_arg_HOAS.coq_arg, Coq_elpi_HOAS.full Coq_elpi_HOAS.co readback = Coq_elpi_arg_HOAS.in_coq_arg; } -let id = { B.string with - API.Conversion.ty = Conv.TyName "id"; - pp_doc = (fun fmt () -> - Format.fprintf fmt "%% [id] is a name that matters, we piggy back on Elpi's strings.@\n"; - Format.fprintf fmt "%% Note: [name] is a name that does not matter.@\n"; - Format.fprintf fmt "typeabbrev id string.@\n@\n") -} - +let id = Coq_elpi_builtins_synterp.id let flag name = { (B.unspec bool) with Conv.ty = Conv.TyName name } @@ -351,32 +314,6 @@ let handle_uinst_option_for_inductive ~depth options i state = let state = update_sigma state (fun sigma -> Evd.merge_context_set UState.univ_flexible_alg sigma ctx) in uinst, state, API.Conversion.Unify (v', lp_uinst) :: extra_goals -type located = - | LocGref of Names.GlobRef.t - | LocModule of Names.ModPath.t - | LocModuleType of Names.ModPath.t - | LocAbbreviation of Globnames.abbreviation - -let located = let open Conv in let open API.AlgebraicData in declare { - ty = TyName "located"; - doc = "Result of coq.locate-all"; - pp = (fun fmt _ -> Format.fprintf fmt ""); - constructors = [ - K("loc-gref","",A(gref,N), - B (fun x -> LocGref x), - M (fun ~ok ~ko -> function LocGref x -> ok x | _ -> ko ())); - K("loc-modpath","",A(modpath,N), - B (fun x -> LocModule x), - M (fun ~ok ~ko -> function LocModule x -> ok x | _ -> ko ())); - K("loc-modtypath","",A(modtypath,N), - B (fun x -> LocModuleType x), - M (fun ~ok ~ko -> function LocModuleType x -> ok x | _ -> ko ())); - K("loc-abbreviation","",A(abbreviation,N), - B (fun x -> LocAbbreviation x), - M (fun ~ok ~ko -> function LocAbbreviation x -> ok x | _ -> ko ())); - ] -} |> CConv.(!<) - (* FIXME PARTIAL API * * Record foo A1..Am := K { f1; .. fn }. -- m params, n fields @@ -566,63 +503,68 @@ let get_instances (env: Environ.env) (sigma: Evd.evar_map) tc : type_class_insta let isnt_of_tc = get_isntances_of_tc env sigma tc in List.map (get_instance env sigma isnt_of_tc) instances_grefs -type scope = ExecutionSite | CurrentModule | Library - -let scope = let open Conv in let open API.AlgebraicData in declare { - ty = TyName "scope"; - doc = "Specify to which module the clause should be attached to"; - pp = (fun fmt _ -> Format.fprintf fmt ""); - constructors = [ - K("execution-site","The module inside which the Elpi program is run",N, - B ExecutionSite, - M (fun ~ok ~ko -> function ExecutionSite -> ok | _ -> ko ())); - K("current","The module being defined (see begin/end-module)",N, - B CurrentModule, - M (fun ~ok ~ko -> function CurrentModule -> ok | _ -> ko ())); - K("library","The outermost module (carrying the file name)",N, - B Library, - M (fun ~ok ~ko -> function Library -> ok | _ -> ko ())) - ] -} |> CConv.(!<) - -let grafting = let open Conv in let open API.AlgebraicData in declare { - ty = TyName "grafting"; - doc = "Specify if the clause has to be grafted before or after a named clause"; - pp = (fun fmt _ -> Format.fprintf fmt ""); - constructors = [ - K("before","",A(id,N), - B (fun x -> (`Before,x)), - M (fun ~ok ~ko -> function (`Before,x) -> ok x | _ -> ko ())); - K("after","",A(id,N), - B (fun x -> (`After,x)), - M (fun ~ok ~ko -> function (`After,x) -> ok x | _ -> ko ())); - ] -} |> CConv.(!<) - -let clause = let open Conv in let open API.AlgebraicData in declare { - ty = TyName "clause"; - doc = {|clauses - -A clause like - :name "foo" :before "bar" foo X Y :- bar X Z, baz Z Y -is represented as - clause "foo" (before "bar") (pi x y z\ foo x y :- bar x z, baz z y) -that is exactly what one would load in the context using =>. - -The name and the grafting specification can be left unspecified.|}; - pp = (fun fmt _ -> Format.fprintf fmt ""); - constructors = [ - K("clause","",A(B.unspec id,A(B.unspec grafting,A(prop,N))), - B (fun id graft c -> unspec2opt id, unspec2opt graft, c), - M (fun ~ok ~ko (id,graft,c) -> ok (opt2unspec id) (opt2unspec graft) c)); - ] -} |> CConv.(!<) - -let set_accumulate_to_db, get_accumulate_to_db = +let set_accumulate_to_db_interp, get_accumulate_to_db_interp = let f = ref (fun _ -> assert false) in (fun x -> f := x), (fun () -> !f) +let err_if_contains_alg_univ ~depth t = + let global_univs = UGraph.domain (Environ.universes (Global.env ())) in + let is_global u = + match Univ.Universe.level u with + | None -> true + | Some l -> Univ.Level.Set.mem l global_univs in + let rec aux ~depth acc t = + match E.look ~depth t with + | E.CData c when isuniv c -> + let u = univout c in + if is_global u then acc + else + begin match Univ.Universe.level u with + | None -> + err Pp.(strbrk "The hypothetical clause contains terms of type univ which are not global, you should abstract them out or replace them by global ones: " ++ + Univ.Universe.pr UnivNames.pr_with_global_universes u) + | _ -> Univ.Universe.Set.add u acc + end + | x -> Coq_elpi_utils.fold_elpi_term aux acc ~depth x + in + let univs = aux ~depth Univ.Universe.Set.empty t in + univs + +let preprocess_clause ~depth clause = + let levels_to_abstract = err_if_contains_alg_univ ~depth clause in + let levels_to_abstract_no = Univ.Universe.Set.cardinal levels_to_abstract in + let rec subst ~depth m t = + match E.look ~depth t with + | E.CData c when isuniv c -> + begin try E.mkBound (Univ.Universe.Map.find (univout c) m) + with Not_found -> t end + | E.App(c,x,xs) -> + E.mkApp c (subst ~depth m x) (List.map (subst ~depth m) xs) + | E.Cons(x,xs) -> + E.mkCons (subst ~depth m x) (subst ~depth m xs) + | E.Lam x -> + E.mkLam (subst ~depth:(depth+1) m x) + | E.Builtin(c,xs) -> + E.mkBuiltin c (List.map (subst ~depth m) xs) + | E.UnifVar _ -> assert false + | E.Const _ | E.Nil | E.CData _ -> t + in + let clause = + let rec bind d map = function + | [] -> + subst ~depth:d map + (API.Utils.move ~from:depth ~to_:(depth + levels_to_abstract_no) clause) + | l :: ls -> + E.mkApp E.Constants.pic (E.mkLam (* pi x\ *) + (bind (d+1) (Univ.Universe.Map.add l d map) ls)) [] + in + bind depth Univ.Universe.Map.empty + (Univ.Universe.Set.elements levels_to_abstract) + in + let vars = collect_term_variables ~depth clause in + vars, clause + let argument_mode = let open Conv in let open API.AlgebraicData in declare { ty = TyName "argument_mode"; doc = "Specify if a predicate argument is in input or output mode"; @@ -638,7 +580,7 @@ let argument_mode = let open Conv in let open API.AlgebraicData in declare { } |> CConv.(!<) -let set_accumulate_text_to_db, get_accumulate_text_to_db = +let set_accumulate_text_to_db_interp, get_accumulate_text_to_db_interp = let f = ref (fun _ _ _ -> assert false) in (fun x -> f := x), (fun () -> !f) @@ -977,7 +919,7 @@ let add_axiom_or_variable api id ty local options state = let uentry = UState.check_univ_decl (Evd.evar_universe_context sigma) udecl ~poly in let kind = Decls.Logical in let impargs = [] in - let loc = to_coq_loc @@ State.get invocation_site_loc state in + let loc = to_coq_loc @@ State.get Coq_elpi_builtins_synterp.invocation_site_loc state in let variable = CAst.(make ~loc @@ Id.of_string id) in if not (is_ground sigma ty) then err Pp.(str"coq.env.add-const: the type must be ground. Did you forge to call coq.typecheck-indt-decl?"); @@ -1147,15 +1089,6 @@ let goption = let open API.AlgebraicData in let open Goptions in declare { ] } |> CConv.(!<) -let module_ast_of_modpath x = - let open Libnames in let open Nametab in - qualid_of_dirpath (dirpath_of_module x) - -let module_ast_of_modtypath x = - let open Constrexpr in let open Libnames in let open Nametab in - CAst.make @@ CMident (qualid_of_path (path_of_modtype x)), - Declaremods.DefaultInline - let find_hint_db s = try Hints.searchtable_map s @@ -1347,7 +1280,6 @@ let eta_contract env sigma t = (*Printf.eprintf "------------- %s\n" Pp.(string_of_ppcmds @@ Printer.pr_econstr_env env sigma t);*) map env t - (*****************************************************************************) (*****************************************************************************) (*****************************************************************************) @@ -1360,13 +1292,8 @@ let eta_contract env sigma t = (*****************************************************************************) (*****************************************************************************) - -let coq_builtins = +let coq_header_builtins = let open API.BuiltIn in - let open Pred in - let open Notation in - let open CConv in - let pp ~depth = P.term depth in [LPCode {|% Coq terms as the object language of elpi and basic API to access Coq % license: GNU Lesser General Public License Version 2.1 or later @@ -1379,9 +1306,11 @@ let coq_builtins = % API to access Coq. |}; + LPCode Coq_elpi_builtins_arg_HOAS.code; LPCode Coq_elpi_builtins_HOAS.code; MLData Coq_elpi_HOAS.record_field_att; MLData Coq_elpi_HOAS.coercion_status; + LPCode {| %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%% builtins %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% @@ -1391,8 +1320,16 @@ let coq_builtins = % The marker *E* means *experimental*, i.e. use at your own risk, it may change % substantially or even disappear in future versions. |}; + ] - LPDoc "-- Misc ---------------------------------------------------------"; +let coq_misc_builtins = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + let pp ~depth = P.term depth in + [ + LPDoc "-- Misc ---------------------------------------------------------"; MLCode(Pred("coq.info", VariadicIn(unit_ctx, !> B.any, "Prints an info message"), @@ -1482,32 +1419,32 @@ line option|}))), let major, minor, patch = coq_version_parser version in !: version +! major +! minor +! patch)), DocAbove); - LPCode {| + ] + +let coq_locate_builtins = + let open API.BuiltIn in + let open Pred in + let open Notation in + [ LPCode {| % %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% % API for objects belonging to the logic % %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%|}; - LPDoc "-- Environment: names -----------------------------------------------"; - LPDoc {|To make the API more precise we use different data types for the names of global objects. + LPDoc "-- Environment: names -----------------------------------------------"; + LPDoc {|To make the API more precise we use different data types for the names of global objects. Note: [ctype \"bla\"] is an opaque data type and by convention it is written [@bla].|}; - - MLData constant; - MLData inductive; - MLData constructor; - MLData gref; - MLData id; - MLData modpath; - MLData modtypath; - ] @ - - [ - LPDoc "-- Environment: read ------------------------------------------------"; - LPDoc "Note: The type [term] is defined in coq-HOAS.elpi"; - - MLData located; + + MLData constant; + MLData inductive; + MLData constructor; + MLData gref; + MLData id; + MLData modpath; + MLData modtypath; + MLData Coq_elpi_builtins_synterp.located; MLCode(Pred("coq.locate-all", In(id, "Name", - Out(B.list located, "Located", + Out(B.list Coq_elpi_builtins_synterp.located, "Located", Easy {|finds all possible meanings of a string. Does not fail.|})), (fun s _ ~depth -> let qualid = Libnames.qualid_of_string s in @@ -1515,16 +1452,16 @@ Note: [ctype \"bla\"] is an opaque data type and by convention it is written [@b let add x = l := !l @ [x] in begin match locate_qualid qualid with - | Some (`Gref gr) -> add @@ LocGref gr - | Some (`Abbrev sd) -> add @@ LocAbbreviation sd + | Some (`Gref gr) -> add @@ Coq_elpi_builtins_synterp.LocGref gr + | Some (`Abbrev sd) -> add @@ Coq_elpi_builtins_synterp.LocAbbreviation sd | None -> () end; begin - try add @@ LocModule (Nametab.locate_module qualid) + try add @@ Coq_elpi_builtins_synterp.LocModule (Nametab.locate_module qualid) with Not_found -> () end; begin - try add @@ LocModuleType (Nametab.locate_modtype qualid) + try add @@ Coq_elpi_builtins_synterp.LocModuleType (Nametab.locate_modtype qualid) with Not_found -> () end; !: !l)), @@ -1541,7 +1478,17 @@ eg "lib:core.bool.true". It's a fatal error if Name cannot be located.|})), (fun s _ ~depth:_ -> !: (locate_gref s))), DocAbove); - +] + +let coq_rest_builtins = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + let pp ~depth = P.term depth in + [ + LPDoc "-- Environment: read ------------------------------------------------"; + LPDoc "Note: The type [term] is defined in coq-HOAS.elpi"; MLCode(Pred("coq.env.typeof", In(gref, "GR", @@ -1874,31 +1821,8 @@ Supported attributes: | Variable v -> raise No_clause)), DocAbove); - MLCode(Pred("coq.locate-module", - In(id, "ModName", - Out(modpath, "ModPath", - Easy "locates a module. It's a fatal error if ModName cannot be located. *E*")), - (fun s _ ~depth -> - let qualid = Libnames.qualid_of_string s in - let mp = - try Nametab.locate_module qualid - with Not_found -> - err Pp.(str "Module not found: " ++ Libnames.pr_qualid qualid) in - !:mp)), - DocAbove); - - MLCode(Pred("coq.locate-module-type", - In(id, "ModName", - Out(modtypath, "ModPath", - Easy "locates a module. It's a fatal error if ModName cannot be located. *E*")), - (fun s _ ~depth -> - let qualid = Libnames.qualid_of_string s in - let mp = - try Nametab.locate_modtype qualid - with Not_found -> - err Pp.(str "Module type not found: " ++ Libnames.pr_qualid qualid) in - !:mp)), - DocAbove); + Coq_elpi_builtins_synterp.locate_module; + Coq_elpi_builtins_synterp.locate_module_type; MLData module_item; @@ -1971,21 +1895,8 @@ Supported attributes: state, !: s, [])), DocAbove); - MLCode(Pred("coq.env.current-path", - Out(list B.string, "Path", - Read(unit_ctx, "lists the current module path")), - (fun _ ~depth _ _ state -> !: (mp2path (Safe_typing.current_modpath (Global.safe_env ()))))), - DocAbove); - - MLCode(Pred("coq.env.current-section-path", - Out(list B.string, "Path", - Read(unit_ctx, "lists the current section path")), - (fun _ ~depth _ _ state -> - let base = Lib.current_dirpath false in - let base_w_sections = Lib.current_dirpath true in - let sections = Libnames.drop_dirpath_prefix base base_w_sections in - !: (mp2path (Names.ModPath.MPfile sections)))), - DocAbove); + Coq_elpi_builtins_synterp.current_path; + Coq_elpi_builtins_synterp.current_section_path; LPCode {|% Deprecated, use coq.env.opaque? pred coq.env.const-opaque? i:constant. @@ -2085,7 +1996,7 @@ Supported attributes: let gr = Declare.declare_definition ~cinfo ~info ~opaque ~body sigma in let () = - let lid = CAst.make ~loc:(to_coq_loc @@ State.get invocation_site_loc state) (Id.of_string id) in + let lid = CAst.make ~loc:(to_coq_loc @@ State.get Coq_elpi_builtins_synterp.invocation_site_loc state) (Id.of_string id) in match scope with | Locality.Discharge -> Dumpglob.dump_definition lid true "var" | Locality.Global _ -> Dumpglob.dump_definition lid false "def" @@ -2149,7 +2060,7 @@ Supported attributes: | [ { Entries.mind_entry_typename = id; mind_entry_consnames = cids }] -> id, cids | _ -> assert false in - let lid_of id = CAst.make ~loc:(to_coq_loc @@ State.get invocation_site_loc state) id in + let lid_of id = CAst.make ~loc:(to_coq_loc @@ State.get Coq_elpi_builtins_synterp.invocation_site_loc state) id in begin match record_info with | None -> (* regular inductive *) Dumpglob.dump_definition (lid_of id) false "ind"; @@ -2212,61 +2123,36 @@ with a number, starting from 1. (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.begin-module-functor", In(id, "The name of the functor", - In(option modtypath, "Its module type", - In(list (pair id modtypath), "Parameters of the functor", - Full(unit_ctx, "Starts a functor *E*")))), - (fun name mp binders_ast ~depth _ _ -> grab_global_env "coq.env.begin-module-functor" (fun state -> - if Global.sections_are_opened () then - err Pp.(str"This elpi code cannot be run within a section since it opens a module"); - let ty = - match mp with - | None -> Declaremods.Check [] - | Some mp -> Declaremods.(Enforce (module_ast_of_modtypath mp)) in - let id = Id.of_string name in - let binders_ast = - List.map (fun (id, mty) -> - [CAst.make (Id.of_string id)], (module_ast_of_modtypath mty)) - binders_ast in - let mp = Declaremods.start_module None id binders_ast ty in - let loc = to_coq_loc @@ State.get invocation_site_loc state in - Dumpglob.dump_moddef ~loc mp "mod"; - + In(B.unspec (option modtypath), "Its module type (optional)", + In(B.unspec (list (pair id modtypath)), "Parameters of the functor (optional)", + Full(unit_ctx, "Starts a functor" ^ Coq_elpi_builtins_synterp.synterp_api_doc)))), + (fun name mp params ~depth _ _ -> grab_global_env "coq.env.begin-module-functor" (fun state -> + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_BeginModule (name,mp,params) state in state, (), []))), DocNext); LPCode {| pred coq.env.begin-module i:id, i:option modtypath. -coq.env.begin-module Name MP :- - coq.env.begin-module-functor Name MP []. +coq.env.begin-module Name MP :- coq.env.begin-module-functor Name MP []. |}; (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.end-module", Out(modpath, "ModPath", - Full(unit_ctx, "end the current module that becomes known as ModPath *E*")), + Full(unit_ctx, "end the current module that becomes known as ModPath *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc)), (fun _ ~depth _ _ -> grab_global_env_drop_sigma "coq.env.end-module" (fun state -> - let mp = Declaremods.end_module () in - state, !: mp, []))), - DocAbove); + let state, mp = Coq_elpi_builtins_synterp.SynterpAction.pop_EndModule () state in + state, ?: mp, []))), + DocAbove); (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.begin-module-type-functor", In(id, "The name of the functor", - In(list (pair id modtypath), "The parameters of the functor", - Full(unit_ctx,"Starts a module type functor *E*"))), - (fun id binders_ast ~depth _ _ -> grab_global_env "coq.env.begin-module-type-functor" (fun state -> - if Global.sections_are_opened () then - err Pp.(str"This elpi code cannot be run within a section since it opens a module"); - let id = Id.of_string id in - let binders_ast = - List.map (fun (id, mty) -> - [CAst.make (Id.of_string id)], (module_ast_of_modtypath mty)) - binders_ast in - let mp = Declaremods.start_modtype id binders_ast [] in - let loc = to_coq_loc @@ State.get invocation_site_loc state in - Dumpglob.dump_moddef ~loc mp "modtype"; - - state, (), []))), + In(B.unspec (list (pair id modtypath)), "The parameters of the functor (optional)", + Full(unit_ctx,"Starts a module type functor *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc))), + (fun id params ~depth _ _ -> grab_global_env "coq.env.begin-module-type-functor" (fun state -> + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_BeginModuleType (id,params) state in + state, (), []))), DocNext); LPCode {| @@ -2278,85 +2164,54 @@ coq.env.begin-module-type Name :- (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.end-module-type", Out(modtypath, "ModTyPath", - Full(unit_ctx, "end the current module type that becomes known as ModPath *E*")), + Full(unit_ctx, "end the current module type that becomes known as ModPath *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc)), (fun _ ~depth _ _ -> grab_global_env_drop_sigma "coq.env.end-module-type" (fun state -> - let mp = Declaremods.end_modtype () in - state, !: mp, []))), + let state, mp = Coq_elpi_builtins_synterp.SynterpAction.pop_EndModuleType () state in + state, ?: mp, []))), DocAbove); MLCode(Pred("coq.env.apply-module-functor", In(id, "The name of the new module", - In(option modtypath, "Its module type", - In(modpath, "The functor being applied", - In(list modpath, "Its arguments", - In(module_inline_default, "Arguments inlining", + In(B.unspec (option modtypath), "Its module type (optional)", + In(B.unspec modpath, "The functor being applied (optional)", + In(B.unspec (list modpath), "Its arguments (optional)", + In(B.unspec module_inline_default, "Arguments inlining (optional)", Out(modpath, "The modpath of the new module", - Full(unit_ctx, "Applies a functor *E*"))))))), + Full(unit_ctx, "Applies a functor *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc))))))), (fun name mp f arguments inline _ ~depth _ _ -> grab_global_env "coq.env.apply-module-functor" (fun state -> - if Global.sections_are_opened () then - err Pp.(str"This elpi code cannot be run within a section since it defines a module"); - let ty = - match mp with - | None -> Declaremods.Check [] - | Some mp -> Declaremods.(Enforce (module_ast_of_modtypath mp)) in - let id = Id.of_string name in - let f = CAst.make (Constrexpr.CMident (module_ast_of_modpath f)) in - let mexpr_ast_args = List.map module_ast_of_modpath arguments in - let mexpr_ast = - List.fold_left (fun hd arg -> CAst.make (Constrexpr.CMapply(hd,arg))) f mexpr_ast_args in - let mp = Declaremods.declare_module id [] ty [mexpr_ast,inline] in - let loc = to_coq_loc @@ State.get invocation_site_loc state in - Dumpglob.dump_moddef ~loc mp "mod"; - state, !: mp, []))), + let state, mp = Coq_elpi_builtins_synterp.SynterpAction.pop_ApplyModule (name,mp,f,arguments,inline) state in + state, ?: mp, []))), DocNext); MLCode(Pred("coq.env.apply-module-type-functor", In(id, "The name of the new module type", - In(modtypath, "The functor", - In(list modpath, "Its arguments", - In(module_inline_default, "Arguments inlining", + In(B.unspec modtypath, "The functor (optional)", + In(B.unspec (list modpath), "Its arguments (optional)", + In(B.unspec module_inline_default, "Arguments inlining (optional)", Out(modtypath, "The modtypath of the new module type", - Full(unit_ctx, "Applies a type functor *E*")))))), + Full(unit_ctx, "Applies a type functor *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc)))))), (fun name f arguments inline _ ~depth _ _ -> grab_global_env "coq.env.apply-module-type-functor" (fun state -> - if Global.sections_are_opened () then - err Pp.(str"This elpi code cannot be run within a section since it defines a module"); - let id = Id.of_string name in - let f,_ = module_ast_of_modtypath f in - let mexpr_ast_args = List.map module_ast_of_modpath arguments in - let mexpr_ast = - List.fold_left (fun hd arg -> CAst.make (Constrexpr.CMapply(hd,arg))) f mexpr_ast_args in - let mp = Declaremods.declare_modtype id [] [] [mexpr_ast,inline] in - let loc = to_coq_loc @@ State.get invocation_site_loc state in - Dumpglob.dump_moddef ~loc mp "modtype"; - state, !: mp, []))), + let state, mp = Coq_elpi_builtins_synterp.SynterpAction.pop_ApplyModuleType (name,f,arguments,inline) state in + state, ?: mp, []))), DocNext); (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.include-module", In(modpath, "ModPath", - In(module_inline_default, "Inline", - Full(unit_ctx, "is like the vernacular Include, Inline can be omitted *E*"))), - (fun mp inline ~depth _ _ -> grab_global_env "coq.env.include-module" (fun state -> - let fpath = match mp with - | ModPath.MPdot(mp,l) -> - Libnames.make_path (ModPath.dp mp) (Label.to_id l) - | _ -> nYI "functors" in - let tname = Constrexpr.CMident (Libnames.qualid_of_path fpath) in - let i = CAst.make tname, inline in - Declaremods.declare_include [i]; - state, (), []))), + In(B.unspec module_inline_default, "Inline (optional)", + Full(unit_ctx, "is like the vernacular Include, Inline can be omitted *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc))), + (fun mp i ~depth _ _ -> grab_global_env "coq.env.include-module" (fun state -> + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_IncludeModule (mp,i) state in + state, (), []))), DocAbove); (* XXX When Coq's API allows it, call vernacentries directly *) MLCode(Pred("coq.env.include-module-type", In(modtypath, "ModTyPath", - In(module_inline_default, "Inline", - Full(unit_ctx, "is like the vernacular Include Type, Inline can be omitted *E*"))), - (fun mp inline ~depth _ _ -> grab_global_env "coq.env.include-module-type" (fun state -> - let fpath = Nametab.path_of_modtype mp in - let tname = Constrexpr.CMident (Libnames.qualid_of_path fpath) in - let i = CAst.make tname, inline in - Declaremods.declare_include [i]; + In(B.unspec module_inline_default, "Inline (optional)", + Full(unit_ctx, "is like the vernacular Include Type, Inline can be omitted *E*" ^ Coq_elpi_builtins_synterp.synterp_api_doc))), + (fun mp i ~depth _ _ -> grab_global_env "coq.env.include-module-type" (fun state -> + let state,_ = Coq_elpi_builtins_synterp.SynterpAction.pop_IncludeModuleType (mp,i) state in state, (), []))), DocAbove); @@ -2364,7 +2219,7 @@ coq.env.begin-module-type Name :- In(modpath, "ModPath", Full(unit_ctx, "is like the vernacular Import *E*")), (fun mp ~depth _ _ -> grab_global_env "coq.env.import-module" (fun state -> - Declaremods.import_module ~export:Lib.Import Libobject.unfiltered mp; + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_ImportModule mp state in state, (), []))), DocAbove); @@ -2372,7 +2227,7 @@ coq.env.begin-module-type Name :- In(modpath, "ModPath", Full(unit_ctx, "is like the vernacular Export *E*")), (fun mp ~depth _ _ -> grab_global_env "coq.env.export-module" (fun state -> - Declaremods.import_module ~export:Lib.Export Libobject.unfiltered mp; + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_ExportModule mp state in state, (), []))), DocAbove); @@ -2396,20 +2251,14 @@ denote the same x as before.|}; In(id, "Name", Full(unit_ctx, "starts a section named Name *E*")), (fun id ~depth _ _ -> grab_global_env "coq.env.begin-section" (fun state -> - let id = Id.of_string id in - let lid = CAst.make ~loc:(to_coq_loc @@ State.get invocation_site_loc state) id in - Dumpglob.dump_definition lid true "sec"; - Lib.open_section id; + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_BeginSection id state in state, (), []))), DocAbove); MLCode(Pred("coq.env.end-section", Full(unit_ctx, "end the current section *E*"), (fun ~depth _ _ -> grab_global_env_drop_sigma "coq.env.end-section" (fun state -> - let loc = to_coq_loc @@ State.get invocation_site_loc state in - Dumpglob.dump_reference ~loc - (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec"; - Lib.close_section (); + let state, _ = Coq_elpi_builtins_synterp.SynterpAction.pop_EndSection () state in state, (), []))), DocAbove); @@ -3312,11 +3161,6 @@ is equivalent to Elpi Export TacName.|})))), MLData attribute_value; MLData attribute; - LPCode {| -% see coq-lib.elpi for coq.parse-attributes generating the options below -type get-option string -> A -> prop. -|}; - LPDoc "-- Coq's pretyper ---------------------------------------------------"; MLCode(Pred("coq.sigma.print", @@ -3984,33 +3828,10 @@ coq.id->name S N :- coq.string->name S N. (fun gr _ ~depth h c state -> !: (gr2path gr))), DocAbove); - MLCode(Pred("coq.modpath->path", - In(modpath, "MP", - Out(B.list B.string, "FullPath", - Read(unit_ctx, "extract the full kernel name, each component is a separate list item"))), - (fun mp _ ~depth h c state -> !: (mp2path mp))), - DocAbove); - - MLCode(Pred("coq.modtypath->path", - In(modtypath, "MTP", - Out(B.list B.string, "FullPath", - Read(unit_ctx, "extract the full kernel name, each component is a separate list item"))), - (fun mtyp _ ~depth h c state -> !: (mp2path mtyp))), - DocAbove); - - MLCode(Pred("coq.modpath->library", - In(modpath, "MP", - Out(modpath, "LibraryPath", - Read(unit_ctx, "extract the enclosing module which can be Required"))), - (fun mp _ ~depth h c state -> !: ModPath.(MPfile (dp mp)))), - DocAbove); - - MLCode(Pred("coq.modtypath->library", - In(modtypath, "MTP", - Out(modpath, "LibraryPath", - Read(unit_ctx, "extract the enclosing module which can be Required"))), - (fun mtyp _ ~depth h c state -> !: ModPath.(MPfile (dp mtyp)))), - DocAbove); + Coq_elpi_builtins_synterp.modpath_to_path; + Coq_elpi_builtins_synterp.modtypath_to_path; + Coq_elpi_builtins_synterp.modpath_to_library; + Coq_elpi_builtins_synterp.modtypath_to_library; MLCode(Pred("coq.term->string", CIn(failsafe_term,"T", @@ -4057,9 +3878,9 @@ Supported attributes: LPDoc "-- Access to Elpi's data --------------------------------------------"; (* Self modification *) - MLData clause; - MLData grafting; - MLData scope; + MLData Coq_elpi_builtins_synterp.clause; + MLData Coq_elpi_builtins_synterp.grafting; + MLData Coq_elpi_builtins_synterp.scope; LPCode {| % see coq.elpi.accumulate-clauses @@ -4068,9 +3889,9 @@ coq.elpi.accumulate S N C :- coq.elpi.accumulate-clauses S N [C]. |}; MLCode(Pred("coq.elpi.accumulate-clauses", - In(B.unspec scope, "Scope", + In(B.unspec Coq_elpi_builtins_synterp.scope, "Scope", In(id, "DbName", - In(B.list clause, "Clauses", + In(B.list Coq_elpi_builtins_synterp.clause, "Clauses", Full (global, {| Declare that, once the program is over, the given clauses has to be added to the given db (see Elpi Db). @@ -4088,61 +3909,11 @@ Clauses cannot be accumulated inside functors. Supported attributes: - @local! (default: false, discard at the end of section or module) - @global! (default: false, always active, only if Scope is execution-site, discouraged)|} )))), - (fun scope dbname clauses ~depth ctx _ state -> - let loc = API.Ast.Loc.initial "(elpi.add_clause)" in - let dbname = Coq_elpi_utils.string_split_on_char '.' dbname in - let clauses scope = - clauses |> CList.rev_map (fun (name,graft,clause) -> - let levels_to_abstract = err_if_contains_alg_univ ~depth clause in - let levels_to_abstract_no = Univ.Universe.Set.cardinal levels_to_abstract in - let rec subst ~depth m t = - match E.look ~depth t with - | E.CData c when isuniv c -> - begin try E.mkBound (Univ.Universe.Map.find (univout c) m) - with Not_found -> t end - | E.App(c,x,xs) -> - E.mkApp c (subst ~depth m x) (List.map (subst ~depth m) xs) - | E.Cons(x,xs) -> - E.mkCons (subst ~depth m x) (subst ~depth m xs) - | E.Lam x -> - E.mkLam (subst ~depth:(depth+1) m x) - | E.Builtin(c,xs) -> - E.mkBuiltin c (List.map (subst ~depth m) xs) - | E.UnifVar _ -> assert false - | E.Const _ | E.Nil | E.CData _ -> t - in - let clause = - let rec bind d map = function - | [] -> - subst ~depth:d map - (API.Utils.move ~from:depth ~to_:(depth + levels_to_abstract_no) clause) - | l :: ls -> - E.mkApp E.Constants.pic (E.mkLam (* pi x\ *) - (bind (d+1) (Univ.Universe.Map.add l d map) ls)) [] - in - bind depth Univ.Universe.Map.empty - (Univ.Universe.Set.elements levels_to_abstract) - in - let vars = collect_term_variables ~depth clause in - let clause = U.clause_of_term ?name ?graft ~depth loc clause in - (dbname,clause,vars,scope)) in - let local = ctx.options.local = Some true in - let super_global = ctx.options.local = Some false in - match scope with - | B.Unspec | B.Given ExecutionSite -> - let scope = if super_global then SuperGlobal else if local then Local else Regular in - State.update clauses_for_later state (fun l -> - clauses scope @ l), (), [] - | B.Given Library -> - if local then CErrors.user_err Pp.(str "coq.elpi.accumulate: library scope is incompatible with @local!"); - State.update clauses_for_later state (fun l -> - clauses Coq_elpi_utils.Global @ l), (), [] - | B.Given CurrentModule -> - let scope = if local then Local else Regular in - let f = get_accumulate_to_db () in - f (clauses scope); - state, (), [] - )), + (fun scope dbname clauses ~depth {options} _ state -> + Coq_elpi_builtins_synterp.accumulate_clauses + ~clauses_for_later:clauses_for_later_interp + ~accumulate_to_db:(get_accumulate_to_db_interp()) ~preprocess_clause + ~scope ~dbname clauses ~depth ~options state)), DocAbove); MLData argument_mode; @@ -4161,7 +3932,7 @@ Supported attributes: - @global! (default: false, always active|}))))), (fun dbname indexing predname spec ~depth ctx _ state -> let dbname = Coq_elpi_utils.string_split_on_char '.' dbname in - let f = get_accumulate_text_to_db () in + let f = get_accumulate_text_to_db_interp () in let local = ctx.options.local = Some true in let super_global = ctx.options.local = Some false in if local && super_global then CErrors.user_err Pp.(str "coq.elpi.add-predicate: @global! incompatible with @local!"); @@ -4196,6 +3967,9 @@ Supported attributes: )), DocAbove); + + ] @ Coq_elpi_builtins_synterp.SynterpAction.builtins_interp @ [ + LPDoc "-- Utils ------------------------------------------------------------"; ] @ gref_set_decl @ diff --git a/src/coq_elpi_builtins.mli b/src/coq_elpi_builtins.mli index 055204608..b0c07a3de 100644 --- a/src/coq_elpi_builtins.mli +++ b/src/coq_elpi_builtins.mli @@ -5,15 +5,17 @@ open Elpi.API open Coq_elpi_utils -val coq_builtins : BuiltIn.declaration list +val coq_header_builtins : BuiltIn.declaration list +val coq_misc_builtins : BuiltIn.declaration list +val coq_locate_builtins : BuiltIn.declaration list +val coq_rest_builtins : BuiltIn.declaration list (* Clauses to be added to elpi programs when the execution is over *) -val clauses_for_later : - (qualified_name * Ast.program * Names.Id.t list * Coq_elpi_utils.clause_scope) list State.component -val set_accumulate_to_db : - (((qualified_name * Ast.program * Names.Id.t list * Coq_elpi_utils.clause_scope) list -> unit)) -> unit -val set_accumulate_text_to_db : ((string list -> string -> Coq_elpi_utils.clause_scope -> unit)) -> unit +val clauses_for_later_interp : (qualified_name * Ast.program * Names.Id.t list * Coq_elpi_utils.clause_scope) list State.component + +val set_accumulate_to_db_interp : (((qualified_name * Ast.program * Names.Id.t list * Coq_elpi_utils.clause_scope) list -> unit)) -> unit +val set_accumulate_text_to_db_interp : ((string list -> string -> Coq_elpi_utils.clause_scope -> unit)) -> unit type attribute_data = | AttributeString of string @@ -28,7 +30,5 @@ val attribute : (string * attribute_value) Conversion.t (* In tactic mode some APIs are disabled *) val tactic_mode : bool State.component -(* To dump glob, we need a quick access to the invocation site loc *) -val invocation_site_loc : Ast.Loc.t State.component - val cache_tac_abbrev : qualified_name -> unit + diff --git a/src/coq_elpi_builtins_arg_HOAS.ml b/src/coq_elpi_builtins_arg_HOAS.ml new file mode 100644 index 000000000..bdc7098b9 --- /dev/null +++ b/src/coq_elpi_builtins_arg_HOAS.ml @@ -0,0 +1,137 @@ +(* Automatically generated from elpi/coq-arg-HOAS.elpi, don't edit *) +(* Regenerate via 'make src/coq_elpi_builtins_arg_HOAS.ml' *) +let code = {| +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%% coq-arg-HOAS %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% This section contains the low level data types linking Coq and elpi. +% In particular the entry points for commands + + +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% +% Entry points +% +% Command and tactic invocation +% %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +% Entry point for commands. Eg. "#[att=true] Elpi mycommand foo 3 (f x)." becomes +% main [str "foo", int 3, trm (app[f,x])] +% in a context where +% attributes [attribute "att" (leaf "true")] +% holds. The encoding of terms is described below. +% See also the coq.parse-attributes utility. +pred main i:list argument. +pred main-interp i:list argument, i:any. +pred main-synterp i:list argument, o:any. +pred usage. +pred attributes o:list attribute. + +% see coq-lib.elpi for coq.parse-attributes generating the options below +type get-option string -> A -> prop. + +% The data type of arguments (for commands or tactics) +kind argument type. +type int int -> argument. % Eg. 1 -2. +type str string -> argument. % Eg. x "y" z.w. or any Coq keyword/symbol +type trm term -> argument. % Eg. (t). + +% Extra arguments for commands. [Definition], [Axiom], [Record] and [Context] +% take precedence over the [str] argument above (when not "quoted"). +% +% Eg. Record or Inductive +type indt-decl indt-decl -> argument. +% Eg. #[universes(polymorphic,...)] Record or Inductive +type upoly-indt-decl indt-decl -> upoly-decl -> argument. +type upoly-indt-decl indt-decl -> upoly-decl-cumul -> argument. +% Eg. Definition or Axiom (when the body is none) +type const-decl id -> option term -> arity -> argument. +% Eg. #[universes(polymorphic,...)] Definition or Axiom +type upoly-const-decl id -> option term -> arity -> upoly-decl -> argument. +% Eg. Context A (b : A). +type ctx-decl context-decl -> argument. + +% Declaration of inductive types %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% + +kind indt-decl type. +kind indc-decl type. +kind record-decl type. + +% An arity is written, in Coq syntax, as: +% (x : T1) .. (xn : Tn) : S1 -> ... -> Sn -> U +% This syntax is used, for example, in the type of an inductive type or +% in the type of constructors. We call the abstractions on the left of ":" +% "parameters" while we call the type following the ":" (proper) arity. + +% Note: in some contexts, like the type of an inductive type constructor, +% Coq makes no distinction between these two writings +% (xn : Tn) : forall y1 : S1, ... and (xn : Tn) (y1 : S1) : ... +% while Elpi is a bit more restrictive, since it understands user directives +% such as the implicit status of an arguments (eg, using {} instead of () around +% the binder), only on parameters. +% Moreover parameters carry the name given by the user as an "id", while binders +% in terms only carry it as a "name", an irrelevant pretty pringintg hint (see +% also the HOAS of terms). A user command can hence only use the names of +% parameters, and not the names of "forall" quantified variables in the arity. +% +% See also the arity->term predicate in coq-lib.elpi + +type parameter id -> implicit_kind -> term -> (term -> arity) -> arity. +type arity term -> arity. + +type parameter id -> implicit_kind -> term -> (term -> indt-decl) -> indt-decl. +type inductive id -> bool -> arity -> (term -> list indc-decl) -> indt-decl. % tt means inductive, ff coinductive +type record id -> term -> id -> record-decl -> indt-decl. + +type constructor id -> arity -> indc-decl. + +type field field-attributes -> id -> term -> (term -> record-decl) -> record-decl. +type end-record record-decl. + +% Example. +% Remark that A is a regular parameter; y is a non-uniform parameter and t +% also features an index of type bool. +% +% Inductive t (A : Type) | (y : nat) : bool -> Type := +% | K1 (x : A) {n : nat} : S n = y -> t A n true -> t A y true +% | K2 : t A y false +% +% is written +% +% (parameter "A" explicit {{ Type }} a\ +% inductive "t" tt (parameter "y" explicit {{ nat }} _\ +% arity {{ bool -> Type }}) +% t\ +% [ constructor "K1" +% (parameter "y" explicit {{ nat }} y\ +% (parameter "x" explicit a x\ +% (parameter "n" maximal {{ nat }} n\ +% arity {{ S lp:n = lp:y -> lp:t lp:n true -> lp:t lp:y true }}))) +% , constructor "K2" +% (parameter "y" explicit {{ nat }} y\ +% arity {{ lp:t lp:y false }}) ]) +% +% Remark that the uniform parameters are not passed to occurrences of t, since +% they never change, while non-uniform parameters are both abstracted +% in each constructor type and passed as arguments to t. +% +% The coq.typecheck-indt-decl API can be used to fill in implicit arguments +% an infer universe constraints in the declaration above (e.g. the hidden +% argument of "=" in the arity of K1). +% +% Note: when and inductive type declaration is passed as an argument to an +% Elpi command non uniform parameters must be separated from the uniform ones +% with a | (a syntax introduced in Coq 8.12 and accepted by coq-elpi since +% version 1.4, in Coq this separator is optional, but not in Elpi). + +% Context declaration (used as an argument to Elpi commands) +kind context-decl type. +% Eg. (x : T) or (x := B), body is optional, type may be a variable +type context-item id -> implicit_kind -> term -> option term -> (term -> context-decl) -> context-decl. +type context-end context-decl. + +typeabbrev field-attributes (list field-attribute). + +macro @global! :- get-option "coq:locality" "global". +macro @local! :- get-option "coq:locality" "local". +|} diff --git a/src/coq_elpi_builtins_synterp.ml b/src/coq_elpi_builtins_synterp.ml new file mode 100644 index 000000000..f9111f2c4 --- /dev/null +++ b/src/coq_elpi_builtins_synterp.ml @@ -0,0 +1,901 @@ +(* coq-elpi: Coq terms as the object language of elpi *) +(* license: GNU Lesser General Public License Version 2.1 or later *) +(* ------------------------------------------------------------------------- *) + +module API = Elpi.API +module State = API.State +module Conv = API.Conversion +module CConv = API.ContextualConversion +module B = struct + include API.BuiltInData + include Elpi.Builtin +end +module Pred = API.BuiltInPredicate +module U = API.Utils + +open Names + +open Coq_elpi_utils +open Coq_elpi_HOAS + +let prop = { B.any with Conv.ty = Conv.TyName "prop" } + +let id = { B.string with + API.Conversion.ty = Conv.TyName "id"; + pp_doc = (fun fmt () -> + Format.fprintf fmt "%% [id] is a name that matters, we piggy back on Elpi's strings.@\n"; + Format.fprintf fmt "%% Note: [name] is a name that does not matter.@\n"; + Format.fprintf fmt "typeabbrev id string.@\n@\n") +} + +type scope = ExecutionSite | CurrentModule | Library + +let options : (options, API.Data.constraints) CConv.ctx_readback = + fun ~depth hyps constraints state -> + state, get_options ~depth hyps state, constraints, [] + +let scope = let open Conv in let open API.AlgebraicData in declare { + ty = TyName "scope"; + doc = "Specify to which module the clause should be attached to"; + pp = (fun fmt _ -> Format.fprintf fmt ""); + constructors = [ + K("execution-site","The module inside which the Elpi program is run",N, + B ExecutionSite, + M (fun ~ok ~ko -> function ExecutionSite -> ok | _ -> ko ())); + K("current","The module being defined (see begin/end-module)",N, + B CurrentModule, + M (fun ~ok ~ko -> function CurrentModule -> ok | _ -> ko ())); + K("library","The outermost module (carrying the file name)",N, + B Library, + M (fun ~ok ~ko -> function Library -> ok | _ -> ko ())) + ] +} |> CConv.(!<) + +let grafting = let open Conv in let open API.AlgebraicData in declare { + ty = TyName "grafting"; + doc = "Specify if the clause has to be grafted before or after a named clause"; + pp = (fun fmt _ -> Format.fprintf fmt ""); + constructors = [ + K("before","",A(id,N), + B (fun x -> (`Before,x)), + M (fun ~ok ~ko -> function (`Before,x) -> ok x | _ -> ko ())); + K("after","",A(id,N), + B (fun x -> (`After,x)), + M (fun ~ok ~ko -> function (`After,x) -> ok x | _ -> ko ())); + ] +} |> CConv.(!<) + +type clause = string option * ([ `After | `Before ] * string) option * API.Data.term + +let clause = let open Conv in let open API.AlgebraicData in declare { + ty = TyName "clause"; + doc = {|clauses + +A clause like + :name "foo" :before "bar" foo X Y :- bar X Z, baz Z Y +is represented as + clause "foo" (before "bar") (pi x y z\ foo x y :- bar x z, baz z y) +that is exactly what one would load in the context using =>. + +The name and the grafting specification can be left unspecified.|}; + pp = (fun fmt _ -> Format.fprintf fmt ""); + constructors = [ + K("clause","",A(B.unspec id,A(B.unspec grafting,A(prop,N))), + B (fun id graft c -> unspec2opt id, unspec2opt graft, c), + M (fun ~ok ~ko (id,graft,c) -> ok (opt2unspec id) (opt2unspec graft) c)); + ] +} |> CConv.(!<) + +let set_accumulate_to_db_synterp, get_accumulate_to_db_synterp = + let f = ref (fun _ -> assert false) in + (fun x -> f := x), + (fun () -> !f) + +let clauses_for_later_synterp : _ State.component = + State.declare_component ~name:"coq-elpi:clauses_for_later" ~descriptor:synterp_state + ~init:(fun () -> []) + ~start:(fun x -> x) + ~pp:(fun fmt l -> + List.iter (fun (dbname, code,vars,scope) -> + Format.fprintf fmt "db:%s code:%a scope:%a\n" + (String.concat "." dbname) + Elpi.API.Pp.Ast.program code Coq_elpi_utils.pp_scope scope) l) () + + +type located = + | LocGref of Names.GlobRef.t + | LocModule of Names.ModPath.t + | LocModuleType of Names.ModPath.t + | LocAbbreviation of Globnames.abbreviation + +let located = let open Conv in let open API.AlgebraicData in declare { + ty = TyName "located"; + doc = "Result of coq.locate-all"; + pp = (fun fmt _ -> Format.fprintf fmt ""); + constructors = [ + K("loc-gref","",A(gref,N), + B (fun x -> LocGref x), + M (fun ~ok ~ko -> function LocGref x -> ok x | _ -> ko ())); + K("loc-modpath","",A(modpath,N), + B (fun x -> LocModule x), + M (fun ~ok ~ko -> function LocModule x -> ok x | _ -> ko ())); + K("loc-modtypath","",A(modtypath,N), + B (fun x -> LocModuleType x), + M (fun ~ok ~ko -> function LocModuleType x -> ok x | _ -> ko ())); + K("loc-abbreviation","",A(abbreviation,N), + B (fun x -> LocAbbreviation x), + M (fun ~ok ~ko -> function LocAbbreviation x -> ok x | _ -> ko ())); + ] +} |> CConv.(!<) + + +let list = B.list +let pair = B.pair +let option = B.option +type accumulation_item = qualified_name * API.Ast.program * Id.t list * Coq_elpi_utils.clause_scope +let accumulate_clauses ~clauses_for_later ~accumulate_to_db ~preprocess_clause ~scope ~dbname clauses ~depth ~options state = + let loc = API.Ast.Loc.initial "(elpi.add_clause)" in + let dbname = Coq_elpi_utils.string_split_on_char '.' dbname in + let clauses scope = + clauses |> CList.rev_map (fun (name,graft,clause) -> + let vars, clause = preprocess_clause ~depth clause in + let clause = U.clause_of_term ?name ?graft ~depth loc clause in + (dbname,clause,vars,scope)) in + let local = (options : options).local = Some true in + let super_global = options.local = Some false in + match scope with + | B.Unspec | B.Given ExecutionSite -> + let scope = if super_global then SuperGlobal else if local then Local else Regular in + State.update clauses_for_later state (fun l -> + clauses scope @ l), (), [] + | B.Given Library -> + if local then CErrors.user_err Pp.(str "coq.elpi.accumulate: library scope is incompatible with @local!"); + State.update clauses_for_later state (fun l -> + clauses Coq_elpi_utils.Global @ l), (), [] + | B.Given CurrentModule -> + let scope = if local then Local else Regular in + let f = accumulate_to_db in + f (clauses scope); + state, (), [] + +let locate_module, locate_module_type = + let open API.BuiltIn in + let open Pred in + let open Notation in + MLCode(Pred("coq.locate-module", + In(id, "ModName", + Out(modpath, "ModPath", + Easy "locates a module. It's a fatal error if ModName cannot be located. *E*")), + (fun s _ ~depth -> + let qualid = Libnames.qualid_of_string s in + let mp = + try Nametab.locate_module qualid + with Not_found -> + err Pp.(str "Module not found: " ++ Libnames.pr_qualid qualid) in + !:mp)), + DocAbove), + MLCode(Pred("coq.locate-module-type", + In(id, "ModName", + Out(modtypath, "ModPath", + Easy "locates a module. It's a fatal error if ModName cannot be located. *E*")), + (fun s _ ~depth -> + let qualid = Libnames.qualid_of_string s in + let mp = + try Nametab.locate_modtype qualid + with Not_found -> + err Pp.(str "Module type not found: " ++ Libnames.pr_qualid qualid) in + !:mp)), + DocAbove) + +let modpath_to_path, modtypath_to_path, modpath_to_library, modtypath_to_library = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + MLCode(Pred("coq.modpath->path", + In(modpath, "MP", + Out(B.list B.string, "FullPath", + Read(unit_ctx, "extract the full kernel name, each component is a separate list item"))), + (fun mp _ ~depth h c state -> !: (mp2path mp))), + DocAbove), + MLCode(Pred("coq.modtypath->path", + In(modtypath, "MTP", + Out(B.list B.string, "FullPath", + Read(unit_ctx, "extract the full kernel name, each component is a separate list item"))), + (fun mtyp _ ~depth h c state -> !: (mp2path mtyp))), + DocAbove), + MLCode(Pred("coq.modpath->library", + In(modpath, "MP", + Out(modpath, "LibraryPath", + Read(unit_ctx, "extract the enclosing module which can be Required"))), + (fun mp _ ~depth h c state -> !: ModPath.(MPfile (dp mp)))), + DocAbove), + MLCode(Pred("coq.modtypath->library", + In(modtypath, "MTP", + Out(modpath, "LibraryPath", + Read(unit_ctx, "extract the enclosing module which can be Required"))), + (fun mtyp _ ~depth h c state -> !: ModPath.(MPfile (dp mtyp)))), + DocAbove) + +let current_path, current_section_path = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + MLCode(Pred("coq.env.current-path", + Out(list B.string, "Path", + Read(unit_ctx, "lists the current module path")), + (fun _ ~depth _ _ state -> !: (mp2path (Lib.current_mp ())))), + DocAbove), + MLCode(Pred("coq.env.current-section-path", + Out(list B.string, "Path", + Read(unit_ctx, "lists the current section path")), + (fun _ ~depth _ _ state -> + let base = Lib.current_dirpath false in + let base_w_sections = Lib.current_dirpath true in + let sections = Libnames.drop_dirpath_prefix base base_w_sections in + !: (mp2path (Names.ModPath.MPfile sections)))), + DocAbove) + +let invocation_site_loc : API.Ast.Loc.t State.component = + State.declare_component ~name:"coq-elpi:invocation-site-loc" ~descriptor:interp_state + ~pp:(fun fmt x -> Format.fprintf fmt "%a" API.Ast.Loc.pp x) + ~init:(fun () -> API.Ast.Loc.initial "(should-not-happen)") + ~start:(fun x -> x) () +let invocation_site_loc_synterp : API.Ast.Loc.t State.component = + State.declare_component ~name:"coq-elpi:invocation-site-loc" ~descriptor:synterp_state + ~pp:(fun fmt x -> Format.fprintf fmt "%a" API.Ast.Loc.pp x) + ~init:(fun () -> API.Ast.Loc.initial "(should-not-happen)") + ~start:(fun x -> x) () + +module SynterpAction = struct + open Declaremods + type action = + | BeginModule of (string * ModPath.t option * (string * ModPath.t) list) * module_params_expr * module_expr module_signature + | BeginModuleType of (string * (string * ModPath.t) list) * module_params_expr * module_expr list + | BeginSection of string + | EndModule of ModPath.t + | EndModuleType of ModPath.t + | EndSection + | ApplyModule of (string * ModPath.t option * ModPath.t * ModPath.t list * Declaremods.inline) * module_params_expr * module_expr list * module_expr module_signature + | ApplyModuleType of (string * ModPath.t * ModPath.t list * Declaremods.inline) * module_params_expr * module_expr list * module_expr list + | IncludeModule of (ModPath.t * Declaremods.inline) * module_expr list + | IncludeModuleType of (ModPath.t * Declaremods.inline) * module_expr list + | ImportModule of ModPath.t + | ExportModule of ModPath.t + + + +(* + | EVernacNotation of { local : bool; decl : Metasyntax.notation_interpretation_decl } + | EVernacSetOption of { export : bool; key : Goptions.option_name; value : Vernacexpr.option_setting } + *) + + type t = { action : action; resulting_state : Vernacstate.Synterp.t } + let synterp_state_after { resulting_state } = resulting_state + + let pp_action = function + | BeginModule ((id,_,_),_,_) -> + Pp.(str "begin-module" ++ spc () ++ qstring id) + | BeginModuleType ((id,_),_,_) -> + Pp.(str "begin-module-type" ++ spc () ++ qstring id) + | EndModule mp -> + Pp.(str "end-module" ++ spc () ++ (str @@ ModPath.to_string mp)) + | EndModuleType mp -> + Pp.(str "end-module-type" ++ spc () ++ (str @@ ModPath.to_string mp)) + | BeginSection id -> + Pp.(str "begin-section" ++ spc () ++ qstring id) + | EndSection -> + Pp.(str "end-section" ++ spc ()) + | ApplyModule ((id,_,_,_,_),_,_,_) -> + Pp.(str "apply-module" ++ spc () ++ qstring id) + | ApplyModuleType ((id,_,_,_),_,_,_) -> + Pp.(str "apply-module-type" ++ spc () ++ qstring id) + | IncludeModule ((mp,_),_) -> + Pp.(str "include-module" ++ spc () ++ (str @@ ModPath.to_string mp)) + | IncludeModuleType ((mp,_),_) -> + Pp.(str "include-module-type" ++ spc () ++ (str @@ ModPath.to_string mp)) + | ImportModule mp -> + Pp.(str "import-module" ++ spc () ++ (str @@ ModPath.to_string mp)) + | ExportModule mp -> + Pp.(str "export-module" ++ spc () ++ (str @@ ModPath.to_string mp)) + let pp { action } = pp_action action + + let action = + let open Conv in let open API.AlgebraicData in declare { + ty = TyName "synterp-action"; + doc = "Action executed during the parsing phase (aka synterp)"; + pp = (fun fmt a -> Pp.pp_with fmt (pp_action a)); + constructors = [ + K("begin-module","",A(id,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function BeginModule ((x,_,_),_,_) -> ok x | _ -> ko ())); + K("begin-module-type","",A(id,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function BeginModuleType ((x,_),_,_) -> ok x | _ -> ko ())); + K("begin-section","",A(id,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function BeginSection x -> ok x | _ -> ko ())); + K("end-module","",A(modpath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function EndModule x -> ok x | _ -> ko ())); + K("end-module-type","",A(modtypath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function EndModuleType x -> ok x | _ -> ko ())); + K("end-section","",N,B EndSection,M (fun ~ok ~ko -> function EndSection -> ok | _ -> ko ())); + K("apply-module-functor","",A(id,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function ApplyModule ((x,_,_,_,_),_,_,_) -> ok x | _ -> ko ())); + K("apply-module-type-functor","",A(id,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function ApplyModuleType ((x,_,_,_),_,_,_) -> ok x | _ -> ko ())); + K("include-module","",A(modpath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function IncludeModule ((x,_),_) -> ok x | _ -> ko ())); + K("include-module-type","",A(modtypath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function IncludeModuleType ((x,_),_) -> ok x | _ -> ko ())); + K("import-module","",A(modpath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function ImportModule x -> ok x | _ -> ko ())); + K("export-module","",A(modpath,N),B(fun x -> nYI "readback action"),M (fun ~ok ~ko -> function ExportModule x -> ok x | _ -> ko ())); + ] + } |> CConv.(!<) + + let log : t list State.component = + State.declare_component ~name:"coq-elpi:synterp-action-write" ~descriptor:synterp_state + ~pp:(fun fmt x -> Format.fprintf fmt "") + ~init:(fun () -> []) + ~start:(fun x -> x) () + + let get_parsing_actions_synterp = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + + [MLData action; + MLCode (Pred("coq.synterp-actions", + Out(list action,"A", + Read(unit_ctx,"Get the list of actions performed during the parsing phase (aka synterp) up to now.")), + (fun _ ~depth _ _ state -> !: (List.map (fun { action } -> action) @@ List.rev (State.get log state)))), + DocAbove) + ] + + let read : t list State.component = + State.declare_component ~name:"coq-elpi:synterp-action-read" ~descriptor:interp_state + ~pp:(fun fmt x -> Format.fprintf fmt "") + ~init:(fun () -> []) + ~start:(fun x -> x) () + + + let push action state = + State.update log state (fun l -> { action; resulting_state = Vernacstate.Synterp.freeze () } :: l) + + let common_err = "Interp actions must match synterp actions. For example if a module was imported during the synterp phase, then it must also be imported during the interp one." + + exception Error of Pp.t + let synterp_interp_error x = raise (Error x) + + let pop case state = + State.update_return read state (function + | x :: xs -> Vernacstate.Synterp.unfreeze x.resulting_state; xs, x.action + | _ -> synterp_interp_error Pp.(strbrk ("The elpi program did perform no (more) actions during the parsing phase (aka synterp), while during the execution phase (aka interp) it tried to perform a " ^ case ^ " action. " ^ common_err))) + + let pop_opt state = + State.update_return read state (function + | x :: xs -> Vernacstate.Synterp.unfreeze x.resulting_state; xs, Some x.action + | [] -> [], None) + + type 'a replay = 'a -> State.t -> State.t * ModPath.t option + + let replay1 action state = + match action with + | BeginModule((name,_,_),binders_ast,ty) -> + if Global.sections_are_opened () then + err Pp.(str"This elpi code cannot be run within a section since it opens a module"); + let id = Id.of_string name in + let mp = Declaremods.Interp.start_module None id binders_ast ty in + let loc = to_coq_loc @@ State.get invocation_site_loc state in + Dumpglob.dump_moddef ~loc mp "mod"; + None + | BeginModuleType((name,_),binders_ast,ty) -> + if Global.sections_are_opened () then + err Pp.(str"This elpi code cannot be run within a section since it opens a module"); + let id = Id.of_string name in + let mp = Declaremods.Interp.start_modtype id binders_ast ty in + let loc = to_coq_loc @@ State.get invocation_site_loc state in + Dumpglob.dump_moddef ~loc mp "modtype"; + None + | EndModule mp -> + let mp1 = Declaremods.Interp.end_module () in + assert(ModPath.equal mp mp1); + Some mp + | EndModuleType mp -> + let mp1 = Declaremods.Interp.end_modtype () in + assert(ModPath.equal mp mp1); + Some mp + | BeginSection name -> + let id = Id.of_string name in + let lid = CAst.make ~loc:(to_coq_loc @@ State.get invocation_site_loc state) id in + Dumpglob.dump_definition lid true "sec"; + Lib.Interp.open_section id; + None + | EndSection -> + let loc = to_coq_loc @@ State.get invocation_site_loc state in + Dumpglob.dump_reference ~loc + (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec"; + Lib.Interp.close_section (); + None + | ImportModule mp -> + Declaremods.import_module ~export:Lib.Import Libobject.unfiltered mp; + None + | ExportModule mp -> + Declaremods.Interp.import_module ~export:Lib.Export Libobject.unfiltered mp; + None + | IncludeModule(_,me) -> + Declaremods.Interp.declare_include me; + None + | IncludeModuleType (_,me) -> + Declaremods.Interp.declare_include me; + None + | ApplyModule ((name,_,_,_,_),params,mexpr_ast,ty) -> + if Global.sections_are_opened () then + err Pp.(str"This elpi code cannot be run within a section since it defines a module"); + let id = Id.of_string name in + let mp = Declaremods.Interp.declare_module id params ty mexpr_ast in + let loc = to_coq_loc @@ State.get invocation_site_loc state in + Dumpglob.dump_moddef ~loc mp "mod"; + Some mp + | ApplyModuleType ((name,_,_,_),params,mexpr_ast1,mexpr_ast2) -> + if Global.sections_are_opened () then + err Pp.(str"This elpi code cannot be run within a section since it defines a module"); + let id = Id.of_string name in + let mp = Declaremods.Interp.declare_modtype id params mexpr_ast1 mexpr_ast2 in + let loc = to_coq_loc @@ State.get invocation_site_loc state in + Dumpglob.dump_moddef ~loc mp "modtype"; + Some mp + + let rec replay state = + let state, action = pop_opt state in + match action with + | None -> state + | Some action -> ignore @@ replay1 action state; replay state + + let wrong_synterp_action x a = + synterp_interp_error Pp.(str "The elpi program did perform" ++ spc () ++ pp_action a ++ spc () ++ + str" at parsing time, while at run time it tried to perform" ++ spc () ++ str x ++ str(". "^common_err)) + let check_inconsistent_synterp_action eq ppx ppy x y a = + if not (eq x y) then + synterp_interp_error Pp.(str "The elpi program did perform" ++ spc () ++ pp_action a ++ spc () ++ + str" at parsing time and runtime, but on different inputs" ++ spc () ++ ppx x ++ str " != " ++ ppy y ++ str(". "^common_err)) + + let check_inconsistent_synterp_action_string = + check_inconsistent_synterp_action (=) Pp.str Pp.str + let check_inconsistent_synterp_action_modpath = + check_inconsistent_synterp_action ModPath.equal (fun x -> Pp.str @@ ModPath.to_string x) (fun x -> Pp.str @@ ModPath.to_string x) + + let eqU f x = function + | B.Given y -> f x y + | B.Unspec -> true + + let check_inconsistent_synterp_action_applym = + let eq (n1,t1,f1,a1,i1) (n2,t2,f2,a2,i2) = + n1 = n2 && + eqU (Option.equal ModPath.equal) t1 t2 && + eqU ModPath.equal f1 f2 && + eqU (List.equal ModPath.equal) a1 a2 && + eqU (=) i1 i2 + in + let ppx (n,t,f,a,i) = Pp.str n in + let ppy (n,t,f,a,i) = Pp.str n in + check_inconsistent_synterp_action eq ppx ppy + + let check_inconsistent_synterp_action_applymt = + let eq (n1,f1,a1,i1) (n2,f2,a2,i2) = + n1 = n2 && + eqU ModPath.equal f1 f2 && + eqU (List.equal ModPath.equal) a1 a2 && + eqU (=) i1 i2 + in + let ppx (n,f,a,i) = Pp.str n in + let ppy (n,f,a,i) = Pp.str n in + check_inconsistent_synterp_action eq ppx ppy + + let check_inconsistent_synterp_action_m = + let eq (n1,t1,a1) (n2,t2,a2) = + n1 = n2 && + eqU (Option.equal ModPath.equal) t1 t2 && + eqU (List.equal (fun (x1,y1) (x2,y2) -> x1 = x2 && ModPath.equal y1 y2)) a1 a2 + in + let ppx (n,t,a) = Pp.str n in + let ppy (n,t,a) = Pp.str n in + check_inconsistent_synterp_action eq ppx ppy + + let check_inconsistent_synterp_action_mt = + let eq (n1,a1) (n2,a2) = + n1 = n2 && + eqU (List.equal (fun (x1,y1) (x2,y2) -> x1 = x2 && ModPath.equal y1 y2)) a1 a2 + in + let ppx (n,a) = Pp.str n in + let ppy (n,a) = Pp.str n in + check_inconsistent_synterp_action eq ppx ppy + + let check_inconsistent_synterp_action_includem = + let eq (n1,t1) (n2,t2) = + ModPath.equal n1 n2 && + eqU (=) t1 t2 + in + let ppx (n,a) = Pp.str @@ ModPath.to_string n in + let ppy (n,a) = Pp.str @@ ModPath.to_string n in + check_inconsistent_synterp_action eq ppx ppy + + let pop_BeginModule (id,_,_ as x) state = + let case = Printf.sprintf "begin-module \"%s\"" id in + let state, a = pop case state in + match a with + | BeginModule(x',_,_) -> check_inconsistent_synterp_action_m x' x a; state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_BeginModuleType (id, _ as x') state = + let case = Printf.sprintf "begin-module-type \"%s\"" id in + let state, a = pop case state in + match a with + | BeginModuleType(x,_,_) -> check_inconsistent_synterp_action_mt x x' a; state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_BeginSection x' state = + let case = Printf.sprintf "begin-section \"%s\"" x' in + let state, a = pop case state in + match a with + | BeginSection x -> check_inconsistent_synterp_action_string x x' a; state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_EndModule () state = + let case = "end-module" in + let state, a = pop case state in + match a with + | EndModule _ -> state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_EndModuleType () state = + let case = "end-module-type" in + let state, a = pop case state in + match a with + | EndModuleType _ -> state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_EndSection () state = + let case = "end-section" in + let state, a = pop case state in + match a with + | EndSection -> state, replay1 a state + | _ -> wrong_synterp_action case a + let pop_ApplyModule a' state = + let case = "apply-module" in + let state, ac = pop case state in + match ac with + | ApplyModule (a,_,_,_) -> check_inconsistent_synterp_action_applym a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + let pop_ApplyModuleType a' state = + let case = "apply-module-type" in + let state, ac = pop case state in + match ac with + | ApplyModuleType (a,_,_,_) -> check_inconsistent_synterp_action_applymt a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + let pop_IncludeModule a' state = + let case = "include-module" in + let state, ac = pop case state in + match ac with + | IncludeModule (a,_) -> check_inconsistent_synterp_action_includem a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + let pop_IncludeModuleType a' state = + let case = "include-module-type" in + let state, ac = pop case state in + match ac with + | IncludeModuleType (a,_) -> check_inconsistent_synterp_action_includem a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + + let pop_ImportModule a' state = + let case = "import-module" in + let state, ac = pop case state in + match ac with + | ImportModule a -> check_inconsistent_synterp_action_modpath a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + let pop_ExportModule a' state = + let case = "export-module" in + let state, ac = pop case state in + match ac with + | ExportModule a -> check_inconsistent_synterp_action_modpath a a' ac; state, replay1 ac state + | _ -> wrong_synterp_action case ac + + + let builtins_interp = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + + [ + + LPDoc "-- Synterp ----------------------------------------------------------"; + + MLData action; + + MLCode (Pred("coq.next-synterp-action", + Out(action,"A", + Read(unit_ctx,"Get the next action performed during parsing (aka synterp), that is also the next action to be performed during execution (aka interp). See also coq.replay-synterp-action")), + (fun _ ~depth _ _ state -> !: (match State.get read state with [] -> raise No_clause | { action } :: _ -> action))), + DocAbove); + + MLCode(Pred("coq.replay-all-missing-synterp-actions", + Full(unit_ctx,"Execute all actions needed in order to match the actions performed at parsing time (aka synterp)"), + (fun ~depth _ _ state -> + let state = replay state in + state, (), [])), + DocAbove); + + + ] + + + +end + +let rec dirpath_of_modpath = function +| ModPath.MPfile d -> DirPath.repr d +| ModPath.MPdot(mp,l) -> Label.to_id l :: dirpath_of_modpath mp +| _ -> assert false + +let module_ast_of_modpath x = + let open Libnames in + qualid_of_dirpath (DirPath.make (dirpath_of_modpath x)) + +let module_ast_of_modtypath x = + let open Constrexpr in let open Libnames in + CAst.make @@ CMident (qualid_of_dirpath (DirPath.make (dirpath_of_modpath x))), + Declaremods.DefaultInline + +let coq_synterp_builtins = + let open API.BuiltIn in + let open Pred in + let open Notation in + let open CConv in + [ + LPCode Coq_elpi_builtins_arg_HOAS.code; + LPDoc "Coq terms are not visible at synterp time, they are always holes"; + LPCode "kind term type."; + LPDoc "-- Parsing time APIs ----------------------------------------------------"; + MLData id; + MLData modpath; + MLData modtypath; + locate_module; + locate_module_type; + LPCode {| +kind located type. +type loc-modpath modpath -> located. +type loc-modtypath modtypath -> located. +|}; + MLCode(Pred("coq.locate-all", + In(id, "Name", + Out(B.list located, "Located", + Easy {|finds all possible meanings of a string. Does not fail.|})), + (fun s _ ~depth -> + let qualid = Libnames.qualid_of_string s in + let l = ref [] in + let add x = l := !l @ [x] in + begin + try add @@ LocModule (Nametab.locate_module qualid) + with Not_found -> () + end; + begin + try add @@ LocModuleType (Nametab.locate_modtype qualid) + with Not_found -> () + end; + !: !l)), + DocAbove); + + MLData module_inline_default; + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.begin-module-functor", + In(id, "The name of the functor", + In(option modtypath, "Its module type", + In(list (pair id modtypath), "Parameters of the functor", + Full(unit_ctx, "Starts a functor *E*")))), + (fun name mp binders ~depth _ _ state -> + let ty = + match mp with + | None -> Declaremods.Check [] + | Some mp -> Declaremods.(Enforce (module_ast_of_modtypath mp)) in + let id = Id.of_string name in + let binders_ast = + List.map (fun (id, mty) -> + [CAst.make (Id.of_string id)], (module_ast_of_modtypath mty)) + binders in + let _, x, y = Declaremods.Synterp.start_module None id binders_ast ty in + let state = SynterpAction.(push (BeginModule((name,mp,binders), x, y))) state in + + state, (), [])), + DocNext); + + LPCode {| +pred coq.env.begin-module i:id, i:option modtypath. +coq.env.begin-module Name MP :- + coq.env.begin-module-functor Name MP []. +|}; + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.end-module", + Out(modpath, "ModPath", + Full(unit_ctx, "end the current module that becomes known as ModPath *E*")), + (fun _ ~depth _ _ state -> + let mp0 = Lib.current_mp () in + let mp = Declaremods.Synterp.end_module () in + let state = SynterpAction.(push (EndModule mp)) state in + assert(ModPath.equal mp0 mp); + state, !: mp, [])), + DocAbove); + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.begin-module-type-functor", + In(id, "The name of the functor", + In(list (pair id modtypath), "The parameters of the functor", + Full(unit_ctx,"Starts a module type functor *E*"))), + (fun name binders ~depth _ _ state -> + let id = Id.of_string name in + let binders_ast = + List.map (fun (id, mty) -> + [CAst.make (Id.of_string id)], (module_ast_of_modtypath mty)) + binders in + let _,y,z = Declaremods.Synterp.start_modtype id binders_ast [] in + let state = SynterpAction.(push (BeginModuleType((name,binders),y,z))) state in + + state, (), [])), + DocNext); + + LPCode {| +pred coq.env.begin-module-type i:id. +coq.env.begin-module-type Name :- + coq.env.begin-module-type-functor Name []. +|}; + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.end-module-type", + Out(modtypath, "ModTyPath", + Full(unit_ctx, "end the current module type that becomes known as ModPath *E*")), + (fun _ ~depth _ _ state -> + let mp0 = Lib.current_mp () in + let _mp = Declaremods.Synterp.end_modtype () in + (* BUG in COQ assert(ModPath.equal mp0 mp);*) + let state = SynterpAction.(push (EndModuleType(mp0))) state in + state, !: mp0, [])), + DocAbove); + + MLCode(Pred("coq.env.apply-module-functor", + In(id, "The name of the new module", + In(option modtypath, "Its module type", + In(modpath, "The functor being applied", + In(list modpath, "Its arguments", + In(module_inline_default, "Arguments inlining", + Out(modpath, "The modpath of the new module", + Full(unit_ctx, "Applies a functor *E*"))))))), + (fun name mp f arguments inline _ ~depth _ _ state -> + let ty = + match mp with + | None -> Declaremods.Check [] + | Some mp -> Declaremods.(Enforce (module_ast_of_modtypath mp)) in + let id = Id.of_string name in + let fa = CAst.make (Constrexpr.CMident (module_ast_of_modpath f)) in + let mexpr_ast_args = List.map module_ast_of_modpath arguments in + let mexpr_ast = + List.fold_left (fun hd arg -> CAst.make (Constrexpr.CMapply(hd,arg))) fa mexpr_ast_args in + let mp1, a,b,c = Declaremods.Synterp.declare_module id [] ty [mexpr_ast,inline] in + let state = SynterpAction.(push (ApplyModule((name,mp,f,arguments,inline),a,b,c))) state in + state, !: mp1, [])), + DocNext); + + MLCode(Pred("coq.env.apply-module-type-functor", + In(id, "The name of the new module type", + In(modtypath, "The functor", + In(list modpath, "Its arguments", + In(module_inline_default, "Arguments inlining", + Out(modtypath, "The modtypath of the new module type", + Full(unit_ctx, "Applies a type functor *E*")))))), + (fun name f arguments inline _ ~depth _ _ state -> + let id = Id.of_string name in + let fa,_ = module_ast_of_modtypath f in + let mexpr_ast_args = List.map module_ast_of_modpath arguments in + let mexpr_ast = + List.fold_left (fun hd arg -> CAst.make (Constrexpr.CMapply(hd,arg))) fa mexpr_ast_args in + let mp, a,b,c = Declaremods.Synterp.declare_modtype id [] [] [mexpr_ast,inline] in + let state = SynterpAction.(push (ApplyModuleType((name,f,arguments,inline),a,c,b (* c,b is intended, see Coq API*)))) state in + state, !: mp, [])), + DocNext); + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.include-module", + In(modpath, "ModPath", + In(module_inline_default, "Inline", + Full(unit_ctx, "is like the vernacular Include, Inline can be omitted *E*"))), + (fun mp inline ~depth _ _ state -> + let fpath = match mp with + | ModPath.MPdot(mp,l) -> + Libnames.make_path (ModPath.dp mp) (Label.to_id l) + | _ -> nYI "functors" in + let tname = Constrexpr.CMident (Libnames.qualid_of_path fpath) in + let i = CAst.make tname, inline in + let me = Declaremods.Synterp.declare_include [i] in + let state = SynterpAction.(push (IncludeModule ((mp,inline),me))) state in + state, (), [])), + DocAbove); + + (* XXX When Coq's API allows it, call vernacentries directly *) + MLCode(Pred("coq.env.include-module-type", + In(modtypath, "ModTyPath", + In(module_inline_default, "Inline", + Full(unit_ctx, "is like the vernacular Include Type, Inline can be omitted *E*"))), + (fun mp inline ~depth _ _ state -> + let fpath = Nametab.path_of_modtype mp in + let tname = Constrexpr.CMident (Libnames.qualid_of_path fpath) in + let i = CAst.make tname, inline in + let me = Declaremods.Synterp.declare_include [i] in + let state = SynterpAction.(push (IncludeModuleType ((mp,inline),me))) state in + state, (), [])), + DocAbove); + + MLCode(Pred("coq.env.import-module", + In(modpath, "ModPath", + Full(unit_ctx, "is like the vernacular Import *E*")), + (fun mp ~depth _ _ state -> + Declaremods.Synterp.import_module ~export:Lib.Import Libobject.unfiltered mp; + let state = SynterpAction.(push (ImportModule mp)) state in + state, (), [])), + DocAbove); + + MLCode(Pred("coq.env.export-module", + In(modpath, "ModPath", + Full(unit_ctx, "is like the vernacular Export *E*")), + (fun mp ~depth _ _ state -> + Declaremods.Synterp.import_module ~export:Lib.Export Libobject.unfiltered mp; + let state = SynterpAction.(push (ExportModule mp)) state in + state, (), [])), + DocAbove); + + MLCode(Pred("coq.env.begin-section", + In(id, "Name", + Full(unit_ctx, "starts a section named Name *E*")), + (fun name ~depth _ _ state -> + let id = Id.of_string name in + Lib.Synterp.open_section id; + let state = SynterpAction.(push (BeginSection name)) state in + state, (), [])), + DocAbove); + + MLCode(Pred("coq.env.end-section", + Full(unit_ctx, "end the current section *E*"), + (fun ~depth _ _ state -> + Lib.Synterp.close_section (); + let state = SynterpAction.(push EndSection) state in + state, (), [])), + DocAbove); + + modpath_to_path; modtypath_to_path; modpath_to_library; modtypath_to_library; + current_path; current_section_path; + + MLData clause; + MLData grafting; + MLData scope; + + LPCode {| +% see coq.elpi.accumulate-clauses +pred coq.elpi.accumulate i:scope, i:id, i:clause. +coq.elpi.accumulate S N C :- coq.elpi.accumulate-clauses S N [C]. +|}; + + MLCode(Pred("coq.elpi.accumulate-clauses", + In(B.unspec scope, "Scope", + In(id, "DbName", + In(B.list clause, "Clauses", + Full (options , {| +Declare that, once the program is over, the given clauses has to be +added to the given db (see Elpi Db). +Clauses usually belong to Coq modules: the Scope argument lets one +select which module: +- execution site (default) is the module in which the pogram is + invoked +- current is the module currently being constructed (see + begin/end-module) +- library is the current file (the module that is named after the file) +The clauses are visible as soon as the enclosing module is Imported. +Clauses cannot be accumulated inside functors. +Supported attributes: +- @local! (default: false, discard at the end of section or module) +- @global! (default: false, always active, only if Scope is execution-site, discouraged)|} )))), + (fun scope dbname clauses ~depth options _ state -> + accumulate_clauses + ~clauses_for_later:clauses_for_later_synterp + ~accumulate_to_db:(get_accumulate_to_db_synterp()) + ~preprocess_clause:(fun ~depth x -> [], x) + ~scope ~dbname clauses ~depth ~options state)), + DocAbove); + + ] @ SynterpAction.get_parsing_actions_synterp + +let synterp_api_doc = ". bla bla" + diff --git a/src/coq_elpi_builtins_synterp.mli b/src/coq_elpi_builtins_synterp.mli new file mode 100644 index 000000000..7d7089660 --- /dev/null +++ b/src/coq_elpi_builtins_synterp.mli @@ -0,0 +1,99 @@ +(* coq-elpi: Coq terms as the object language of elpi *) +(* license: GNU Lesser General Public License Version 2.1 or later *) +(* ------------------------------------------------------------------------- *) + +module SynterpAction : sig + type t + val builtins_interp : Elpi.API.BuiltIn.declaration list + + val log : t list Elpi.API.State.component + val read : t list Elpi.API.State.component + val pp : t -> Pp.t + + exception Error of Pp.t + + open Elpi.API + open Elpi.Builtin + open Names + open Declaremods + + val synterp_state_after : t -> Vernacstate.Synterp.t + + type 'a replay = 'a -> State.t -> State.t * ModPath.t option + + val pop_BeginModule : (string * ModPath.t option unspec * (string * ModPath.t) list unspec) replay + val pop_EndModule : unit replay + val pop_BeginModuleType : (string * (string * ModPath.t) list unspec) replay + val pop_EndModuleType : unit replay + + val pop_ApplyModule : + (string * ModPath.t option unspec * ModPath.t unspec * ModPath.t list unspec * inline unspec) replay + + val pop_ApplyModuleType : (string * ModPath.t unspec * ModPath.t list unspec * inline unspec) replay + val pop_IncludeModule : (ModPath.t * inline unspec) replay + val pop_IncludeModuleType : (ModPath.t * inline unspec) replay + val pop_ImportModule : ModPath.t replay + val pop_ExportModule : ModPath.t replay + val pop_BeginSection : string replay + val pop_EndSection : unit replay + +end + +open Elpi.API +open Coq_elpi_utils +open Names + +val clauses_for_later_synterp : + (qualified_name * Ast.program * Id.t list * Coq_elpi_utils.clause_scope) list State.component + +val set_accumulate_to_db_synterp : + ((qualified_name * Ast.program * Id.t list * Coq_elpi_utils.clause_scope) list -> unit) -> unit + +val prop : Data.term Conversion.t +val id : string Conversion.t + +type clause = string option * ([ `After | `Before ] * string) option * Data.term + +val clause : clause Conversion.t + +type scope = ExecutionSite | CurrentModule | Library + +val scope : scope Conversion.t +val grafting : ([ `After | `Before ] * string) Conversion.t +val options : (Coq_elpi_HOAS.options, Data.constraints) ContextualConversion.ctx_readback +val locate_module : BuiltIn.declaration +val locate_module_type : BuiltIn.declaration +val current_path : BuiltIn.declaration +val current_section_path : BuiltIn.declaration +val modpath_to_path : BuiltIn.declaration +val modtypath_to_path : BuiltIn.declaration +val modpath_to_library : BuiltIn.declaration +val modtypath_to_library : BuiltIn.declaration +val synterp_api_doc : string +val coq_synterp_builtins : BuiltIn.declaration list + +type located = + | LocGref of GlobRef.t + | LocModule of ModPath.t + | LocModuleType of ModPath.t + | LocAbbreviation of Globnames.abbreviation + +val located : located Conversion.t + +type accumulation_item = qualified_name * Ast.program * Id.t list * Coq_elpi_utils.clause_scope + +val accumulate_clauses : + clauses_for_later:accumulation_item list State.component -> + accumulate_to_db:(accumulation_item list -> unit) -> + preprocess_clause:(depth:int -> Data.term -> Id.t list * Data.term) -> + scope:scope Elpi.Builtin.unspec -> + dbname:string -> + clause list -> + depth:int -> + options:Coq_elpi_HOAS.options -> + State.t -> + State.t * unit * Conversion.extra_goals + + (* To dump glob, we need a quick access to the invocation site loc *) +val invocation_site_loc : Ast.Loc.t State.component +val invocation_site_loc_synterp : Ast.Loc.t State.component diff --git a/src/coq_elpi_glob_quotation.ml b/src/coq_elpi_glob_quotation.ml index bd1e2fbb4..a164de375 100644 --- a/src/coq_elpi_glob_quotation.ml +++ b/src/coq_elpi_glob_quotation.ml @@ -26,10 +26,10 @@ let get_elpi_code_appArg = ref (fun _ -> assert false) let get_ctx, set_ctx, _update_ctx = let bound_vars = - S.declare ~name:"coq-elpi:glob-quotation-bound-vars" + S.declare_component ~name:"coq-elpi:glob-quotation-bound-vars" ~descriptor:interp_state ~init:(fun () -> None) ~pp:(fun fmt -> function Some (x,_) -> () | None -> ()) - ~start:(fun x -> x) + ~start:(fun x -> x) () in S.(get bound_vars, set bound_vars, update bound_vars) @@ -60,8 +60,9 @@ let is_restricted_name = let glob_environment : Environ.env S.component = - S.declare ~name:"coq-elpi:glob-environment" - ~pp:(fun _ _ -> ()) ~init:Global.env ~start:(fun _ -> Global.env ()) + S.declare_component ~name:"coq-elpi:glob-environment" ~descriptor:interp_state + ~pp:(fun _ _ -> ()) ~init:(fun () -> Global.env ()) + ~start:(fun _ -> Global.env ()) () let push_env state name = let open Context.Rel.Declaration in @@ -400,10 +401,9 @@ let coq_quotation ~depth state loc src = (* Install the quotation *) -let () = Q.set_default_quotation coq_quotation -let () = Q.register_named_quotation ~name:"coq" coq_quotation - -let () = API.Quotation.register_named_quotation ~name:"gref" +let () = Q.set_default_quotation coq_quotation ~descriptor:interp_quotations +let () = Q.register_named_quotation ~name:"coq" coq_quotation ~descriptor:interp_quotations +let () = API.Quotation.register_named_quotation ~name:"gref" ~descriptor:interp_quotations (fun ~depth state _loc src -> let gr = locate_gref src in let state, gr, gls = gref.API.Conversion.embed ~depth state gr in diff --git a/src/coq_elpi_name_quotation.ml b/src/coq_elpi_name_quotation.ml index a1cf1cc01..0cd7d2e09 100644 --- a/src/coq_elpi_name_quotation.ml +++ b/src/coq_elpi_name_quotation.ml @@ -3,6 +3,7 @@ (* ------------------------------------------------------------------------- *) module API = Elpi.API +open Coq_elpi_utils open Coq_elpi_HOAS open Names @@ -11,11 +12,14 @@ let to_name src = else in_elpi_name (Name.Name (Id.of_string src)) (* Install the quotation *) -let () = API.Quotation.register_named_quotation ~name:"name" - (fun ~depth state _loc src -> state, to_name src) -;; +let () = + let f ~depth state _loc src = state, to_name src in + API.Quotation.register_named_quotation ~descriptor:interp_quotations ~name:"name" f; + API.Quotation.register_named_quotation ~descriptor:synterp_quotations ~name:"name" f -let () = API.Quotation.declare_backtick ~name:"Name.t" - (fun state s -> - let src = String.sub s 1 (String.length s - 2) in - state, to_name src) +let () = + let f state s = + let src = String.sub s 1 (String.length s - 2) in + state, to_name src in + API.Quotation.declare_backtick ~descriptor:interp_quotations ~name:"Name.t" f; + API.Quotation.declare_backtick ~descriptor:synterp_quotations ~name:"Name.t" f diff --git a/src/coq_elpi_programs.ml b/src/coq_elpi_programs.ml index 022dc8c64..aa71fdf53 100644 --- a/src/coq_elpi_programs.ml +++ b/src/coq_elpi_programs.ml @@ -8,6 +8,277 @@ module EP = API.Parse open Coq_elpi_utils type program_name = Loc.t * qualified_name +type cunit = Names.KerName.t * EC.compilation_unit + +module SLMap = Map.Make(struct type t = qualified_name let compare = compare_qualified_name end) +module SLSet = Set.Make(struct type t = qualified_name let compare = compare_qualified_name end) + +type src = + | File of src_file + | EmbeddedString of src_string + | Database of qualified_name +and src_file = { + fname : string; + fast : cunit; +} +and src_string = { + sloc : API.Ast.Loc.t; + sdata : string; + sast : cunit; +} + +let alpha = 65599 +let combine_hash h1 h2 = h1 * alpha + h2 + +let hash_cunit (kn,_) = Names.KerName.hash kn + +let hash_list f z l = List.fold_left (fun acc x -> combine_hash (f x) acc) z l + +module Chunk = struct + type t = + | Base of { + hash : int; + base : cunit; + } + | Snoc of { + source_rev : cunit list; + prev : t; + hash : int + } + let hash = function + | Base { hash } -> hash + | Snoc { hash } -> hash + + let snoc source_rev prev = + let hash = (hash_list hash_cunit (hash prev) source_rev) in + Snoc { hash; prev; source_rev } + + let eq x y = x == y +end + +module Code = struct +type 'db t = + | Base of { + hash : int; + base : cunit; + } + | Snoc of { + source : cunit; + prev : 'db t; + hash : int; + cacheme: bool; + } + | Snoc_db of { + chunks : 'db; + prev : 'db t; + hash : int + } +let hash = function + | Base { hash } -> hash + | Snoc { hash } -> hash + | Snoc_db { hash } -> hash + +let cache = function + | Base _ -> true + | Snoc { cacheme } -> cacheme + | Snoc_db _ -> false + +let snoc source prev = +let hash = combine_hash (hash prev) (hash_cunit source) in + Snoc { hash; prev; source; cacheme = cache prev } + +let snoc_opt source prev = + match prev with + | Some prev -> snoc source prev + | None -> Base { hash = (hash_cunit source); base = source } + +let snoc_db f chunks prev = + let hash = hash prev in + let hash = combine_hash (f chunks) hash in + Snoc_db { hash; prev; chunks } + +let snoc_db_opt chunks prev = + match prev with + | Some prev -> snoc_db (fun _ -> 0) chunks prev + | None -> assert false + +let rec map f = function + | Base x -> Base x + | Snoc { prev; source; cacheme = true } -> + (* no need to map, since prev is constant *) + let prev = Obj.magic prev in + snoc source prev + | Snoc { prev; source } -> + let prev = map f prev in + snoc source prev + | Snoc_db { prev; chunks } -> + let prev = map f prev in + let chunks = f chunks in + snoc_db Chunk.hash chunks prev + +let rec eq c x y = + x == y || + match x,y with + | Base x, Base y -> Names.KerName.equal (fst x.base) (fst y.base) + | Snoc x, Snoc y -> Names.KerName.equal (fst x.source) (fst y.source) && eq c x.prev y.prev + | Snoc_db x, Snoc_db y -> c x.chunks y.chunks && eq c x.prev y.prev + | _ -> false + +end + +type db = { + sources_rev : Chunk.t; + units : Names.KNset.t; +} + +type nature = Command of { raw_args : bool } | Tactic | Program of { raw_args : bool } + +type program = { + sources_rev : qualified_name Code.t; + units : Names.KNset.t; + dbs : SLSet.t; + empty : bool; (* it is empty, if it only contains default code *) +} + +type from = Initialization | User + +type snippet = { + program : qualified_name; + code :cunit list; + scope : Coq_elpi_utils.clause_scope; + vars : Names.Id.t list; +} + +let group_clauses l = + let rec aux acc l = + match acc, l with + | _, [] -> List.rev acc + | [], (dbname,ast,vs,scope) :: xs -> aux [dbname,[ast],vs,scope] xs + | (dbname,asts,vs,scope) :: tl as acc, (dbname1,ast,vs1,scope1) :: xs -> + if dbname = dbname1 && vs = vs1 && scope = scope1 then + aux ((dbname,asts@[ast],vs,scope) :: tl) xs + else + aux ((dbname1,[ast],vs1,scope1) :: acc) xs + in + aux [] l + +(**********************************************************************) + + +module type Programs = sig + val debug_vars : API.Compile.StrSet.t ref + val cc_flags : unit -> API.Compile.flags + val unit_from_file : elpi:API.Setup.elpi -> string -> cunit + val unit_from_string : elpi:API.Setup.elpi -> API.Ast.Loc.t -> string -> cunit + val assemble_units : elpi:API.Setup.elpi -> API.Compile.compilation_unit list -> API.Compile.program + val extend_w_units : base:API.Compile.program -> API.Compile.compilation_unit list -> API.Compile.program + val parse_goal : elpi:API.Setup.elpi -> API.Ast.Loc.t -> string -> API.Ast.query + val intern_unit : (string option * API.Compile.compilation_unit * API.Compile.flags) -> cunit + + val db_exists : qualified_name -> bool + val program_exists : qualified_name -> bool + val declare_db : program_name -> unit + val declare_program : program_name -> nature -> unit + val get_nature : qualified_name -> nature + + val init_program : program_name -> src -> unit + val init_db : program_name -> cunit -> unit + + val accumulate : qualified_name -> src list -> unit + val accumulate_to_db : qualified_name -> cunit list -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit + val load_checker : string -> unit + val load_printer : string -> unit + val load_command : string -> unit + val load_tactic : string -> unit + + val ensure_initialized : unit -> API.Setup.elpi + + val checker : unit -> API.Compile.compilation_unit list + val printer : unit -> API.Compile.compilation_unit + val tactic_init : unit -> src + val command_init : unit -> src + + val code : ?even_if_empty:bool -> qualified_name -> Chunk.t Code.t option + + val in_stage : string -> string + val stage : Summary.Stage.t + + val db_inspect : qualified_name -> int + +end + + +module type Stage = sig + val stage : Summary.Stage.t + val in_stage : string -> string + val init : unit -> API.Setup.elpi +end + +module SourcesStorage(S : Stage) = struct + open S + + module Libobject = struct + include Libobject + let declare_named_object o = Libobject.declare_named_object { o with + Libobject.object_stage = stage; + Libobject.object_name = in_stage (o.Libobject.object_name) + } + let declare_object o = Libobject.declare_object { o with + Libobject.object_stage = stage; + Libobject.object_name = in_stage (o.Libobject.object_name) + } + end + + module Summary = struct + include Summary + let ref ?local ~name x = Summary.ref ~stage ?local ~name:(in_stage name) x + end + + let program_name : nature SLMap.t ref = + Summary.ref ~name:("elpi-programs") SLMap.empty + let program_exists name = SLMap.mem name !program_name + + let get_nature p = + try SLMap.find p !program_name + with Not_found -> + CErrors.user_err + Pp.(str "No Elpi Program named " ++ pr_qualified_name p) + + let in_program_name : qualified_name * nature -> Libobject.obj = + let open Libobject in + declare_object @@ (superglobal_object_nodischarge "elpi-programs-names" + ~cache:(fun (name,nature) -> + program_name := SLMap.add name nature !program_name) + ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet")))) + + let declare_program name nature = + let obj = in_program_name (name,nature) in + Lib.add_leaf obj + + let db_name : SLSet.t ref = Summary.ref ~name:("elpi-dbs") SLSet.empty + let db_exists name = SLSet.mem name !db_name + + let in_db_name : qualified_name -> Libobject.obj = + let open Libobject in + declare_object @@ + (superglobal_object_nodischarge "elpi-db-names" + ~cache:(fun name -> db_name := SLSet.add name !db_name) + ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet")))) + let declare_db program = + ignore @@ Lib.add_leaf (in_db_name program) + + + let declare_program (loc,qualid) nature = + if program_exists qualid || db_exists qualid then + CErrors.user_err Pp.(str "Program/Tactic/Db " ++ pr_qualified_name qualid ++ str " already exists") + else + declare_program qualid nature + + let declare_db (loc,qualid) = + if program_exists qualid || db_exists qualid then + CErrors.user_err Pp.(str "Program/Tactic/Db " ++ pr_qualified_name qualid ++ str " already exists") + else + declare_db qualid let debug_vars = Summary.ref ~name:"elpi-debug" EC.StrSet.empty @@ -36,7 +307,6 @@ let in_source : Names.Id.t -> string option * EC.compilation_unit * EC.flags -> } ;; -type cunit = Names.KerName.t * EC.compilation_unit let dig = ref 0 let intern_unit (hash,u,flags) = @@ -106,154 +376,227 @@ let extend_w_units ~base units = let loc = Option.map Coq_elpi_utils.to_coq_loc oloc in CErrors.user_err ?loc Pp.(str (Option.default "" @@ Option.map API.Ast.Loc.show oloc) ++ str msg) -module SLMap = Map.Make(struct type t = qualified_name let compare = compare_qualified_name end) -module SLSet = Set.Make(struct type t = qualified_name let compare = compare_qualified_name end) - - -type src = -| File of src_file -| EmbeddedString of src_string -| Database of qualified_name -and src_file = { - fname : string; - fast : cunit; -} -and src_string = { - sloc : API.Ast.Loc.t; - sdata : string; - sast : cunit; -} - -let alpha = 65599 -let combine_hash h1 h2 = h1 * alpha + h2 - -let hash_cunit (kn,_) = Names.KerName.hash kn - -let hash_list f z l = List.fold_left (fun acc x -> combine_hash (f x) acc) z l - -module Chunk = struct -type t = -| Base of { - hash : int; - base : cunit; -} -| Snoc of { - source_rev : cunit list; - prev : t; - hash : int -} -let hash = function -| Base { hash } -> hash -| Snoc { hash } -> hash - -let snoc source_rev prev = -let hash = (hash_list hash_cunit (hash prev) source_rev) in -Snoc { hash; prev; source_rev } - -let eq x y = x == y -end - -module Code = struct -type 'db t = -| Base of { - hash : int; - base : cunit; -} -| Snoc of { - source : cunit; - prev : 'db t; - hash : int; - cacheme: bool; -} -| Snoc_db of { - chunks : 'db; - prev : 'db t; - hash : int -} -let hash = function -| Base { hash } -> hash -| Snoc { hash } -> hash -| Snoc_db { hash } -> hash +let program_src : program SLMap.t ref = + Summary.ref ~name:("elpi-programs-src") SLMap.empty -let cache = function -| Base _ -> true -| Snoc { cacheme } -> cacheme -| Snoc_db _ -> false +let db_name_src : db SLMap.t ref = + Summary.ref ~name:("elpi-db-src") SLMap.empty -let snoc source prev = -let hash = combine_hash (hash prev) (hash_cunit source) in -Snoc { hash; prev; source; cacheme = cache prev } - -let snoc_opt source prev = -match prev with -| Some prev -> snoc source prev -| None -> Base { hash = (hash_cunit source); base = source } + (* Setup called *) +let elpi = Stdlib.ref None -let snoc_db f chunks prev = -let hash = hash prev in -let hash = combine_hash (f chunks) hash in -Snoc_db { hash; prev; chunks } - -let snoc_db_opt chunks prev = -match prev with -| Some prev -> snoc_db (fun _ -> 0) chunks prev -| None -> assert false +let ensure_initialized = + let init = lazy (let e = S.init () in elpi := Some e; e) in + fun () -> Lazy.force init -let rec map f = function -| Base x -> Base x -| Snoc { prev; source; cacheme = true } -> - (* no need to map, since prev is constant *) - let prev = Obj.magic prev in - snoc source prev -| Snoc { prev; source } -> - let prev = map f prev in - snoc source prev -| Snoc_db { prev; chunks } -> - let prev = map f prev in - let chunks = f chunks in - snoc_db Chunk.hash chunks prev +let get ?(fail_if_not_exists=false) p = + let _elpi = ensure_initialized () in + let nature = get_nature p in + try + let { sources_rev; units; dbs; empty } = SLMap.find p !program_src in + units, dbs, Some nature, Some sources_rev, empty + with Not_found -> + if fail_if_not_exists then + CErrors.user_err + Pp.(str "No Elpi Program named " ++ pr_qualified_name p) + else + Names.KNset.empty, SLSet.empty, None, None, true + + let append_to_prog from name src = + let units, dbs, _, prog, empty = get name in + let units, dbs, prog = + match src with + (* undup *) + | File { fast = (kn,_) } when Names.KNset.mem kn units -> units, dbs, prog + | EmbeddedString { sast = (kn,_) } when Names.KNset.mem kn units -> units, dbs, prog + | Database n when SLSet.mem n dbs -> units, dbs, prog + (* add *) + | File { fast = (kn,_ as u) } -> (Names.KNset.add kn units), dbs, Some (Code.snoc_opt u prog) + | EmbeddedString { sast = (kn,_ as u) } -> (Names.KNset.add kn units), dbs, Some (Code.snoc_opt u prog) + | Database n -> units, SLSet.add n dbs, Some (Code.snoc_db_opt n prog) + in + let prog = Option.get prog in + { units; dbs; sources_rev = prog; empty = empty && from = Initialization } + -let rec eq c x y = - x == y || - match x,y with - | Base x, Base y -> Names.KerName.equal (fst x.base) (fst y.base) - | Snoc x, Snoc y -> Names.KerName.equal (fst x.source) (fst y.source) && eq c x.prev y.prev - | Snoc_db x, Snoc_db y -> c x.chunks y.chunks && eq c x.prev y.prev - | _ -> false + let in_program : qualified_name * src * from -> Libobject.obj = + let open Libobject in + declare_object @@ superglobal_object_nodischarge "ELPI" + ~cache:(fun (name,src_ast,from) -> + program_src := + SLMap.add name (append_to_prog from name src_ast) !program_src) + ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet"))) + + + let init_program name u = + let obj = in_program (name, u, Initialization) in + Lib.add_leaf obj + ;; + + let add_to_program name v = + let obj = in_program (name, v, User) in + Lib.add_leaf obj + + let code ?(even_if_empty=false) n : Chunk.t Code.t option = + let _,_,_,sources, empty = get n in + if empty && not even_if_empty then None else + sources |> Option.map (fun sources -> sources |> Code.map (fun name -> + try + let { sources_rev } : db = SLMap.find name !db_name_src in + sources_rev + with + Not_found -> + CErrors.user_err Pp.(str "Unknown Db " ++ str (show_qualified_name name)))) + + let append_to_db name kname c = + try + let (db : db) = SLMap.find name !db_name_src in + if Names.KNset.mem kname db.units then db + else { sources_rev = Chunk.snoc c db.sources_rev; units = Names.KNset.add kname db.units } + with Not_found -> + match c with + | [] -> assert false + | [base] -> + { sources_rev = Chunk.Base { hash = hash_cunit base; base }; units = Names.KNset.singleton kname } + | _ -> assert false + + let in_db : Names.Id.t -> snippet -> Libobject.obj = + let open Libobject in + let cache ((_,kn),{ program = name; code = p; _ }) = + db_name_src := SLMap.add name (append_to_db name kn p) !db_name_src in + let import i (_,s as o) = if Int.equal i 1 || s.scope = Coq_elpi_utils.SuperGlobal then cache o in + declare_named_object @@ { (default_object "ELPI-DB") with + classify_function = (fun { scope; program; _ } -> + match scope with + | Coq_elpi_utils.Local -> Dispose + | Coq_elpi_utils.Regular -> Substitute + | Coq_elpi_utils.Global -> Keep + | Coq_elpi_utils.SuperGlobal -> Keep); + load_function = import; + cache_function = cache; + subst_function = (fun (_,o) -> o); + open_function = simple_open import; + discharge_function = (fun (({ scope; program; vars; } as o)) -> + if scope = Coq_elpi_utils.Local || (List.exists (fun x -> Lib.is_in_section (Names.GlobRef.VarRef x)) vars) then None + else Some o); + } + + let accum = ref 0 + let add_to_db program code vars scope = + ignore @@ Lib.add_leaf + (in_db (Names.Id.of_string (incr accum; Printf.sprintf "_ELPI_%d" !accum)) { program; code; scope; vars }) + + let init_db qualid init = + add_to_db qualid [init] [] Coq_elpi_utils.SuperGlobal + + (* templates *) + let lp_command_ast = Summary.ref ~name:("elpi-lp-command") None + let in_lp_command_src : src -> Libobject.obj = + let open Libobject in + declare_object { (default_object "ELPI-LP-COMMAND") with + load_function = (fun _ x -> lp_command_ast := Some x); + } + let load_command s = + let elpi = ensure_initialized () in + let ast = File { + fname = s; + fast = unit_from_file ~elpi s + } in + lp_command_ast := Some ast; + Lib.add_leaf (in_lp_command_src ast) + let command_init () = + match !lp_command_ast with + | None -> CErrors.user_err Pp.(str "Elpi CommandTemplate was not called") + | Some ast -> ast + + let lp_tactic_ast = Summary.ref ~name:("elpi-lp-tactic") None + let in_lp_tactic_ast : src -> Libobject.obj = + let open Libobject in + declare_object { (default_object "ELPI-LP-TACTIC") with + load_function = (fun _ x -> lp_tactic_ast := Some x); + } + let load_tactic s = + let elpi = ensure_initialized () in + let ast = File { + fname = s; + fast = unit_from_file ~elpi s + } in + lp_tactic_ast := Some ast; + Lib.add_leaf (in_lp_tactic_ast ast) + let tactic_init () = + match !lp_tactic_ast with + | None -> CErrors.user_err Pp.(str "Elpi TacticTemplate was not called") + | Some ast -> ast + + let init_program (loc,qualid) (init : src) = + if stage = Summary.Stage.Interp && Global.sections_are_opened () then + CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside sections") + else + init_program qualid init + + let init_db (loc,qualid) (init : cunit) = + if stage = Summary.Stage.Interp && Global.sections_are_opened () then + CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside sections") + else if stage = Summary.Stage.Interp && match Global.current_modpath () with Names.ModPath.MPdot (Names.ModPath.MPfile _,_) -> true | _ -> false then + CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside modules") + else + init_db qualid init + ;; + + let lp_checker_ast = Summary.ref ~name:("elpi-lp-checker") None + let in_lp_checker_ast : cunit list -> Libobject.obj = + let open Libobject in + declare_object { (default_object "ELPI-LP-CHECKER") with + load_function = (fun _ x -> lp_checker_ast := Some x); + } + + let load_checker s = + let elpi = ensure_initialized () in + let basic_checker = unit_from_string ~elpi (Elpi.API.Ast.Loc.initial "(elpi-checker)") Elpi.Builtin_checker.code in + let coq_checker = unit_from_file ~elpi s in + let p = [basic_checker;coq_checker] in + Lib.add_leaf (in_lp_checker_ast p) + let checker () = + match !lp_checker_ast with + | None -> CErrors.user_err Pp.(str "Elpi Checker was not called") + | Some l -> List.map snd l + + let lp_printer_ast = Summary.ref ~name:("elpi-lp-printer") None + let in_lp_printer_ast : cunit -> Libobject.obj = + let open Libobject in + declare_object { (default_object "ELPI-LP-PRINTER") with + load_function = (fun _ x -> lp_printer_ast := Some x); + } + let load_printer s = + let elpi = ensure_initialized () in + let ast = unit_from_file ~elpi s in + lp_printer_ast := Some ast; + Lib.add_leaf (in_lp_printer_ast ast) + let printer () = + match !lp_printer_ast with + | None -> CErrors.user_err Pp.(str "Elpi Printer was not called") + | Some l -> snd l + + let accumulate p (v : src list) = + if not (program_exists p) then + CErrors.user_err Pp.(str "No Elpi Program named " ++ pr_qualified_name p); + v |> List.iter (add_to_program p) + + let accumulate_to_db p v vs ~scope = + if not (db_exists p) then + CErrors.user_err Pp.(str "No Elpi Db " ++ pr_qualified_name p); + add_to_db p v vs scope + + let in_stage = in_stage + let stage = stage + let db_inspect name = + try Names.KNset.cardinal (SLMap.find name !db_name_src).units + with Not_found -> -1 + end -type db = { - sources_rev : Chunk.t; - units : Names.KNset.t; -} - -type nature = Command of { raw_args : bool } | Tactic | Program of { raw_args : bool } - -let current_program = Summary.ref ~name:"elpi-cur-program-name" None - - -type program = { - sources_rev : qualified_name Code.t; - units : Names.KNset.t; - dbs : SLSet.t; -} - -let program_name : nature SLMap.t ref = - Summary.ref ~stage:Summary.Stage.Synterp ~name:"elpi-programs-synterp" SLMap.empty - -let program_src : program SLMap.t ref = - Summary.ref ~name:"elpi-programs" SLMap.empty - -let program_exists name = SLMap.mem name !program_name - -let db_name : SLSet.t ref = Summary.ref ~stage:Summary.Stage.Synterp ~name:"elpi-db-synterp" SLSet.empty -let db_name_src : db SLMap.t ref = Summary.ref ~name:"elpi-db" SLMap.empty -let db_exists name = SLSet.mem name !db_name - -(* Setup called *) -let elpi = Stdlib.ref None - +(***********************************************************************) let elpi_builtins = API.BuiltIn.declare ~file_name:"elpi-builtin.elpi" @@ -270,10 +613,24 @@ let elpi_builtins = [] ) -let coq_builtins = +let coq_interp_builtins = API.BuiltIn.declare ~file_name:"coq-builtin.elpi" - Coq_elpi_builtins.coq_builtins + begin + Coq_elpi_builtins.coq_header_builtins @ + Coq_elpi_builtins.coq_misc_builtins @ + Coq_elpi_builtins.coq_locate_builtins @ + Coq_elpi_builtins.coq_rest_builtins + end + +let coq_synterp_builtins = + API.BuiltIn.declare + ~file_name:"coq-builtin-synterp.elpi" + begin + Coq_elpi_builtins.coq_misc_builtins @ + Coq_elpi_builtins_synterp.coq_synterp_builtins + end + let file_resolver = let error_cannot_resolve dp file = @@ -335,291 +692,40 @@ fun ?cwd ~unit:file () -> | [] -> error_cannot_resolve dp file else legacy_resolver ?cwd ~unit:file () ;; - -let init () = - let e = API.Setup.init ~builtins:[coq_builtins;elpi_builtins] ~file_resolver () in - elpi := Some e; - e -;; - -let ensure_initialized = - let init = lazy (init ()) in - fun () -> Lazy.force init -;; - -let document_builtins () = - API.BuiltIn.document_file coq_builtins; - API.BuiltIn.document_file ~header:"% Generated file, do not edit" elpi_builtins - -(* We load pervasives and coq-lib once and forall at the beginning *) -let get_nature p = - try SLMap.find p !program_name - with Not_found -> - CErrors.user_err - Pp.(str "No Elpi Program named " ++ pr_qualified_name p) - -let get ?(fail_if_not_exists=false) p = - let _elpi = ensure_initialized () in - let nature = get_nature p in - try - let { sources_rev; units; dbs } = SLMap.find p !program_src in - units, dbs, Some nature, Some sources_rev - with Not_found -> - if fail_if_not_exists then - CErrors.user_err - Pp.(str "No Elpi Program named " ++ pr_qualified_name p) - else - Names.KNset.empty, SLSet.empty, None, None - -let code n : Chunk.t Code.t = - let _,_,_,sources = get n in - match sources with - | None -> CErrors.user_err Pp.(str "Unknown Program " ++ str (show_qualified_name n)) - | Some sources -> sources |> Code.map (fun name -> - try - let { sources_rev } : db = SLMap.find name !db_name_src in - sources_rev - with - Not_found -> - CErrors.user_err Pp.(str "Unknown Db " ++ str (show_qualified_name name))) - -let in_program_name : qualified_name * nature -> Libobject.obj = - let open Libobject in - declare_object @@ { (superglobal_object_nodischarge "ELPI_synterp" - ~cache:(fun (name,nature) -> - program_name := SLMap.add name nature !program_name) - ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet")))) - with object_stage = Summary.Stage.Synterp } - -let append_to_prog name src = - let units, dbs, _, prog = get name in - let units, dbs, prog = - match src with - (* undup *) - | File { fast = (kn,_) } when Names.KNset.mem kn units -> units, dbs, prog - | EmbeddedString { sast = (kn,_) } when Names.KNset.mem kn units -> units, dbs, prog - | Database n when SLSet.mem n dbs -> units, dbs, prog - (* add *) - | File { fast = (kn,_ as u) } -> (Names.KNset.add kn units), dbs, Some (Code.snoc_opt u prog) - | EmbeddedString { sast = (kn,_ as u) } -> (Names.KNset.add kn units), dbs, Some (Code.snoc_opt u prog) - | Database n -> units, SLSet.add n dbs, Some (Code.snoc_db_opt n prog) - in - let prog = Option.get prog in - { units; dbs; sources_rev = prog } - -let in_program : qualified_name * src -> Libobject.obj = - let open Libobject in - declare_object @@ superglobal_object_nodischarge "ELPI" - ~cache:(fun (name,src_ast) -> - program_src := - SLMap.add name (append_to_prog name src_ast) !program_src) - ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet"))) - - -let declare_program name nature = - let obj = in_program_name (name,nature) in - Lib.add_leaf obj - -let init_program name u = - let obj = in_program (name, u) in - Lib.add_leaf obj -;; - -let add_to_program name v = - let obj = in_program (name, v) in - Lib.add_leaf obj -;; - -let append_to_db name kname c = - try - let (db : db) = SLMap.find name !db_name_src in - if Names.KNset.mem kname db.units then db - else { sources_rev = Chunk.snoc c db.sources_rev; units = Names.KNset.add kname db.units } - with Not_found -> - match c with - | [] -> assert false - | [base] -> - { sources_rev = Chunk.Base { hash = hash_cunit base; base }; units = Names.KNset.singleton kname } - | _ -> assert false - -type snippet = { - program : qualified_name; - code :cunit list; - scope : Coq_elpi_utils.clause_scope; - vars : Names.Id.t list; -} - -let in_db : Names.Id.t -> snippet -> Libobject.obj = - let open Libobject in - let cache ((_,kn),{ program = name; code = p; _ }) = - db_name_src := SLMap.add name (append_to_db name kn p) !db_name_src in - let import i (_,s as o) = if Int.equal i 1 || s.scope = Coq_elpi_utils.SuperGlobal then cache o in - declare_named_object @@ { (default_object "ELPI-DB") with - classify_function = (fun { scope; program; _ } -> - match scope with - | Coq_elpi_utils.Local -> Dispose - | Coq_elpi_utils.Regular -> Substitute - | Coq_elpi_utils.Global -> Keep - | Coq_elpi_utils.SuperGlobal -> Keep); - load_function = import; - cache_function = cache; - subst_function = (fun (_,o) -> o); - open_function = simple_open import; - discharge_function = (fun (({ scope; program; vars; } as o)) -> - if scope = Coq_elpi_utils.Local || (List.exists (fun x -> Lib.is_in_section (Names.GlobRef.VarRef x)) vars) then None - else Some o); - } - -let in_db_name : qualified_name -> Libobject.obj = - let open Libobject in - declare_object @@ { - (superglobal_object_nodischarge "ELPI-DB_synterp" - ~cache:(fun name -> db_name := SLSet.add name !db_name) - ~subst:(Some (fun _ -> CErrors.user_err Pp.(str"elpi: No functors yet")))) - with - object_stage = Summary.Stage.Synterp } - -let declare_db program = - ignore @@ Lib.add_leaf (in_db_name program) - -let accum = ref 0 -let add_to_db program code vars scope = - ignore @@ Lib.add_leaf - (in_db (Names.Id.of_string (incr accum; Printf.sprintf "_ELPI_%d" !accum)) { program; code; scope; vars }) - -let lp_command_ast = Summary.ref ~name:"elpi-lp-command" None -let in_lp_command_src : src -> Libobject.obj = - let open Libobject in - declare_object { (default_object "ELPI-LP-COMMAND") with - load_function = (fun _ x -> lp_command_ast := Some x); - } -let load_command s = - let elpi = ensure_initialized () in - let ast = File { - fname = s; - fast = unit_from_file ~elpi s - } in - lp_command_ast := Some ast; - Lib.add_leaf (in_lp_command_src ast) -let command_init () = - match !lp_command_ast with - | None -> CErrors.user_err Pp.(str "Elpi CommandTemplate was not called") - | Some ast -> ast - -let lp_tactic_ast = Summary.ref ~name:"elpi-lp-tactic" None -let in_lp_tactic_ast : src -> Libobject.obj = - let open Libobject in - declare_object { (default_object "ELPI-LP-TACTIC") with - load_function = (fun _ x -> lp_tactic_ast := Some x); - } -let load_tactic s = - let elpi = ensure_initialized () in - let ast = File { - fname = s; - fast = unit_from_file ~elpi s - } in - lp_tactic_ast := Some ast; - Lib.add_leaf (in_lp_tactic_ast ast) -let tactic_init () = - match !lp_tactic_ast with - | None -> CErrors.user_err Pp.(str "Elpi TacticTemplate was not called") - | Some ast -> ast - -let declare_program (loc,qualid) nature = - if program_exists qualid || db_exists qualid then - CErrors.user_err Pp.(str "Program/Tactic/Db " ++ pr_qualified_name qualid ++ str " already exists") - else - declare_program qualid nature +(***********************************************************************) + +module Synterp : Programs = struct + module S = struct + let stage = Summary.Stage.Synterp + let in_stage x = x ^ "-synterp" + let init () = + API.Setup.init ~state:synterp_state ~hoas:synterp_hoas + ~quotations:synterp_quotations ~builtins:[coq_synterp_builtins;elpi_builtins] ~file_resolver () + end + include SourcesStorage(S) + + let () = Coq_elpi_builtins_synterp.set_accumulate_to_db_synterp (fun clauses_to_add -> + let elpi = ensure_initialized () in + let flags = cc_flags () in + let clauses_to_add = clauses_to_add |> group_clauses |> + List.map (fun (dbname,asts,vs,scope) -> + let units = List.map (unit_from_ast ~elpi ~flags None) asts in + dbname,units,vs,scope) in + clauses_to_add |> List.iter (fun (dbname,units,vs,scope) -> + accumulate_to_db dbname units vs ~scope)) -let init_program (loc,qualid) (init : src) = - if Global.sections_are_opened () then - CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside sections") - else - init_program qualid init - -let declare_db (loc,qualid) = - if program_exists qualid || db_exists qualid then - CErrors.user_err Pp.(str "Program/Tactic/Db " ++ pr_qualified_name qualid ++ str " already exists") - else - declare_db qualid - -let init_db (loc,qualid) (init : cunit) = - if Global.sections_are_opened () then - CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside sections") - else if match Global.current_modpath () with Names.ModPath.MPdot (Names.ModPath.MPfile _,_) -> true | _ -> false then - CErrors.user_err Pp.(str "Program/Tactic/Db cannot be declared inside modules") - else - add_to_db qualid [init] [] Coq_elpi_utils.SuperGlobal -;; - -let set_current_program n = - let _ = ensure_initialized () in - current_program := Some n - -let current_program () = - match !current_program with - | None -> CErrors.user_err Pp.(str "No current Elpi Program") - | Some x -> x - - -let lp_checker_ast = Summary.ref ~name:"elpi-lp-checker" None -let in_lp_checker_ast : cunit list -> Libobject.obj = - let open Libobject in - declare_object { (default_object "ELPI-LP-CHECKER") with - load_function = (fun _ x -> lp_checker_ast := Some x); - } - -let load_checker s = - let elpi = ensure_initialized () in - let basic_checker = unit_from_string ~elpi (Elpi.API.Ast.Loc.initial "(elpi-checker)") Elpi.Builtin_checker.code in - let coq_checker = unit_from_file ~elpi s in - let p = [basic_checker;coq_checker] in - Lib.add_leaf (in_lp_checker_ast p) -let checker () = - match !lp_checker_ast with - | None -> CErrors.user_err Pp.(str "Elpi Checker was not called") - | Some l -> List.map snd l - -let lp_printer_ast = Summary.ref ~name:"elpi-lp-printer" None -let in_lp_printer_ast : cunit -> Libobject.obj = - let open Libobject in - declare_object { (default_object "ELPI-LP-PRINTER") with - load_function = (fun _ x -> lp_printer_ast := Some x); - } -let load_printer s = - let elpi = ensure_initialized () in - let ast = unit_from_file ~elpi s in - lp_printer_ast := Some ast; - Lib.add_leaf (in_lp_printer_ast ast) -let printer () = - match !lp_printer_ast with - | None -> CErrors.user_err Pp.(str "Elpi Printer was not called") - | Some l -> snd l - -let accumulate p (v : src list) = - if not (program_exists p) then - CErrors.user_err Pp.(str "No Elpi Program named " ++ pr_qualified_name p); - v |> List.iter (add_to_program p) - -let accumulate_to_db p v vs ~scope = - if not (db_exists p) then - CErrors.user_err Pp.(str "No Elpi Db " ++ pr_qualified_name p); - add_to_db p v vs scope - -let group_clauses l = - let rec aux acc l = - match acc, l with - | _, [] -> List.rev acc - | [], (dbname,ast,vs,scope) :: xs -> aux [dbname,[ast],vs,scope] xs - | (dbname,asts,vs,scope) :: tl as acc, (dbname1,ast,vs1,scope1) :: xs -> - if dbname = dbname1 && vs = vs1 && scope = scope1 then - aux ((dbname,asts@[ast],vs,scope) :: tl) xs - else - aux ((dbname1,[ast],vs1,scope1) :: acc) xs - in - aux [] l -let () = Coq_elpi_builtins.set_accumulate_to_db (fun clauses_to_add -> +end +module Interp : Programs = struct + include SourcesStorage(struct + let stage = Summary.Stage.Interp + let in_stage x = x ^ "-interp" + let init () = + API.Setup.init ~state:interp_state ~hoas:interp_hoas + ~quotations:interp_quotations ~builtins:[coq_interp_builtins;elpi_builtins] ~file_resolver () + end) + +let () = Coq_elpi_builtins.set_accumulate_to_db_interp (fun clauses_to_add -> let elpi = ensure_initialized () in let flags = cc_flags () in let clauses_to_add = clauses_to_add |> group_clauses |> @@ -629,8 +735,16 @@ let () = Coq_elpi_builtins.set_accumulate_to_db (fun clauses_to_add -> clauses_to_add |> List.iter (fun (dbname,units,vs,scope) -> accumulate_to_db dbname units vs ~scope)) -let () = Coq_elpi_builtins.set_accumulate_text_to_db (fun n txt scope -> +let () = Coq_elpi_builtins.set_accumulate_text_to_db_interp (fun n txt scope -> let elpi = ensure_initialized () in let loc = API.Ast.Loc.initial "(elpi.add_predicate)" in let u = unit_from_string ~elpi loc txt in accumulate_to_db n [u] [] ~scope) + +end + +let document_builtins () = + API.BuiltIn.document_file coq_interp_builtins; + API.BuiltIn.document_file coq_synterp_builtins; + API.BuiltIn.document_file ~header:"% Generated file, do not edit" elpi_builtins + \ No newline at end of file diff --git a/src/coq_elpi_programs.mli b/src/coq_elpi_programs.mli index 63879f2fe..2512754ac 100644 --- a/src/coq_elpi_programs.mli +++ b/src/coq_elpi_programs.mli @@ -23,33 +23,6 @@ and src_string = { } type nature = Command of { raw_args : bool } | Tactic | Program of { raw_args : bool } -val get_nature : qualified_name -> nature - -val declare_program : program_name -> nature -> unit -val init_program : program_name -> src -> unit - -val declare_db : program_name -> unit -val init_db : program_name -> cunit -> unit - -val set_current_program : qualified_name -> unit -val current_program : unit -> qualified_name -val accumulate : qualified_name -> src list -> unit -val accumulate_to_db : qualified_name -> cunit list -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit -val group_clauses : - (qualified_name * Ast.program * Names.variable list * clause_scope) list -> - (qualified_name * Ast.program list * Names.variable list * clause_scope) list -val load_checker : string -> unit -val load_printer : string -> unit -val load_command : string -> unit -val load_tactic : string -> unit -val document_builtins : unit -> unit -val ensure_initialized : unit -> Setup.elpi -val db_exists : qualified_name -> bool -val checker : unit -> Compile.compilation_unit list -val printer : unit -> Compile.compilation_unit -val tactic_init : unit -> src -val command_init : unit -> src -val combine_hash : int -> int -> int module Chunk : sig type t = @@ -86,16 +59,61 @@ module Code : sig val hash : 'db t -> int val cache : 'db t -> bool val eq : ('db -> 'db -> bool) -> 'db t -> 'db t -> bool + val snoc_opt : cunit -> 'db t option -> 'db t end -val code : qualified_name -> Chunk.t Code.t -val debug_vars : Compile.StrSet.t ref -val cc_flags : unit -> Compile.flags -val unit_from_file : elpi:Setup.elpi -> string -> cunit -val unit_from_string : elpi:Setup.elpi -> Ast.Loc.t -> string -> cunit -val assemble_units : elpi:Setup.elpi -> Compile.compilation_unit list -> Compile.program -val extend_w_units : base:Compile.program -> Compile.compilation_unit list -> Compile.program -val parse_goal : elpi:Setup.elpi -> Ast.Loc.t -> string -> Ast.query -val intern_unit : (string option * Compile.compilation_unit * Compile.flags) -> cunit - module SLMap : Map.S with type key = qualified_name + +val combine_hash : int -> int -> int + +(* runtime *) + +module type Programs = sig + + val debug_vars : Compile.StrSet.t ref + val cc_flags : unit -> Compile.flags + val unit_from_file : elpi:Setup.elpi -> string -> cunit + val unit_from_string : elpi:Setup.elpi -> Ast.Loc.t -> string -> cunit + val assemble_units : elpi:Setup.elpi -> Compile.compilation_unit list -> Compile.program + val extend_w_units : base:Compile.program -> Compile.compilation_unit list -> Compile.program + val parse_goal : elpi:Setup.elpi -> Ast.Loc.t -> string -> Ast.query + val intern_unit : (string option * Compile.compilation_unit * Compile.flags) -> cunit + + val db_exists : qualified_name -> bool + val program_exists : qualified_name -> bool + val declare_db : program_name -> unit + val declare_program : program_name -> nature -> unit + val get_nature : qualified_name -> nature + + val init_program : program_name -> src -> unit + val init_db : program_name -> cunit -> unit + + val accumulate : qualified_name -> src list -> unit + val accumulate_to_db : qualified_name -> cunit list -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit + + val load_checker : string -> unit + val load_printer : string -> unit + val load_command : string -> unit + val load_tactic : string -> unit + + val ensure_initialized : unit -> Setup.elpi + + val checker : unit -> Compile.compilation_unit list + val printer : unit -> Compile.compilation_unit + val tactic_init : unit -> src + val command_init : unit -> src + + val code : ?even_if_empty:bool -> qualified_name -> Chunk.t Code.t option + + val in_stage : string -> string + val stage : Summary.Stage.t + val db_inspect : qualified_name -> int +end + +module Synterp : Programs +module Interp : Programs + +val group_clauses : + (qualified_name * Ast.program * Names.variable list * clause_scope) list -> + (qualified_name * Ast.program list * Names.variable list * clause_scope) list +val document_builtins : unit -> unit diff --git a/src/coq_elpi_utils.ml b/src/coq_elpi_utils.ml index 3a63fc9a4..1aecabd27 100644 --- a/src/coq_elpi_utils.ml +++ b/src/coq_elpi_utils.ml @@ -4,6 +4,15 @@ module API = Elpi.API +let synterp_quotations = API.Quotation.new_quotations_descriptor () +let synterp_hoas = API.RawData.new_hoas_descriptor () +let synterp_state = API.State.new_state_descriptor () + +let interp_quotations = API.Quotation.new_quotations_descriptor () +let interp_hoas = API.RawData.new_hoas_descriptor () +let interp_state = API.State.new_state_descriptor () + + let of_coq_loc l = { API.Ast.Loc.source_name = (match l.Loc.fname with Loc.InFile {file} -> file | Loc.ToplevelInput -> "(stdin)"); @@ -289,7 +298,7 @@ let mp2path x = let gr2path gr = let open Names in match gr with - | Names.GlobRef.VarRef v -> mp2path (Safe_typing.current_modpath (Global.safe_env ())) + | Names.GlobRef.VarRef v -> mp2path (Lib.current_mp ()) | Names.GlobRef.ConstRef c -> mp2path @@ Constant.modpath c | Names.GlobRef.IndRef (i,_) -> mp2path @@ MutInd.modpath i | Names.GlobRef.ConstructRef ((i,_),j) -> mp2path @@ MutInd.modpath i diff --git a/src/coq_elpi_utils.mli b/src/coq_elpi_utils.mli index f2bd1e5dd..03c351e37 100644 --- a/src/coq_elpi_utils.mli +++ b/src/coq_elpi_utils.mli @@ -2,6 +2,13 @@ (* license: GNU Lesser General Public License Version 2.1 or later *) (* ------------------------------------------------------------------------- *) +val synterp_hoas : Elpi.API.Setup.hoas_descriptor +val synterp_quotations : Elpi.API.Setup.quotations_descriptor +val synterp_state : Elpi.API.Setup.state_descriptor + +val interp_hoas : Elpi.API.Setup.hoas_descriptor +val interp_quotations : Elpi.API.Setup.quotations_descriptor +val interp_state : Elpi.API.Setup.state_descriptor val to_coq_loc : Elpi.API.Ast.Loc.t -> Loc.t val of_coq_loc : Loc.t -> Elpi.API.Ast.Loc.t diff --git a/src/coq_elpi_vernacular.ml b/src/coq_elpi_vernacular.ml index 5a1f86503..6943ef3e8 100644 --- a/src/coq_elpi_vernacular.ml +++ b/src/coq_elpi_vernacular.ml @@ -9,73 +9,100 @@ module EU = API.Utils module ET = API.RawData open Coq_elpi_utils +open Coq_elpi_arg_HOAS module Programs = Coq_elpi_programs (* ******************** Vernacular commands ******************************* *) open Programs -let load_command = load_command -let load_tactic = load_tactic -let load_printer = load_printer -let load_checker = load_checker -let document_builtins = document_builtins - -module Synterp = struct - let create_command ?(raw_args=false) n = - declare_program n (Command { raw_args }) - let create_tactic n = - declare_program n Tactic +let default_max_step = max_int - let create_program ?(raw_args=false) n = - declare_program n (Program { raw_args }) +let main_quotedc = ET.Constants.declare_global_symbol "main-quoted" +let mainc = ET.Constants.declare_global_symbol "main" +let main_interpc = ET.Constants.declare_global_symbol "main-interp" +let main_synterpc = ET.Constants.declare_global_symbol "main-synterp" +let attributesc = ET.Constants.declare_global_symbol "attributes" - let create_db n = declare_db n -end +let atts2impl loc ~depth state atts q = + let open Coq_elpi_builtins in + let rec convert_att_r = function + | (name,Attributes.VernacFlagEmpty) -> name, AttributeEmpty + | (name,Attributes.VernacFlagList l) -> name, AttributeList (convert_atts l) + | (name,Attributes.VernacFlagLeaf v) -> name, AttributeLeaf (convert_att_value v) + and convert_att att = convert_att_r att.CAst.v + and convert_atts l = List.map convert_att l + and convert_att_value = function + Attributes.FlagIdent s | Attributes.FlagString s -> AttributeString s + in + let atts = ("elpi.loc", AttributeLeaf (AttributeLoc loc)) :: convert_atts atts in + let atts = + match Sys.getenv_opt "COQ_ELPI_ATTRIBUTES" with + | None -> atts + | Some txt -> + match Pcoq.parse_string (Pvernac.main_entry None) (Printf.sprintf "#[%s] Qed." txt) |> Option.map (fun x -> x.CAst.v) with + | None -> atts + | Some { Vernacexpr.attrs ; _ } -> List.map (fun {CAst.v=(name,v)} -> convert_att_r ("elpi."^name,v)) attrs @ atts + | exception Gramlib.Stream.Error msg -> + CErrors.user_err Pp.(str"Environment variable COQ_ELPI_ATTRIBUTES contains ill formed value:" ++ spc () ++ str txt ++ cut () ++ str msg) in + let state, atts, _ = EU.map_acc (Coq_elpi_builtins.attribute.API.Conversion.embed ~depth) state atts in + let atts = ET.mkApp attributesc (EU.list_to_lp_list atts) [] in + state, ET.mkApp ET.Constants.implc atts [q] -let create_command n = - let _ = ensure_initialized () in - init_program n (command_init()); - set_current_program (snd n) +let same_phase y x = + match x, y with + | _, Both -> true + | Summary.Stage.Interp, Interp -> true + | Summary.Stage.Synterp, Synterp -> true + | _ -> false -let create_tactic n = - let _ = ensure_initialized () in - init_program n (tactic_init ()); - set_current_program (snd n) +let skip ?only ~ph ~phase f x = + let m rex = Str.string_match rex Coq_config.version 0 in + let exec1 = + match only with + | None -> true + | Some (None, None) -> true + | Some (Some skip, None) -> not (List.exists m skip) + | Some (None, Some keep) -> List.exists m keep + | Some (Some _, Some _) -> CErrors.user_err Pp.(str "Attributes #[skip] and #[only] cannot be used at the same time") + in + let exec2 = + match ph with + | None -> same_phase Interp phase + | Some ph -> same_phase ph phase + in + if exec1 && exec2 then f x else () -let create_program n ~init:(loc,s) = - let elpi = ensure_initialized () in - let unit = unit_from_string ~elpi loc s in - let init = EmbeddedString { sloc = loc; sdata = s; sast = unit} in - init_program n init; - set_current_program (snd n) +module Compiler(P : Programs) = struct -let create_db n ~init:(loc,s) = - let elpi = ensure_initialized () in - let unit = unit_from_string ~elpi loc s in - init_db n unit + module Summary = struct + include Summary + let ref ?local ~name x = Summary.ref ~stage:P.stage ?local ~name:(P.in_stage name) x + end -let default_max_step = max_int + let skip = skip ~phase:P.stage -let trace_options = Summary.ref ~name:"elpi-trace" [] -let max_steps = Summary.ref ~name:"elpi-steps" default_max_step + let trace_options = Summary.ref ~name:"elpi-trace" [] + let max_steps = Summary.ref ~name:"elpi-steps" default_max_step -let debug vl = debug_vars := List.fold_right EC.StrSet.add vl EC.StrSet.empty + let debug vl = P.debug_vars := List.fold_right EC.StrSet.add vl EC.StrSet.empty + let debug ~atts:ph vl = skip ~ph debug vl -let bound_steps n = - if n <= 0 then max_steps := default_max_step else max_steps := n + let bound_steps n = + if n <= 0 then max_steps := default_max_step else max_steps := n + let bound_steps ~atts:ph n = skip ~ph bound_steps n (* Units are marshalable, but programs are not *) let compiler_cache_code = Summary.ref ~local:true - ~name:"elpi-compiler-cache-code" + ~name:("elpi-compiler-cache-code") Int.Map.empty let compiler_cache_chunk = Summary.ref ~local:true - ~name:"elpi-compiler-cache-chunk" + ~name:("elpi-compiler-cache-chunk") Int.Map.empty let programs_tip = Summary.ref ~local:true - ~name:"elpi-compiler-cache-gc" + ~name:("elpi-compiler-cache-gc") SLMap.empty (* lookup/cache for hash h shifted over base b *) @@ -108,28 +135,26 @@ let recache_chunk b h1 h2 p src = uncache b h1 compiler_cache_chunk; cache_chunk b h2 p src -let get_and_compile name = - let src = code name in - let prog = - let rec compile_code src = - let h = Code.hash src in - try - lookup_code 0 h src - with Not_found -> - match src with - | Code.Base { base = (k,u) } -> - let elpi = ensure_initialized () in - let prog = assemble_units ~elpi [u] in - cache_code 0 h prog src - | Code.Snoc { prev; source } -> - let base = compile_code prev in - let prog = extend_w_units ~base [snd source] in - if Code.cache src then cache_code 0 h prog src else prog - | Code.Snoc_db { prev; chunks } -> - let base = compile_code prev in - let prog = compile_chunk (Code.hash prev) base chunks in - prog - and compile_chunk bh base src = +let compile src = + let rec compile_code src = + let h = Code.hash src in + try + lookup_code 0 h src + with Not_found -> + match src with + | Code.Base { base = (k,u) } -> + let elpi = P.ensure_initialized () in + let prog = P.assemble_units ~elpi [u] in + cache_code 0 h prog src + | Code.Snoc { prev; source } -> + let base = compile_code prev in + let prog = P.extend_w_units ~base [snd source] in + if Code.cache src then cache_code 0 h prog src else prog + | Code.Snoc_db { prev; chunks } -> + let base = compile_code prev in + let prog = compile_chunk (Code.hash prev) base chunks in + prog + and compile_chunk bh base src = (* DBs have to be re-based on top of base, hence bh *) let h = Chunk.hash src in try @@ -137,29 +162,33 @@ let get_and_compile name = with Not_found -> match src with | Chunk.Base { base = (k,u) } -> - let prog = extend_w_units ~base [u] in + let prog = P.extend_w_units ~base [u] in cache_chunk bh h prog src | Chunk.Snoc { prev; source_rev } -> let base = compile_chunk bh base prev in - let prog = extend_w_units ~base (List.rev_map snd source_rev) in + let prog = P.extend_w_units ~base (List.rev_map snd source_rev) in recache_chunk bh (Chunk.hash prev) h prog src - in - let prog = compile_code src in + in + compile_code src + +let get_and_compile ?even_if_empty name : (EC.program * bool) option = + P.code ?even_if_empty name |> Option.map (fun src -> + let prog = compile src in let new_hash = Code.hash src in let old_hash = try SLMap.find name !programs_tip with Not_found -> 0 in programs_tip := SLMap.add name new_hash !programs_tip; - recache_code 0 old_hash new_hash prog src in - let raw_args = - match get_nature name with - | Command { raw_args } -> raw_args - | Program { raw_args } -> raw_args - | Tactic -> true in - prog, raw_args + let prog = recache_code 0 old_hash new_hash prog src in + let raw_args = + match P.get_nature name with + | Command { raw_args } -> raw_args + | Program { raw_args } -> raw_args + | Tactic -> true in + (prog, raw_args)) let run_static_check query = let checker = - let elpi = ensure_initialized () in - assemble_units ~elpi (checker()) in + let elpi = P.ensure_initialized () in + P.assemble_units ~elpi (P.checker()) in (* We turn a failure into a proper error in etc/coq-elpi_typechecker.elpi *) ignore (EC.static_check ~checker query) @@ -167,7 +196,7 @@ let run ~static_check program query = let t1 = Unix.gettimeofday () in let query = match query with - | `Ast query_ast -> EC.query program query_ast + | `Ast (f,query_ast) -> API.RawQuery.compile_ast program query_ast f | `Fun query_builder -> API.RawQuery.compile program query_builder in let t2 = Unix.gettimeofday () in let _ = API.Setup.trace [] in @@ -192,11 +221,12 @@ let elpi_fails program_name = let kind = "tactic/command" in let name = show_qualified_name program_name in CErrors.user_err (strbrk (String.concat " " [ - "The elpi"; kind; name ; "failed without giving a specific error message."; + "The elpi"; kind; name ; + "failed without giving a specific error message" ^ (Summary.Stage.(match P.stage with Interp -> "." | Synterp -> " (during the parsing phase, aka synterp).")); "Please report this inconvenience to the authors of the program." ])) -let run_and_print ~print ~static_check program_name program_ast query_ast : unit = +let run_and_print ~print ~static_check program_name program_ast query_ast : _ * Coq_elpi_builtins_synterp.SynterpAction.t list = let open API.Data in let open Coq_elpi_utils in match run ~static_check program_ast query_ast with @@ -204,7 +234,7 @@ let run_and_print ~print ~static_check program_name program_ast query_ast : unit | API.Execute.NoMoreSteps -> CErrors.user_err Pp.(str "elpi run out of steps ("++int !max_steps++str")") | API.Execute.Success { - assignments ; constraints; state; pp_ctx + assignments ; constraints; state; pp_ctx; relocate_assignment_to_runtime; } -> if print then begin if not (StrMap.is_empty assignments) then begin @@ -219,158 +249,375 @@ let run_and_print ~print ~static_check program_name program_ast query_ast : unit let scst = pp2string (EPP.constraints pp_ctx) constraints in if scst <> "" then Feedback.msg_notice Pp.(str"Syntactic constraints:" ++ spc()++str scst); - let sigma = Coq_elpi_HOAS.get_sigma state in - let ccst = Evd.evar_universe_context sigma in - if not (UState.is_empty ccst) then - Feedback.msg_notice Pp.(str"Universe constraints:" ++ spc() ++ - Termops.pr_evar_universe_context ccst) + if P.stage = Summary.Stage.Interp then begin + let sigma = Coq_elpi_HOAS.get_sigma state in + let ccst = Evd.evar_universe_context sigma in + if not (UState.is_empty ccst) then + Feedback.msg_notice Pp.(str"Universe constraints:" ++ spc() ++ + Termops.pr_evar_universe_context ccst); + end end; (* We add clauses declared via coq.elpi.accumulate *) let clauses_to_add = - API.State.get Coq_elpi_builtins.clauses_for_later + API.State.get + (match P.stage with + | Summary.Stage.Synterp -> Coq_elpi_builtins_synterp.clauses_for_later_synterp + | Summary.Stage.Interp -> Coq_elpi_builtins.clauses_for_later_interp) state in - let elpi = ensure_initialized () in - let flags = cc_flags() in + let elpi = P.ensure_initialized () in + let flags = P.cc_flags() in let clauses_to_add = clauses_to_add |> group_clauses |> List.map (fun (dbname,asts,vs,scope) -> let units = asts |> List.map (fun ast -> EC.unit ~elpi ~flags ast) in - let units = units |> List.map (fun unit -> intern_unit (None,unit,flags)) in + let units = units |> List.map (fun unit -> P.intern_unit (None,unit,flags)) in dbname,units,vs,scope) in clauses_to_add |> List.iter (fun (dbname,units,vs,scope) -> - accumulate_to_db dbname units vs ~scope) -;; + P.accumulate_to_db dbname units vs ~scope); + relocate_assignment_to_runtime, + if P.stage = Summary.Stage.Synterp then API.State.get Coq_elpi_builtins_synterp.SynterpAction.log state |> List.rev + else API.State.get Coq_elpi_builtins_synterp.SynterpAction.read state + +let current_program = Summary.ref ~name:"elpi-cur-program-name" None +let set_current_program n = + current_program := Some n + +let typecheck_program ?program () = + match !current_program with + | None -> () + | Some program -> + let elpi = P.ensure_initialized () in + get_and_compile program |> Option.iter (fun (program, _) -> + let query_ast = P.parse_goal ~elpi (API.Ast.Loc.initial "(typecheck)") "true." in + let query = EC.query program query_ast in + let _ = API.Setup.trace !trace_options in + run_static_check query) + +let current_program () = + match !current_program with + | None -> CErrors.user_err Pp.(str "No current Elpi Program") + | Some x -> x + +let run_in_program ?(program = current_program ()) ?(st_setup=fun x -> x) (loc, query) = + let elpi = P.ensure_initialized () in + get_and_compile ~even_if_empty:true program |> Option.map (fun (program_ast, _) -> + let query_ast = `Ast (st_setup, P.parse_goal ~elpi loc query) in + run_and_print ~print:true ~static_check:true program program_ast query_ast) + + let accumulate_extra_deps ?(program=current_program()) ids = + let elpi = P.ensure_initialized () in + let s = ids |> List.map (fun id -> + try ComExtraDeps.query_extra_dep id + with + | Not_found when P.stage = Summary.Stage.Interp -> CErrors.anomaly Pp.(str"wtf") + | Not_found -> + err Pp.(str"File " ++ Names.Id.print id ++ + str" is unknown; please add a directive like 'From .. Extra Dependency .. as " ++ + Names.Id.print id ++ str"'.")) in + try + let new_src_ast = List.map (fun fname -> + File { + fname; + fast = P.unit_from_file ~elpi fname; + }) s in + P.accumulate program new_src_ast + with Failure s -> CErrors.user_err Pp.(str s) + let accumulate_extra_deps ~atts:(only,ph) ?program ids = skip ~only ~ph (accumulate_extra_deps ?program) ids + + let accumulate_files ?(program=current_program()) s = + let elpi = P.ensure_initialized () in + try + let new_src_ast = List.map (fun fname -> + File { + fname; + fast = P.unit_from_file ~elpi fname; + }) s in + P.accumulate program new_src_ast + with Failure s -> CErrors.user_err Pp.(str s) + let accumulate_files ~atts:(only,ph) ?program s = skip ~only ~ph (accumulate_files ?program) s + + let accumulate_string ?(program=current_program()) (loc,s) = + let elpi = P.ensure_initialized () in + let new_ast = P.unit_from_string ~elpi loc s in + if P.db_exists program then + P.accumulate_to_db program [new_ast] [] ~scope:Coq_elpi_utils.Regular + else + P.accumulate program [EmbeddedString { sloc = loc; sdata = s; sast = new_ast}] + let accumulate_string ~atts:(only,ph) ?program loc = skip ~only ~ph (accumulate_string ?program) loc + + + let accumulate_db ?(program=current_program()) name = + let _ = P.ensure_initialized () in + if P.db_exists name then P.accumulate program [Database name] + else CErrors.user_err Pp.(str "Db " ++ pr_qualified_name name ++ str" not found") + let accumulate_db ~atts:(only,ph) ?program name = skip ~only ~ph (accumulate_db ?program) name + + + let accumulate_to_db db (loc,s) idl ~scope = + let elpi = P.ensure_initialized () in + let new_ast = P.unit_from_string ~elpi loc s in + if P.db_exists db then P.accumulate_to_db db [new_ast] idl ~scope + else CErrors.user_err Pp.(str "Db " ++ pr_qualified_name db ++ str" not found") + let accumulate_to_db ~atts:(only,ph) db loc idl ~scope = skip ~only ~ph (accumulate_to_db db loc ~scope) idl + -let run_in_program ?(program = current_program ()) (loc, query) = - let elpi = ensure_initialized () in - let program_ast, _ = get_and_compile program in - let query_ast = `Ast (parse_goal ~elpi loc query) in - run_and_print ~print:true ~static_check:true program program_ast query_ast -;; + let mk_trace_opts start stop preds = + [ "-trace-on" + ; "-trace-at"; "run"; string_of_int start; string_of_int stop + ; "-trace-only"; "\\(run\\|select\\|user:\\)" + ; "-trace-tty-maxbox"; "30" + ] @ List.(flatten (map (fun x -> ["-trace-only-pred"; x]) preds)) + + let trace start stop preds opts = + if start = 0 && stop = 0 then trace_options := [] + else trace_options := mk_trace_opts start stop preds @ opts + let trace ~atts start stop preds opts = skip ~ph:atts (trace start stop preds) opts + + let trace_browser _opts = + trace_options := + [ "-trace-on"; "json"; "/tmp/traced.tmp.json" + ; "-trace-at"; "run"; "0"; string_of_int max_int + ; "-trace-only"; "user" + ]; + Feedback.msg_notice + Pp.(strbrk "Now click \"Start watching\" in the Elpi Trace Browser panel and then execute the Command/Tactic/Query you want to trace. Also try \"F1 Elpi\".") + let trace_browser ~atts opts = skip ~ph:atts trace_browser opts + + let print name args = + let elpi = P.ensure_initialized () in + get_and_compile name |> Option.iter (fun (program, _) -> + let args, fname, fname_txt = + let default_fname = String.concat "." name ^ ".html" in + let default_fname_txt = String.concat "." name ^ ".txt" in + let default_blacklist = [ + "elaborator.elpi";"reduction.elpi"; + "coq-builtin.elpi";"coq-lib.elpi";"coq-HOAS.elpi" + ] in + match args with + | [] -> default_blacklist, default_fname, default_fname_txt + | [x] -> default_blacklist, x ^ ".html", x ^ ".txt" + | x :: xs -> xs, x ^ ".html", x ^ ".txt" in + let args = List.map API.RawOpaqueData.of_string args in + let query_ast = Interp.parse_goal ~elpi (API.Ast.Loc.initial "(print)") "true." in + let query = EC.query program query_ast in + let oc = open_out fname_txt in + let fmt = Format.formatter_of_out_channel oc in + EPP.program fmt query; + Format.pp_print_flush fmt (); + close_out oc; + let loc = { API.Ast.Loc. + source_name = "(Elpi Print)"; + source_start = 0; + source_stop = 0; + line = -1; + line_starts_at = 0; } in + let q ~depth state = + let state, quotedP, _ = API.Quotation.quote_syntax_compiletime state query in + assert(depth=0); (* else, we should lift the terms down here *) + let q = ET.mkApp main_quotedc + (EU.list_to_lp_list quotedP) + [API.RawOpaqueData.of_string fname; EU.list_to_lp_list args] in + state, (loc,q), [] in + ignore @@ run_and_print ~print:false ~static_check:false ["Elpi";"Print"] + (P.assemble_units ~elpi [P.printer ()]) (`Fun q)) + let print ~atts name args = skip ~ph:atts (print name) args + + + let create_command ~atts:(raw_args) n = + let raw_args = Option.default false raw_args in + let _ = P.ensure_initialized () in + P.declare_program n (Command { raw_args }); + P.init_program n (P.command_init()); + set_current_program (snd n) -let typecheck_program ?(program = current_program ()) () = - let elpi = ensure_initialized () in - let program, _ = get_and_compile program in - let query_ast = parse_goal ~elpi (API.Ast.Loc.initial "(typecheck)") "true." in - let query = EC.query program query_ast in - let _ = API.Setup.trace !trace_options in - run_static_check query -;; + let create_tactic n = + let _ = P.ensure_initialized () in + P.declare_program n Tactic; + if P.stage = Summary.Stage.Interp then P.init_program n (P.tactic_init ()); + set_current_program (snd n) + + let create_program ~atts:(raw_args) n ~init:(loc,s) = + let raw_args = Option.default false raw_args in + let elpi = P.ensure_initialized () in + P.declare_program n (Program { raw_args }); + let unit = P.unit_from_string ~elpi loc s in + let init = EmbeddedString { sloc = loc; sdata = s; sast = unit} in + P.init_program n init; + set_current_program (snd n) + + let create_db ~atts n ~init:(loc,s) = + let do_init = + match atts with + | None -> same_phase Interp P.stage + | Some phase -> same_phase phase P.stage in + let elpi = P.ensure_initialized () in + P.declare_db n; + if do_init then + let unit = P.unit_from_string ~elpi loc s in + P.init_db n unit + + let load_command = P.load_command + let load_tactic = P.load_tactic + let load_printer = P.load_printer + let load_checker = P.load_checker + + let get_and_compile qn = get_and_compile ?even_if_empty:None qn -let mainc = ET.Constants.declare_global_symbol "main" -let attributesc = ET.Constants.declare_global_symbol "attributes" +end -let atts2impl loc ~depth state atts q = - let open Coq_elpi_builtins in - let rec convert_att_r = function - | (name,Attributes.VernacFlagEmpty) -> name, AttributeEmpty - | (name,Attributes.VernacFlagList l) -> name, AttributeList (convert_atts l) - | (name,Attributes.VernacFlagLeaf v) -> name, AttributeLeaf (convert_att_value v) - and convert_att att = convert_att_r att.CAst.v - and convert_atts l = List.map convert_att l - and convert_att_value = function - Attributes.FlagIdent s | Attributes.FlagString s -> AttributeString s - in - let atts = ("elpi.loc", AttributeLeaf (AttributeLoc loc)) :: convert_atts atts in - let atts = - match Sys.getenv_opt "COQ_ELPI_ATTRIBUTES" with - | None -> atts - | Some txt -> - match Pcoq.parse_string (Pvernac.main_entry None) (Printf.sprintf "#[%s] Qed." txt) |> Option.map (fun x -> x.CAst.v) with - | None -> atts - | Some { Vernacexpr.attrs ; _ } -> List.map (fun {CAst.v=(name,v)} -> convert_att_r ("elpi."^name,v)) attrs @ atts - | exception Gramlib.Stream.Error msg -> - CErrors.user_err Pp.(str"Environment variable COQ_ELPI_ATTRIBUTES contains ill formed value:" ++ spc () ++ str txt ++ cut () ++ str msg) in - let state, atts, _ = EU.map_acc (Coq_elpi_builtins.attribute.API.Conversion.embed ~depth) state atts in - let atts = ET.mkApp attributesc (EU.list_to_lp_list atts) [] in - state, ET.mkApp ET.Constants.implc atts [q] -;; -let run_program loc name ~atts args = - let program, raw_args = get_and_compile name in - let loc = Coq_elpi_utils.of_coq_loc loc in - let env = Global.env () in - let sigma = Evd.from_env env in - let args = args - |> List.map (Coq_elpi_arg_HOAS.Cmd.glob (Genintern.empty_glob_sign ~strict:true env)) - |> List.map (Coq_elpi_arg_HOAS.Cmd.interp (Ltac_plugin.Tacinterp.default_ist ()) env sigma) - in - let query ~depth state = - let state, args, gls = EU.map_acc - (Coq_elpi_arg_HOAS.in_elpi_cmd ~depth ~raw:raw_args Coq_elpi_HOAS.(mk_coq_context ~options:(default_options ()) state)) - state args in - let state, q = atts2impl loc ~depth state atts (ET.mkApp mainc (EU.list_to_lp_list args) []) in - let state = API.State.set Coq_elpi_builtins.tactic_mode state false in - let state = API.State.set Coq_elpi_builtins.invocation_site_loc state loc in - state, (loc, q), gls in - - run_and_print ~print:false ~static_check:false name program (`Fun query) -;; +module type Common = sig + val typecheck_program : ?program:qualified_name -> unit -> unit + val get_and_compile : + qualified_name -> (Elpi.API.Compile.program * bool) option + val run : static_check:bool -> + Elpi.API.Compile.program -> + [ `Ast of (Elpi.API.State.t -> Elpi.API.State.t) * Elpi.API.Ast.query + | `Fun of + depth:int -> + Elpi.API.State.t -> + Elpi.API.State.t * + (Elpi.API.Ast.Loc.t * Elpi.API.Data.term) * + Elpi.API.Conversion.extra_goals ] -> + unit Elpi.API.Execute.outcome + + val accumulate_files : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> string list -> unit + val accumulate_extra_deps : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> Names.Id.t list -> unit + val accumulate_string : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> unit + val accumulate_db : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> qualified_name -> unit + val accumulate_to_db : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> qualified_name -> Elpi.API.Ast.Loc.t * string -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit + + val load_checker : string -> unit + val load_printer : string -> unit + val load_tactic : string -> unit + val load_command : string -> unit + + val debug : atts:phase option -> string list -> unit + val trace : atts:phase option -> int -> int -> string list -> string list -> unit + val trace_browser : atts:phase option -> string list -> unit + val bound_steps : atts:phase option -> int -> unit + val print : atts:phase option -> qualified_name -> string list -> unit + + val create_program : atts:bool option -> program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit + val create_command : atts:bool option -> program_name -> unit + val create_tactic : program_name -> unit + val create_db : atts:phase option -> program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit -let mk_trace_opts start stop preds = - [ "-trace-on" - ; "-trace-at"; "run"; string_of_int start; string_of_int stop - ; "-trace-only"; "\\(run\\|select\\|user:\\)" - ; "-trace-tty-maxbox"; "30" - ] @ List.(flatten (map (fun x -> ["-trace-only-pred"; x]) preds)) - -let trace start stop preds opts = - if start = 0 && stop = 0 then trace_options := [] - else trace_options := mk_trace_opts start stop preds @ opts - -let trace_browser _opts = - trace_options := - [ "-trace-on"; "json"; "/tmp/traced.tmp.json" - ; "-trace-at"; "run"; "0"; string_of_int max_int - ; "-trace-only"; "user" - ]; - Feedback.msg_notice - Pp.(strbrk "Now click \"Start watching\" in the Elpi Trace Browser panel and then execute the Command/Tactic/Query you want to trace. Also try \"F1 Elpi\".") +end -let main_quotedc = ET.Constants.declare_global_symbol "main-quoted" +module Synterp = struct + include Compiler(Coq_elpi_programs.Synterp) -let print name args = - let elpi = ensure_initialized () in - let args, fname, fname_txt = - let default_fname = String.concat "." name ^ ".html" in - let default_fname_txt = String.concat "." name ^ ".txt" in - let default_blacklist = [ - "elaborator.elpi";"reduction.elpi"; - "coq-builtin.elpi";"coq-lib.elpi";"coq-HOAS.elpi" - ] in - match args with - | [] -> default_blacklist, default_fname, default_fname_txt - | [x] -> default_blacklist, x ^ ".html", x ^ ".txt" - | x :: xs -> xs, x ^ ".html", x ^ ".txt" in - let args = List.map API.RawOpaqueData.of_string args in - let program, _ = get_and_compile name in - let query_ast = parse_goal ~elpi (API.Ast.Loc.initial "(print)") "true." in - let query = EC.query program query_ast in - let oc = open_out fname_txt in - let fmt = Format.formatter_of_out_channel oc in - EPP.program fmt query; - Format.pp_print_flush fmt (); - close_out oc; - let loc = { API.Ast.Loc. - source_name = "(Elpi Print)"; - source_start = 0; - source_stop = 0; - line = -1; - line_starts_at = 0; } in - let q ~depth state = - let state, quotedP, _ = API.Quotation.quote_syntax_compiletime state query in - assert(depth=0); (* else, we should lift the terms down here *) - let q = ET.mkApp main_quotedc - (EU.list_to_lp_list quotedP) - [API.RawOpaqueData.of_string fname; EU.list_to_lp_list args] in - state, (loc,q), [] in - run_and_print ~print:false ~static_check:false ["Elpi";"Print"] - (assemble_units ~elpi [printer ()]) (`Fun q) -;; + let main_syterp args syndata = + ET.mkApp API.RawData.Constants.orc + (ET.mkApp main_synterpc (EU.list_to_lp_list args) [syndata]) + [ET.mkApp mainc (EU.list_to_lp_list args) []] + + let run_program loc name ~atts args = + get_and_compile name |> Option.map (fun (program, raw_args) -> + let loc = Coq_elpi_utils.of_coq_loc loc in + let query ~depth state = + let state, args, gls = EU.map_acc + (Coq_elpi_arg_HOAS.in_elpi_cmd_synterp ~depth) + state args in + let state, ek = API.FlexibleData.Elpi.make ~name:"NewData" state in + let data = ET.mkUnifVar ek ~args:[] state in + let state, q = atts2impl loc ~depth state atts (main_syterp args data) in + let state = API.State.set Coq_elpi_builtins_synterp.invocation_site_loc_synterp state loc in + state, (loc, q), gls in + try + let relocate, synterplog = run_and_print ~print:false ~static_check:false name program (`Fun query) in + relocate "NewData", synterplog + with Coq_elpi_builtins_synterp.SynterpAction.Error x -> err x) + + let run_in_program ?program query = + try run_in_program ?program query |> Option.map (fun (_,x) -> x) + with Coq_elpi_builtins_synterp.SynterpAction.Error x -> err x + +end + +let document_builtins = document_builtins + +module Interp = struct + + include Compiler(Coq_elpi_programs.Interp) + + let main_interp args syndata = + ET.mkApp API.RawData.Constants.orc + (ET.mkApp main_interpc (EU.list_to_lp_list args) [syndata]) + [ET.mkApp mainc (EU.list_to_lp_list args) []] + + let missing_synterp = let open Pp in + cut() ++ + str "The command lacks code for the synterp phase. In order to add code to this phase use '#[synterp] Elpi Accumulate'. See also https://lpcic.github.io/coq-elpi/tutorial_coq_elpi_command.html#parsing-and-execution" + + + let run_program loc name ~atts ~syndata args = + get_and_compile name |> Option.iter (fun (program, raw_args) -> + let loc = Coq_elpi_utils.of_coq_loc loc in + let env = Global.env () in + let sigma = Evd.from_env env in + let args = args + |> List.map (Coq_elpi_arg_HOAS.Cmd.glob (Genintern.empty_glob_sign ~strict:true env)) + |> List.map (Coq_elpi_arg_HOAS.Cmd.interp (Ltac_plugin.Tacinterp.default_ist ()) env sigma) + in + let synterplog, synterm = + match syndata with + | None -> [], fun ~target:_ ~depth -> Stdlib.Result.Ok ET.mkDiscard + | Some (relocate_term,log) -> log, relocate_term in + let initial_synterp_state = + match synterplog with + | [] -> Vernacstate.Synterp.freeze () + | x :: _ -> Coq_elpi_builtins_synterp.SynterpAction.synterp_state_after x in + let query ~depth state = + let state, args, gls = EU.map_acc + (Coq_elpi_arg_HOAS.in_elpi_cmd ~depth ~raw:raw_args Coq_elpi_HOAS.(mk_coq_context ~options:(default_options ()) state)) + state args in + let synterm = + match synterm ~target:state ~depth with + | Stdlib.Result.Ok x -> x + | Stdlib.Result.Error s -> err Pp.(str"Data returned by the synterp phase cannot be passed to the interp phase due to unknown symbol: " ++ str s) in + + let state, q = atts2impl loc ~depth state atts (main_interp args synterm) in + let state = API.State.set Coq_elpi_builtins.tactic_mode state false in + let state = API.State.set Coq_elpi_builtins_synterp.invocation_site_loc state loc in + let state = API.State.set Coq_elpi_builtins_synterp.SynterpAction.read state synterplog in + state, (loc, q), gls in + let final_synterp_state = Vernacstate.Synterp.freeze () in + Vernacstate.Synterp.unfreeze initial_synterp_state; + try begin + try + let _, synterp_left = run_and_print ~print:false ~static_check:false name program (`Fun query) in + match synterp_left with + | [] -> Vernacstate.Synterp.unfreeze final_synterp_state + | x :: _ -> + err Pp.(strbrk"The interpretation phase did not consume all the parsing time effects. Next in the queue is " ++ Coq_elpi_builtins_synterp.SynterpAction.pp x) + with + | Coq_elpi_builtins_synterp.SynterpAction.Error x when syndata = None -> + err Pp.(x ++ missing_synterp) + | Coq_elpi_builtins_synterp.SynterpAction.Error x -> + err x + end with e -> + let e = Exninfo.capture e in + Vernacstate.Synterp.unfreeze final_synterp_state; + Exninfo.iraise e) + + let run_in_program ?program ~syndata query = + let st_setup state = + let syndata = Option.default [] syndata in + API.State.set Coq_elpi_builtins_synterp.SynterpAction.read state syndata in + try ignore @@ run_in_program ?program ~st_setup query + with + | Coq_elpi_builtins_synterp.SynterpAction.Error x when syndata = None -> err Pp.(x ++ missing_synterp) + | Coq_elpi_builtins_synterp.SynterpAction.Error x -> err x + + +end open Tacticals let run_tactic_common loc ?(static_check=false) program ~main ?(atts=[]) () = let open Proofview in let open Notations in + let open Interp in Unsafe.tclGETGOALS >>= fun gls -> let gls = CList.map Proofview.drop_state gls in Proofview.tclEVARMAP >>= fun sigma -> @@ -380,75 +627,24 @@ let run_tactic_common loc ?(static_check=false) program ~main ?(atts=[]) () = ~in_elpi_tac_arg:Coq_elpi_arg_HOAS.in_elpi_tac ~depth state in let state, qatts = atts2impl loc ~depth state atts q in let state = API.State.set Coq_elpi_builtins.tactic_mode state true in - let state = API.State.set Coq_elpi_builtins.invocation_site_loc state loc in + let state = API.State.set Coq_elpi_builtins_synterp.invocation_site_loc state loc in state, (loc, qatts), gls in - let cprogram, _ = get_and_compile program in - match run ~static_check cprogram (`Fun query) with - | API.Execute.Success solution -> Coq_elpi_HOAS.tclSOLUTION2EVD sigma solution - | API.Execute.NoMoreSteps -> CErrors.user_err Pp.(str "elpi run out of steps") - | API.Execute.Failure -> elpi_fails program - | exception (Coq_elpi_utils.LtacFail (level, msg)) -> tclFAILn level msg + get_and_compile program |> Option.cata (fun (cprogram, _) -> + match run ~static_check cprogram (`Fun query) with + | API.Execute.Success solution -> Coq_elpi_HOAS.tclSOLUTION2EVD sigma solution + | API.Execute.NoMoreSteps -> CErrors.user_err Pp.(str "elpi run out of steps") + | API.Execute.Failure -> elpi_fails program + | exception (Coq_elpi_utils.LtacFail (level, msg)) -> tclFAILn level msg) + tclIDTAC let run_tactic loc program ~atts _ist args = let loc = Coq_elpi_utils.of_coq_loc loc in run_tactic_common loc program ~main:(Coq_elpi_HOAS.Solve args) ~atts () -let run_in_tactic ?(program = current_program ()) (loc,query) _ist = +let run_in_tactic ?(program = Interp.current_program ()) (loc,query) _ist = run_tactic_common loc ~static_check:true program ~main:(Coq_elpi_HOAS.Custom query) () -let accumulate_extra_deps ?(program=current_program()) ids = - let elpi = ensure_initialized () in - let s = ids |> List.map (fun id -> - try ComExtraDeps.query_extra_dep id - with Not_found -> - err Pp.(str"File " ++ Names.Id.print id ++ - str" is unknown; please add a directive like 'From .. Extra Dependency .. as " ++ - Names.Id.print id ++ str"'.")) in - try - let new_src_ast = List.map (fun fname -> - File { - fname; - fast = unit_from_file ~elpi fname; - }) s in - accumulate program new_src_ast - with Failure s -> CErrors.user_err Pp.(str s) - ;; - -let accumulate_files ?(program=current_program()) s = - let elpi = ensure_initialized () in - try - let new_src_ast = List.map (fun fname -> - File { - fname; - fast = unit_from_file ~elpi fname; - }) s in - accumulate program new_src_ast - with Failure s -> CErrors.user_err Pp.(str s) - ;; - -let accumulate_string ?(program=current_program()) (loc,s) = - let elpi = ensure_initialized () in - let new_ast = unit_from_string ~elpi loc s in - if db_exists program then - accumulate_to_db program [new_ast] [] ~scope:Coq_elpi_utils.Regular - else - accumulate program [EmbeddedString { sloc = loc; sdata = s; sast = new_ast}] -;; - -let accumulate_db ?(program=current_program()) name = - let _ = ensure_initialized () in - if db_exists name then accumulate program [Database name] - else CErrors.user_err - Pp.(str "Db " ++ pr_qualified_name name ++ str" not found") - -let accumulate_to_db db (loc,s) = - let elpi = ensure_initialized () in - let new_ast = unit_from_string ~elpi loc s in - if db_exists db then accumulate_to_db db [new_ast] - else CErrors.user_err - Pp.(str "Db " ++ pr_qualified_name db ++ str" not found") - let loc_merge l1 l2 = try Loc.merge l1 l2 with Failure _ -> l1 @@ -476,8 +672,11 @@ let cache_program (nature,p,p_str) = (Extend.TUentry (Genarg.get_arg_tag Coq_elpi_arg_syntax.wit_elpi_loc), Vernacextend.TyNil))))) - (fun loc0 args loc1 ?loc ~atts () -> Vernacextend.vtdefault (fun () -> - run_program (Option.default (loc_merge loc0 loc1) loc) p ~atts args)) + (fun loc0 args loc1 ?loc ~atts () -> + let loc = Option.default (loc_merge loc0 loc1) loc in + let syndata = Synterp.run_program loc p ~atts args in + Vernacextend.vtdefault (fun () -> + Interp.run_program loc p ~atts ~syndata args)) in Egramml.extend_vernac_command_grammar ~undoable:true ext @@ -487,7 +686,7 @@ let cache_program (nature,p,p_str) = CErrors.user_err Pp.(str "elpi: Only commands and tactics can be exported") let subst_program = function - | _, (Command _, _, _) -> CErrors.user_err Pp.(str"elpi: No functors yet") + | _, (Command _,_,_) -> CErrors.user_err Pp.(str"elpi: No functors yet") | _, (Tactic,_,_ as x) -> x | _, (Program _,_,_) -> assert false @@ -499,16 +698,6 @@ let in_exported_program : nature * qualified_name * string -> Libobject.obj = let export_command p = let p_str = String.concat "." p in - let nature = get_nature p in + let nature = Coq_elpi_programs.Synterp.get_nature p in Lib.add_leaf (in_exported_program (nature,p,p_str)) -let skip ~atts:(skip,only) f x = - let m rex = Str.string_match rex Coq_config.version 0 in - let exec = - match skip, only with - | None, None -> true - | Some skip, None -> not (List.exists m skip) - | None, Some keep -> List.exists m keep - | Some _, Some _ -> CErrors.user_err Pp.(str "Attributes #[skip] and #[only] cannot be used at the same time") - in - if exec then f x else () diff --git a/src/coq_elpi_vernacular.mli b/src/coq_elpi_vernacular.mli index 64a132d82..8f8b90e66 100644 --- a/src/coq_elpi_vernacular.mli +++ b/src/coq_elpi_vernacular.mli @@ -4,65 +4,69 @@ open Coq_elpi_utils open Coq_elpi_programs +open Coq_elpi_arg_HOAS -module Synterp : sig -val create_program : ?raw_args:bool -> program_name -> unit -val create_command : ?raw_args:bool -> program_name -> unit -val create_tactic : program_name -> unit -val create_db : program_name -> unit -end +val atts2impl : +Elpi.API.Ast.Loc.t -> depth:int -> Elpi.API.State.t -> Attributes.vernac_flags -> + Elpi.API.Data.term -> Elpi.API.State.t * Elpi.API.Data.term + +module type Common = sig + val typecheck_program : ?program:qualified_name -> unit -> unit + + val get_and_compile : + qualified_name -> (Elpi.API.Compile.program * bool) option + val run : static_check:bool -> + Elpi.API.Compile.program -> + [ `Ast of (Elpi.API.State.t -> Elpi.API.State.t) * Elpi.API.Ast.query + | `Fun of + depth:int -> + Elpi.API.State.t -> + Elpi.API.State.t * + (Elpi.API.Ast.Loc.t * Elpi.API.Data.term) * + Elpi.API.Conversion.extra_goals ] -> + unit Elpi.API.Execute.outcome + + val accumulate_files : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> string list -> unit + val accumulate_extra_deps : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> Names.Id.t list -> unit + val accumulate_string : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> unit + val accumulate_db : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> ?program:qualified_name -> qualified_name -> unit + val accumulate_to_db : atts:((Str.regexp list option * Str.regexp list option) * phase option) -> qualified_name -> Elpi.API.Ast.Loc.t * string -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit -val create_program : program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit -val create_command : program_name -> unit -val create_tactic : program_name -> unit -val create_db : program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit + (* Setup *) + val load_checker : string -> unit + val load_printer : string -> unit + val load_tactic : string -> unit + val load_command : string -> unit + + val debug : atts:phase option -> string list -> unit + val trace : atts:phase option -> int -> int -> string list -> string list -> unit + val trace_browser : atts:phase option -> string list -> unit + val bound_steps : atts:phase option -> int -> unit + val print : atts:phase option -> qualified_name -> string list -> unit -val typecheck_program : ?program:qualified_name -> unit -> unit + val create_program : atts:bool option -> program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit + val create_command : atts:bool option -> program_name -> unit + val create_tactic : program_name -> unit + val create_db : atts:phase option -> program_name -> init:(Elpi.API.Ast.Loc.t * string) -> unit -val accumulate_files : ?program:qualified_name -> string list -> unit -val accumulate_extra_deps : ?program:qualified_name -> Names.Id.t list -> unit -val accumulate_string : ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> unit -val accumulate_db : ?program:qualified_name -> qualified_name -> unit +end -val accumulate_to_db : qualified_name -> Elpi.API.Ast.Loc.t * string -> Names.Id.t list -> scope:Coq_elpi_utils.clause_scope -> unit +module Synterp : sig include Common +val run_program : Loc.t -> qualified_name -> atts:Attributes.vernac_flags -> Cmd.raw list -> ((target:Elpi.API.State.t -> depth:int -> (Elpi.API.Data.term,string) Stdlib.Result.t) * Coq_elpi_builtins_synterp.SynterpAction.t list) option +val run_in_program : ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> Coq_elpi_builtins_synterp.SynterpAction.t list option -val skip : atts:(Str.regexp list option * Str.regexp list option) -> ('a -> unit) -> 'a -> unit +end +module Interp : sig include Common +val run_program : Loc.t -> qualified_name -> atts:Attributes.vernac_flags -> syndata:((target:Elpi.API.State.t -> depth:int -> (Elpi.API.Data.term,string) Stdlib.Result.t) * Coq_elpi_builtins_synterp.SynterpAction.t list) option -> Cmd.raw list -> unit +val run_in_program : ?program:qualified_name -> syndata:Coq_elpi_builtins_synterp.SynterpAction.t list option -> Elpi.API.Ast.Loc.t * string -> unit +end -(* Setup *) -val load_checker : string -> unit -val load_printer : string -> unit -val load_tactic : string -> unit -val load_command : string -> unit val document_builtins : unit -> unit -(* Debug *) -val debug : string list -> unit -val trace : int -> int -> string list -> string list -> unit -val trace_browser : string list -> unit -val bound_steps : int -> unit -val print : qualified_name -> string list -> unit -open Coq_elpi_arg_HOAS -val run_program : Loc.t -> qualified_name -> atts:Attributes.vernac_flags -> Cmd.raw list -> unit -val run_in_program : ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> unit val run_tactic : Loc.t -> qualified_name -> atts:Attributes.vernac_flags -> Geninterp.interp_sign -> Tac.top list -> unit Proofview.tactic val run_in_tactic : ?program:qualified_name -> Elpi.API.Ast.Loc.t * string -> Geninterp.interp_sign -> unit Proofview.tactic +(* move to synterp *) val export_command : qualified_name -> unit - -val atts2impl : - Elpi.API.Ast.Loc.t -> depth:int -> Elpi.API.State.t -> Attributes.vernac_flags -> - Elpi.API.Data.term -> Elpi.API.State.t * Elpi.API.Data.term -val get_and_compile : - qualified_name -> Elpi.API.Compile.program * bool -val run : static_check:bool -> - Elpi.API.Compile.program -> - [ `Ast of Elpi.API.Ast.query - | `Fun of - depth:int -> - Elpi.API.State.t -> - Elpi.API.State.t * - (Elpi.API.Ast.Loc.t * Elpi.API.Data.term) * - Elpi.API.Conversion.extra_goals ] -> - unit Elpi.API.Execute.outcome \ No newline at end of file diff --git a/src/coq_elpi_vernacular_syntax.mlg b/src/coq_elpi_vernacular_syntax.mlg index 4c7c8e4fc..42c36ab81 100644 --- a/src/coq_elpi_vernacular_syntax.mlg +++ b/src/coq_elpi_vernacular_syntax.mlg @@ -51,7 +51,7 @@ let rec inlogpath q = function | x :: xs -> ("coq://" ^ Libnames.string_of_qualid q ^ "/" ^ x) :: inlogpath q xs let warning_legacy_accumulate_gen = - CWarnings.create ~name:"elpi.accumulate-syntax" ~category:Coq_elpi_utils.elpi_depr_cat (fun has_from -> + CWarnings.create ~default:CWarnings.AsError ~name:"elpi.accumulate-syntax" ~category:Coq_elpi_utils.elpi_depr_cat (fun has_from -> if has_from then Pp.strbrk "The syntax 'Elpi Accumulate File \"path\"' is deprecated, use 'Elpi Accumulate File \"path\" From logpath'" else Pp.strbrk "The syntax 'Elpi Accumulate File \"path\" From logpath' is deprecated, use 'From logpath Extra Dependency \"path\" as name. Elpi Accumulate File name.'") @@ -98,147 +98,262 @@ END (* Syntax **************************************************************** *) VERNAC COMMAND EXTEND Elpi CLASSIFIED AS SIDEFF -| #[ atts = skip_attribute ] [ "Elpi" "Accumulate" "File" ne_ident_list(ids) ] -> { - EV.skip ~atts EV.accumulate_extra_deps ids } -| #[ atts = skip_attribute ] [ "Elpi" "Accumulate" "Files" ne_ident_list(ids) ] -> { - EV.skip ~atts EV.accumulate_extra_deps ids } -| #[ atts = skip_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" ne_ident_list(ids) ] -> { - EV.skip ~atts (EV.accumulate_extra_deps ~program:(snd p)) ids } -| #[ atts = skip_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" ne_ident_list(ids) ] -> { - EV.skip ~atts (EV.accumulate_extra_deps ~program:(snd p)) ids } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "File" ne_ident_list(ids) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_extra_deps ~atts ids; + atts + } -> { + EV.Interp.accumulate_extra_deps ~atts ids } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Files" ne_ident_list(ids) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_extra_deps ~atts ids; + atts + } -> { + EV.Interp.accumulate_extra_deps ~atts ids } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" ne_ident_list(ids) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_extra_deps ~atts ~program:(snd p) ids; + atts + } -> { + EV.Interp.accumulate_extra_deps ~atts ~program:(snd p) ids } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" ne_ident_list(ids) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_extra_deps ~atts ~program:(snd p) ids; + atts + } -> { + EV.Interp.accumulate_extra_deps ~atts ~program:(snd p) ids } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "File" string_list(s) ] -> { warning_legacy_accumulate (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_files s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "File" string_list(s) "From" global(g) ] -> { warning_legacy_accumulate2 (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_files (inlogpath g s) } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Files" string_list(s) ] -> { warning_legacy_accumulate (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_files s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Files" string_list(s) "From" global(g) ] -> { warning_legacy_accumulate2 (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_files (inlogpath g s) } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" elpi_string(s) ] -> { - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_string s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" string_list(s) ] -> { warning_legacy_accumulate (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_files ~program:(snd p)) s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" string_list(s) "From" global(g) ] -> { warning_legacy_accumulate2 (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_files ~program:(snd p)) (inlogpath g s) } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" string_list(s) ] -> { warning_legacy_accumulate (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_files ~program:(snd p)) s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" string_list(s) "From" global(g) ] -> { warning_legacy_accumulate2 (); - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_files ~program:(snd p)) (inlogpath g s) } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) elpi_string(s) ] -> { - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_string ~program:(snd p)) s } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Db" qualified_name(d) ] -> { - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts EV.accumulate_db (snd d) } -| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Db" qualified_name(d) ] -> { - let atts = validate_attributes skip_attribute atts in - EV.skip ~atts (EV.accumulate_db ~program:(snd p)) (snd d) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "File" string_list(s) ] SYNTERP AS atts { warning_legacy_accumulate (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts s; + atts + } -> { + EV.Interp.accumulate_files ~atts s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "File" string_list(s) "From" global(g) ] SYNTERP AS atts { warning_legacy_accumulate2 (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts (inlogpath g s); + atts + } -> { + EV.Interp.accumulate_files ~atts (inlogpath g s) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Files" string_list(s) ] SYNTERP AS atts { warning_legacy_accumulate (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts s; + atts + } -> { + EV.Interp.accumulate_files ~atts s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Files" string_list(s) "From" global(g) ] SYNTERP AS atts { warning_legacy_accumulate2 (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts (inlogpath g s); + atts + } -> { + EV.Interp.accumulate_files ~atts (inlogpath g s) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" elpi_string(s) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_string ~atts s; + atts + } -> { + EV.Interp.accumulate_string ~atts s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" string_list(s) ] SYNTERP AS atts { warning_legacy_accumulate (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts ~program:(snd p) s; + atts + } -> { + EV.Interp.accumulate_files ~atts ~program:(snd p) s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "File" string_list(s) "From" global(g) ] SYNTERP AS atts { warning_legacy_accumulate2 (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts ~program:(snd p) (inlogpath g s); + atts + } -> { + EV.Interp.accumulate_files ~atts ~program:(snd p) (inlogpath g s) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" string_list(s) ] SYNTERP AS atts { warning_legacy_accumulate (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts ~program:(snd p) s; + atts + } -> { + EV.Interp.accumulate_files ~atts ~program:(snd p) s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Files" string_list(s) "From" global(g) ] SYNTERP AS atts { warning_legacy_accumulate2 (); + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_files ~atts ~program:(snd p) (inlogpath g s); + atts + } -> { + EV.Interp.accumulate_files ~atts ~program:(snd p) (inlogpath g s) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) elpi_string(s) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_string ~atts ~program:(snd p) s; + atts + } -> { + EV.Interp.accumulate_string ~atts ~program:(snd p) s } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" "Db" qualified_name(d) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_db ~atts (snd d); + atts + } -> { + EV.Interp.accumulate_db ~atts (snd d) } +| #[ atts = any_attribute ] [ "Elpi" "Accumulate" qualified_name(p) "Db" qualified_name(d) ] SYNTERP AS atts { + let atts = validate_attributes skip_and_synterp_attributes atts in + EV.Synterp.accumulate_db ~atts ~program:(snd p) (snd d); + atts + } -> { + EV.Interp.accumulate_db ~atts ~program:(snd p) (snd d) } -| #[ atts = any_attribute ] [ "Elpi" "Debug" string_list(s) ] -> { - let () = ignore_unknown_attributes atts in - EV.debug s } -| #[ atts = any_attribute ] [ "Elpi" "Trace" "Browser" string_list(s) ] -> { - let () = ignore_unknown_attributes atts in - EV.trace_browser s } -| #[ atts = any_attribute ] [ "Elpi" "Trace" string_list(s) ] -> { - let () = ignore_unknown_attributes atts in - EV.trace 1 max_int s [] } -| #[ atts = any_attribute ] [ "Elpi" "Trace" string_list(s) "/" string_list(o) ] -> { - let () = ignore_unknown_attributes atts in - EV.trace 1 max_int s o } -| #[ atts = any_attribute ] [ "Elpi" "Trace" int(start) int(stop) string_list(s) ] -> { - let () = ignore_unknown_attributes atts in - EV.trace start stop s [] } -| #[ atts = any_attribute ] [ "Elpi" "Trace" int(start) int(stop) string_list(s) "/" string_list(o) ] -> { - let () = ignore_unknown_attributes atts in - EV.trace start stop s o } +| #[ atts = any_attribute ] [ "Elpi" "Debug" string_list(s) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.debug ~atts s; + atts + } -> { + EV.Interp.debug ~atts s } +| #[ atts = any_attribute ] [ "Elpi" "Trace" "Browser" string_list(s) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace_browser ~atts s; + atts + } -> { + EV.Interp.trace_browser ~atts s } +| #[ atts = any_attribute ] [ "Elpi" "Trace" string_list(s) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace ~atts 1 max_int s []; + atts + } -> { + EV.Interp.trace ~atts 1 max_int s [] } +| #[ atts = any_attribute ] [ "Elpi" "Trace" string_list(s) "/" string_list(o) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace ~atts 1 max_int s o; + atts + } -> { + EV.Interp.trace ~atts 1 max_int s o } +| #[ atts = any_attribute ] [ "Elpi" "Trace" int(start) int(stop) string_list(s) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace ~atts start stop s []; + atts + } -> { + EV.Interp.trace ~atts start stop s [] } +| #[ atts = any_attribute ] [ "Elpi" "Trace" int(start) int(stop) string_list(s) "/" string_list(o) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace ~atts start stop s o; + atts + } -> { + EV.Interp.trace ~atts start stop s o } -| #[ atts = any_attribute ] [ "Elpi" "Trace" "Off" ] -> { - let () = ignore_unknown_attributes atts in - EV.trace 0 0 [] [] } -| #[ atts = any_attribute ] [ "Elpi" "Bound" "Steps" int(steps) ] -> { - let () = ignore_unknown_attributes atts in - EV.bound_steps steps } +| #[ atts = any_attribute ] [ "Elpi" "Trace" "Off" ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.trace ~atts 0 0 [] []; + atts + } -> { + EV.Interp.trace ~atts 0 0 [] [] } +| #[ atts = any_attribute ] [ "Elpi" "Bound" "Steps" int(steps) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.bound_steps ~atts steps; + atts + } -> { + EV.Interp.bound_steps ~atts steps } -| #[ atts = any_attribute ] [ "Elpi" "Print" qualified_name(p) string_list(s) ] -> { - let () = ignore_unknown_attributes atts in - EV.print (snd p) s } +| #[ atts = any_attribute ] [ "Elpi" "Print" qualified_name(p) string_list(s) ] SYNTERP AS atts { + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.print ~atts (snd p) s; + atts + } -> { + EV.Interp.print ~atts (snd p) s } -| #[ atts = any_attribute ] [ "Elpi" "Program" qualified_name(p) elpi_string(s) ] SYNTERP AS _ +| #[ atts = any_attribute ] [ "Elpi" "Program" qualified_name(p) elpi_string(s) ] SYNTERP AS atts { - let raw_args = validate_attributes raw_args_attribute atts in - EV.Synterp.create_program ?raw_args p + let atts = validate_attributes raw_args_attributes atts in + EV.Synterp.create_program ~atts ~init:s p; + atts } -> { - EV.create_program p ~init:s + EV.Interp.create_program ~atts ~init:s p } -| #[ atts = any_attribute ] [ "Elpi" "Command" qualified_name(p) ] SYNTERP AS _ +| #[ atts = any_attribute ] [ "Elpi" "Command" qualified_name(p) ] SYNTERP AS atts { - let raw_args = validate_attributes raw_args_attribute atts in - EV.Synterp.create_command ?raw_args p + let atts = validate_attributes raw_args_attributes atts in + EV.Synterp.create_command ~atts p; + atts } -> { - EV.create_command p } + EV.Interp.create_command ~atts p } | #[ atts = any_attribute ] [ "Elpi" "Tactic" qualified_name(p) ] SYNTERP AS _ { let () = ignore_unknown_attributes atts in EV.Synterp.create_tactic p } -> { - EV.create_tactic p } -| #[ atts = any_attribute ] [ "Elpi" "Db" qualified_name(d) elpi_string(s) ] SYNTERP AS _ + EV.Interp.create_tactic p } +| #[ atts = any_attribute ] [ "Elpi" "Db" qualified_name(d) elpi_string(s) ] SYNTERP AS atts { - let () = ignore_unknown_attributes atts in - EV.Synterp.create_db d + let atts = validate_attributes synterp_attribute atts in + EV.Synterp.create_db ~atts d ~init:s; + atts } -> { - EV.create_db d ~init:s } + EV.Interp.create_db ~atts d ~init:s } -| #[ atts = any_attribute ] [ "Elpi" "Typecheck" ] -> { +| #[ atts = any_attribute ] [ "Elpi" "Typecheck" ] SYNTERP AS _ + { let () = ignore_unknown_attributes atts in - EV.typecheck_program () } -| #[ atts = any_attribute ] [ "Elpi" "Typecheck" qualified_name(p) ] -> { + EV.Synterp.typecheck_program () + } -> { + EV.Interp.typecheck_program () } +| #[ atts = any_attribute ] [ "Elpi" "Typecheck" qualified_name(p) ] SYNTERP AS _ + { let () = ignore_unknown_attributes atts in - EV.typecheck_program ~program:(snd p) () } + EV.Synterp.typecheck_program ~program:(snd p) () + } -> { + let () = ignore_unknown_attributes atts in + EV.Interp.typecheck_program ~program:(snd p) () } | #[ atts = any_attribute ] [ "Elpi" "Document" "Builtins" ] -> { let () = ignore_unknown_attributes atts in EV.document_builtins () } -| #[ atts = any_attribute ] [ "Elpi" "Checker" string(s) ] -> { +| #[ atts = any_attribute ] [ "Elpi" "Checker" string(syn) string(int) ] SYNTERP AS _ { let () = ignore_unknown_attributes atts in - EV.load_checker s } -| #[ atts = any_attribute ] [ "Elpi" "Printer" string(s) ] -> { + EV.Synterp.load_checker syn + } -> { + EV.Interp.load_checker int } +| #[ atts = any_attribute ] [ "Elpi" "Printer" string(syn) string(int) ] SYNTERP AS _ { let () = ignore_unknown_attributes atts in - EV.load_printer s } -| #[ atts = any_attribute ] [ "Elpi" "Template" "Command" string(s) ] -> { + EV.Synterp.load_printer syn + } -> { + EV.Interp.load_printer int } +| #[ atts = any_attribute ] [ "Elpi" "Template" "Command" string(syn) string(int) ] SYNTERP AS _ { let () = ignore_unknown_attributes atts in - EV.load_command s } -| #[ atts = any_attribute ] [ "Elpi" "Template" "Tactic" string(s) ] -> { + EV.Synterp.load_command syn + } -> { + EV.Interp.load_command int } +| #[ atts = any_attribute ] [ "Elpi" "Template" "Tactic" string(s) ] SYNTERP AS _ { let () = ignore_unknown_attributes atts in - EV.load_tactic s } + () (* no synterp tactics EV.Synterp.load_tactic syn *) + } -> { + EV.Interp.load_tactic s } END VERNAC COMMAND EXTEND ElpiRun CLASSIFIED BY { fun _ -> Vernacextend.(VtSideff ([], VtNow)) } -| #[ atts = any_attribute ] [ "Elpi" "Query" elpi_string(s) ] -> { +| #[ atts = any_attribute ] [ "Elpi" "Query" elpi_string(synterp) elpi_string_opt(interp) ] SYNTERP AS syndata + { let () = ignore_unknown_attributes atts in - EV.run_in_program s } -| #[ atts = any_attribute ] [ "Elpi" "Query" qualified_name(p) elpi_string(s) ] -> { + let synterp, interp = match interp with None -> None, synterp | Some interp -> Some synterp, interp in + match synterp with + | Some synterp -> EV.Synterp.run_in_program synterp + | None -> None + } -> { + let synterp, interp = match interp with None -> None, synterp | Some interp -> Some synterp, interp in + EV.Interp.run_in_program ~syndata interp } +| #[ atts = any_attribute ] [ "Elpi" "Query" qualified_name(p) elpi_string(synterp) elpi_string_opt(interp) ] SYNTERP AS syndata + { let () = ignore_unknown_attributes atts in - EV.run_in_program ~program:(snd p) s } + let synterp, interp = match interp with None -> None, synterp | Some interp -> Some synterp, interp in + match synterp with + | Some synterp -> EV.Synterp.run_in_program synterp + | None -> None + } -> { + let synterp, interp = match interp with None -> None, synterp | Some interp -> Some synterp, interp in + EV.Interp.run_in_program ~program:(snd p) ~syndata interp } | #[ atts = any_attribute ] [ "Elpi" "Export" qualified_name(p) ] => { Vernacextend.(VtSideff ([],VtNow)) } SYNTERP AS _ { EV.export_command (snd p) } -> { let () = ignore_unknown_attributes atts in () } -| #[ atts = any_attribute ] [ "Elpi" qualified_name(p) elpi_cmd_arg_list(args) ] -> { - EV.run_program (fst p) (snd p) ~atts args } +| #[ atts = any_attribute ] [ "Elpi" qualified_name(p) elpi_cmd_arg_list(args) ] SYNTERP AS syndata + { + let () = ignore_unknown_attributes atts in + EV.Synterp.run_program (fst p) (snd p) ~atts args + } -> { + EV.Interp.run_program (fst p) (snd p) ~atts ~syndata args } END TACTIC EXTEND elpi_tac diff --git a/src/elpi_plugin.mlpack b/src/elpi_plugin.mlpack index 7c8692070..6c47f092c 100644 --- a/src/elpi_plugin.mlpack +++ b/src/elpi_plugin.mlpack @@ -6,7 +6,9 @@ Coq_elpi_glob_quotation Coq_elpi_name_quotation Coq_elpi_arg_HOAS Coq_elpi_arg_syntax +Coq_elpi_builtins_arg_HOAS Coq_elpi_builtins_HOAS +Coq_elpi_builtins_synterp Coq_elpi_builtins Coq_elpi_programs Coq_elpi_vernacular diff --git a/tests/test_API.v b/tests/test_API.v index b01122799..eae23466d 100644 --- a/tests/test_API.v +++ b/tests/test_API.v @@ -20,20 +20,27 @@ Elpi Query lp:{{ (****** warnings *************************************) -Set Warnings "-elpi,-category". +Set Warnings "-elpi,-category,+unknown-warning". Elpi Query lp:{{ coq.warn "this is a generic warning". }}. + Elpi Query lp:{{ + coq.say "A", + coq.warning "category" "name" "this is a warning with a name an category". +}} lp:{{ + coq.say "B", coq.warning "category" "name" "this is a warning with a name an category". }}. + Set Warnings "+category". + Elpi Query lp:{{ coq.warning "category" "name" "this is a warning with a name an category". -}}. +}} lp:{{ coq.warning "category" "name" "this is a warning with a name an category". }}. Fail Elpi Query lp:{{ coq.warning "category" "name" "this is another warning with a name an category". -}}. +}} lp:{{ coq.warning "category" "name" "this is a warning with a name an category". }}. Set Warnings "elpi,category". (****** locate **********************************) @@ -83,15 +90,19 @@ Elpi Query lp:{{coq.sort.sup X Y}}. Elpi Db test.db lp:{{type foo string -> prop.}}. Elpi Command test.use.db. Elpi Accumulate Db test.db. +#[synterp] Elpi Accumulate lp:{{ + main [] :- coq.env.begin-module "test_db_accumulate" none, coq.env.end-module _. +}}. Elpi Accumulate lp:{{ main [] :- coq.elpi.accumulate _ "test.db" (clause _ _ (pi x\ foo x :- x = "here")), - coq.env.begin-module "test_db_accumulate" none, + coq.env.begin-module "test_db_accumulate" _, coq.elpi.accumulate current "test.db" (clause _ _ (pi x\ foo x :- x = "there")), coq.env.end-module _. }}. +Elpi Typecheck. Fail Elpi Query lp:{{foo _}}. Elpi test.use.db. @@ -143,18 +154,23 @@ Elpi Db global.db lp:{{ }}. Elpi Command declare. Elpi Accumulate Db global.db. +#[synterp] Elpi Accumulate lp:{{ +main [str "library",_] :- coq.env.begin-module "ClausesL" none, coq.env.end-module _. +main [str "current",_] :- coq.env.begin-module "ClausesC" none, coq.env.end-module _. +main [str "execution-site",_] :- coq.env.begin-module "ClausesE" none, coq.env.end-module _. +}}. Elpi Accumulate lp:{{ main [str "library", str I] :- - coq.env.begin-module "ClausesL" none, + coq.env.begin-module "ClausesL" _, coq.elpi.accumulate library "global.db" (clause _ (before "init") (declared I)), coq.env.end-module _. main [str "current", str I] :- - coq.env.begin-module "ClausesC" none, + coq.env.begin-module "ClausesC" _, coq.elpi.accumulate current "global.db" (clause _ (before "init") (declared I)), coq.env.end-module _. main [str "execution-site", str I] :- - coq.env.begin-module "ClausesE" none, + coq.env.begin-module "ClausesE" _, coq.elpi.accumulate execution-site "global.db" (clause _ (before "init") (declared I)), coq.env.end-module _. diff --git a/tests/test_API2.v b/tests/test_API2.v index 4f721469c..8d28238c2 100644 --- a/tests/test_API2.v +++ b/tests/test_API2.v @@ -171,10 +171,12 @@ Elpi Query lp:{{ #[local] Hint Opaque x : core. Elpi Query lp:{{ + std.do! [coq.env.begin-module "xx" none, coq.env.end-module XX, coq.env.import-module XX ] +}} lp:{{ std.do! [ {{ x }} = global (const C1), coq.hints.opaque C1 "core" @opaque!, - coq.env.begin-module "xx" none, + coq.env.begin-module "xx" _, @local! => coq.hints.set-opaque C1 "core" @transparent!, coq.env.end-module M, coq.hints.opaque C1 "core" @opaque!, @@ -185,10 +187,12 @@ Elpi Query lp:{{ }}. Elpi Query lp:{{ + std.do! [coq.env.begin-module "xx2" none, coq.env.end-module XX, coq.env.import-module XX] +}} lp:{{ std.do! [ {{ x }} = global (const C1), coq.hints.opaque C1 "core" @opaque!, - coq.env.begin-module "xx2" none, + coq.env.begin-module "xx2" _, coq.hints.set-opaque C1 "core" @transparent!, coq.env.end-module M, coq.hints.opaque C1 "core" @opaque!, @@ -201,10 +205,12 @@ Elpi Query lp:{{ #[local] Hint Opaque x : core. Elpi Query lp:{{ + std.do! [coq.env.begin-module "xx3" none, coq.env.end-module _] +}} lp:{{ std.do! [ {{ x }} = global (const C1), coq.hints.opaque C1 "core" @opaque!, - coq.env.begin-module "xx3" none, + coq.env.begin-module "xx3" _, @global! => coq.hints.set-opaque C1 "core" @transparent!, coq.env.end-module M, coq.hints.opaque C1 "core" @transparent!, @@ -245,6 +251,8 @@ Goal 0 = 1. solve [debug eauto with xxx]. Abort. (*Module Type T. Axiom Inline(1) T : Type. End T.*) Elpi Query lp:{{ + std.do! [ coq.env.begin-module-type "T", coq.env.end-module-type _, ] +}} lp:{{ coq.env.begin-module-type "T", @inline-at! 1 => coq.env.add-axiom "T" {{ Type }} _, coq.env.end-module-type _ @@ -253,9 +261,9 @@ Elpi Query lp:{{ (*Module F(P : T). Definition id (a : P.T) := a. End F.*) Elpi Query lp:{{ - - coq.locate-module-type "T" T, - coq.env.begin-module-functor "F" none [pr "P" T], + std.do! [coq.env.begin-module-functor "F" none [pr "P" {coq.locate-module-type "T"}], coq.env.end-module _] +}} lp:{{ + coq.env.begin-module-functor "F" _ _, coq.locate "P.T" GR, coq.env.add-const "id" (fun `a` (global GR) a\ a) _ _ _, coq.env.end-module _. @@ -264,8 +272,10 @@ Elpi Query lp:{{ (*Module X. Definition T := nat. End X.*) Elpi Query lp:{{ + std.do! [coq.env.begin-module "X" none, coq.env.end-module _] +}} lp:{{ - coq.env.begin-module "X" none, + coq.env.begin-module "X" _, coq.env.add-const "T" {{ nat }} _ _ _, coq.env.end-module _. @@ -273,18 +283,18 @@ Elpi Query lp:{{ (*Module G := F X [no inline].*) Elpi Query lp:{{ - coq.locate-module "F" F, - coq.locate-module "X" X, - coq.env.apply-module-functor "G" none F [X] coq.inline.no G, + std.do! [coq.env.apply-module-functor "G" none {coq.locate-module "F"} [{coq.locate-module "X"}] coq.inline.no _] +}} lp:{{ + coq.env.apply-module-functor "G" _ _ _ _ G, coq.say G }}. Print Module G. (*Module H := F X [inline at leve 2].*) Elpi Query lp:{{ - coq.locate-module "F" F, - coq.locate-module "X" X, - coq.env.apply-module-functor "H" none F [X] (coq.inline.at 2) H, + std.do! [coq.env.apply-module-functor "H" none {coq.locate-module "F"} [{coq.locate-module "X"}] (coq.inline.at 2) _] +}} lp:{{ + coq.env.apply-module-functor "H" _ _ _ _ H, coq.say H }}. Print Module H. @@ -303,9 +313,10 @@ Abort. (*Module Type FT (P : T). Axiom idT : P.T -> P.T. End FT.*) Elpi Query lp:{{ + std.do! [coq.env.begin-module-type-functor "FT" [pr "P" {coq.locate-module-type "T"}], coq.env.end-module-type _] +}} lp:{{ - coq.locate-module-type "T" T, - coq.env.begin-module-type-functor "FT" [pr "P" T], + coq.env.begin-module-type-functor "FT" _, coq.locate "P.T" GR, coq.env.add-axiom "idT" (prod _ (global GR) _\ (global GR)) _, coq.env.end-module-type _. @@ -315,18 +326,18 @@ Print Module Type FT. (*Module Type GT := FT X [no inline].*) Elpi Query lp:{{ - coq.locate-module-type "FT" F, - coq.locate-module "X" X, - coq.env.apply-module-type-functor "GT" F [X] coq.inline.no G, + std.do! [coq.env.apply-module-type-functor "GT" {coq.locate-module-type "FT"} [{coq.locate-module "X"}] coq.inline.no _] +}} lp:{{ + coq.env.apply-module-type-functor "GT" _ _ _ G, coq.say G }}. Print Module Type GT. (*Module Type HT := FT X [inline at leve 2].*) Elpi Query lp:{{ - coq.locate-module-type "FT" F, - coq.locate-module "X" X, - coq.env.apply-module-type-functor "HT" F [X] (coq.inline.at 2) H, + std.do! [coq.env.apply-module-type-functor "HT" {coq.locate-module-type "FT"} [{coq.locate-module "X"}] (coq.inline.at 2) _] +}} lp:{{ + coq.env.apply-module-type-functor "HT" _ _ _ H, coq.say H }}. Print Module Type HT. diff --git a/tests/test_API_env.v b/tests/test_API_env.v index 341b55fee..b1bd25e6e 100644 --- a/tests/test_API_env.v +++ b/tests/test_API_env.v @@ -361,6 +361,9 @@ Set Universe Polymorphism. Elpi Query lp:{{ coq.env.begin-module "Test" none, + coq.env.end-module _. +}} lp:{{ + coq.env.begin-module "Test" _, Decl = record "Rec" {{ Type }} "BuildRec" (field [] "f" {{ Type }} (_\ end-record)), @univpoly! => coq.env.add-indt Decl _, diff --git a/tests/test_API_module.v b/tests/test_API_module.v index 8da351cc8..417def432 100644 --- a/tests/test_API_module.v +++ b/tests/test_API_module.v @@ -51,14 +51,23 @@ Elpi Query lp:{{ }}. Elpi Query lp:{{ + std.do! [ + coq.env.begin-module-type "TA", + coq.env.end-module-type Mp_ta, + coq.env.begin-module "A" (some Mp_ta), + coq.env.begin-module "B" none, + coq.env.end-module _, + coq.env.end-module _, + ] +}} lp:{{ std.do! [ coq.env.begin-module-type "TA", coq.env.add-axiom "z" {{nat}} _, coq.env.add-axiom "i" {{Type}} _, coq.env.end-module-type MP_TA, - coq.env.begin-module "A" (some MP_TA), + coq.env.begin-module "A" _, coq.env.add-const "x" {{3}} _ _ _, - coq.env.begin-module "B" none, + coq.env.begin-module "B" _, coq.env.add-const "y" {{3}} _ _ GRy, coq.env.end-module _, coq.env.add-const "z" (global (const GRy)) _ _ _, @@ -76,12 +85,22 @@ Print A.i. Fail Check A.i1_ind. Elpi Query lp:{{ + coq.locate-module-type "TA" MP_TA, + std.do! [ + coq.env.begin-module-type "TF", + coq.env.end-module-type TF, + coq.env.begin-module-functor "F" (some TF) [pr "a" MP_TA, pr "b" MP_TA], + coq.locate-module "a" A, + coq.env.import-module A, + coq.env.end-module _, + ] +}} lp:{{ std.do! [ coq.env.begin-module-type "TF", coq.env.add-axiom "w" {{nat}} _, coq.env.end-module-type MP_TF, coq.locate-module-type "TA" MP_TA, - coq.env.begin-module-functor "F" (some MP_TF) [pr "a" MP_TA, pr "b" MP_TA], + coq.env.begin-module-functor "F" _ _, coq.env.import-module {coq.locate-module "a"}, coq.env.add-const "w" (global {coq.locate "z"}) _ _ _, coq.env.end-module _ @@ -93,16 +112,27 @@ Print B. Print B.w. Elpi Query lp:{{ + coq.locate-module-type "TA" MP_TA, + std.do! [ + coq.env.begin-module-type-functor "TB" [pr "A" MP_TA], + coq.env.end-module-type _, + ] +}} lp:{{ std.do! [ - coq.locate-module-type "TA" MP_TA, - coq.env.begin-module-type-functor "TB" [pr "A" MP_TA], + coq.env.begin-module-type-functor "TB" _, coq.env.end-module-type _ ] }}. Print TB. -Elpi Query lp:{{ +Elpi Query lp:{{ std.do! [ + coq.locate-module "A" A, coq.env.begin-module "IA" none, + coq.env.include-module A _, + coq.env.end-module _, +] +}} lp:{{ + coq.env.begin-module "IA" _, coq.env.include-module {coq.locate-module "A"} _, coq.env.end-module _. }}. @@ -110,13 +140,24 @@ Elpi Query lp:{{ Print IA. Module Tmp. -Elpi Query lp:{{ coq.env.import-module { coq.locate-module "IA" } }}. +Elpi Query +lp:{{ + std.do! [coq.env.import-module { coq.locate-module "IA" }] +}} lp:{{ + coq.env.import-module { coq.locate-module "IA" } +}}. Check i. End Tmp. Elpi Query lp:{{ + std.do! [ + coq.env.begin-module-type "ITA", + coq.env.include-module-type {coq.locate-module-type "TA"} (coq.inline.at 2), + coq.env.end-module-type _, + ] +}} lp:{{ coq.env.begin-module-type "ITA", - coq.env.include-module-type {coq.locate-module-type "TA"} (coq.inline.at 2), + coq.env.include-module-type {coq.locate-module-type "TA"} _, coq.env.end-module-type _. }}. diff --git a/tests/test_API_section.v b/tests/test_API_section.v index cac2c39e2..7cfb32c0c 100644 --- a/tests/test_API_section.v +++ b/tests/test_API_section.v @@ -30,6 +30,8 @@ Check eq_refl : e2 = 3. End SA. Elpi Query lp:{{ + std.do! [ coq.env.begin-section "Foo", coq.env.end-section ] +}} lp:{{ coq.env.begin-section "Foo", coq.env.add-section-variable "x" {{ nat }} X, coq.env.section [X], diff --git a/tests/test_HOAS.v b/tests/test_HOAS.v index 3c660f4f1..889e97352 100644 --- a/tests/test_HOAS.v +++ b/tests/test_HOAS.v @@ -54,7 +54,11 @@ Elpi kwd fun in as 4 end match return => : := { } ; , | "x" 1 H (fun x => match End kwd. +Elpi Command test. + Elpi Query lp:{{ + std.do! [coq.env.begin-section "xxxxx", coq.env.end-section] +}} lp:{{ coq.env.begin-section "xxxxx", coq.univ.new U, T = sort (typ U), @@ -600,6 +604,8 @@ Polymorphic Inductive tree (A : Type) := Set Printing Universes. Print tree. Elpi Query lp:{{ + std.do! [coq.env.begin-module "M" none, coq.env.end-module _] +}} lp:{{ pglobal (indt I) _ = {{ tree }}, coq.env.indt-decl I D, coq.env.begin-module "M" none, diff --git a/tests/test_glob.v b/tests/test_glob.v index 803a732c8..9ee60ea74 100644 --- a/tests/test_glob.v +++ b/tests/test_glob.v @@ -7,6 +7,9 @@ Section A. Variable v : nat. End A. Module N1. End N1. Module M1 := N1. Elpi Command test. +#[synterp] Elpi Accumulate lp:{{ + main _ :- std.do! [ coq.env.begin-section "A", coq.env.end-section, coq.env.begin-module "N2" none, coq.env.end-module _]. +}}. Elpi Accumulate lp:{{ main _ :- coq.env.add-const "d2" {{ 3 }} _ _ _, diff --git a/tests/test_libobject_A.v b/tests/test_libobject_A.v index c5750f723..b3253daba 100644 --- a/tests/test_libobject_A.v +++ b/tests/test_libobject_A.v @@ -1,6 +1,6 @@ From elpi Require Import elpi. -Elpi Db a.db lp:{{ +#[interp] Elpi Db a.db lp:{{ pred a o:term. :name "init" a {{ 0 }}. a {{ 1 }}. diff --git a/tests/test_link_order_import3.ref b/tests/test_link_order_import3.ref index a96138e45..2b3045596 100644 --- a/tests/test_link_order_import3.ref +++ b/tests/test_link_order_import3.ref @@ -116,6 +116,8 @@ std.split-at A0 [A1 | A2] [A1 | A3] A4 :- std.split-at _ _ _ _ :- std.fatal-error split-at run out of list items. std.fold [] A0 _ A0. std.fold [A0 | A1] A2 A3 A4 :- A3 A0 A2 A5 , std.fold A1 A5 A3 A4. +std.fold-right [] A0 _ A0. +std.fold-right [A0 | A1] A2 A3 A4 :- std.fold-right A1 A2 A3 A5 , A3 A0 A5 A4. std.fold2 [] [_ | _] _ _ _ :- std.fatal-error fold2 on lists of different length. std.fold2 [_ | _] [] _ _ _ :- diff --git a/tests/test_synterp.v b/tests/test_synterp.v new file mode 100644 index 000000000..56b77a3c5 --- /dev/null +++ b/tests/test_synterp.v @@ -0,0 +1,147 @@ +From Coq Require Import Bool. +From elpi Require Import elpi. + +Elpi Command foo. +#[synterp] Elpi Accumulate lp:{{ main-synterp X X :- coq.say "synterp" X. }}. +Elpi Accumulate lp:{{ main-interp X Y :- coq.say "interp" X, std.assert! (X = Y) "bug". }}. +Elpi Typecheck. + +Elpi foo X. +Elpi foo 1. +Elpi foo (1). +Elpi foo Definition x (P Q : bool) : nat := 0. +Elpi foo Axiom x (P Q : bool) : nat. +Elpi foo Record x (P Q : bool) := K { f1 : nat; f2 : f1 = f1 }. +Elpi foo Inductive x (P : bool) | (Q : bool) : nat -> Type := K : nat -> x Q 3 | R (w : bool) : x Q 1 . +Elpi foo Context (A : nat) (B : bool). + +Elpi Export foo. + +foo X. +foo 1. +foo (1). +foo Definition x (P Q : bool) : nat := 0. +foo Axiom x (P Q : bool) : nat. +foo Record x (P Q : bool) := K { f1 : nat; f2 : f1 = f1 }. +foo Inductive x (P : bool) | (Q : bool) : nat -> Type := K : nat -> x Q 3 | R (w : bool) : x Q 1 . +foo Context (A : nat) (B : bool). + +Elpi Command start. +#[synterp] Elpi Accumulate lp:{{ + main-synterp [str X] _ :- coq.env.begin-module X none. +}}. +Elpi Accumulate lp:{{ + main [str X] :- coq.env.begin-module X _. +}}. + +Elpi Command stop. +#[synterp] Elpi Accumulate lp:{{ + main-synterp [str X] _ :- coq.env.end-module _. +}}. +Elpi Accumulate lp:{{ + main [str X] :- coq.env.end-module _. +}}. +Elpi start "X". +Definition a:= 1. +About a. +(* End X. bug vernacextend push/pop *) +Elpi stop "X". + + +Elpi Command module. +#[synterp] Elpi Accumulate lp:{{ + main [str X] :- coq.env.begin-module X none, coq.env.end-module _. +}}. +Elpi Accumulate lp:{{ + main [str X] :- coq.env.begin-module X _, coq.env.add-const "a" {{1}} _ _ _, coq.env.end-module _. +}}. +Elpi module "A". + +Print Module X. +Print Module A. + +Elpi Command err. +Elpi Accumulate lp:{{ + main [str X] :- coq.env.begin-module X none, coq.env.end-module _. +}}. +Fail Elpi err "C". + + + + +Elpi Command module2. +#[synterp] Elpi Accumulate lp:{{ + main [str X, str Y] :- + coq.env.begin-module X none, + coq.env.end-module M, + coq.env.import-module M, + coq.env.begin-module Y none, + coq.env.end-module _. +}}. +Elpi Accumulate lp:{{ + main [str X, str Y] :- + coq.env.begin-module X _, + coq.env.add-const "a" {{1}} _ _ _, + coq.env.end-module M, + coq.env.import-module M, + coq.env.begin-module Y _, + coq.env.end-module _. + +}}. +Elpi Typecheck. +Elpi module2 "B" "C". +Check a. + +#[both] Elpi Db acc.db lp:{{ +pred p o:int. +}}. +Elpi Command acc. +#[both] Elpi Accumulate Db acc.db. +#[synterp] Elpi Accumulate lp:{{ + main [int N] :- + @local! => coq.elpi.accumulate _ "acc.db" (clause _ _ (p N)), + coq.env.begin-module "TMP" none, + true + . +}}. +#[interp] Elpi Accumulate lp:{{ + main _ :- + coq.env.begin-module "TMP" none, + true. +}}. +Elpi Export acc. + +Elpi Command pr. +#[both] Elpi Accumulate Db acc.db. +#[synterp] Elpi Accumulate lp:{{ + main [int N] :- + std.findall (p X_) L, coq.say "L=" L, std.assert! (std.length L N) "wtf", + coq.env.end-module _, + true. +}}. +#[interp] Elpi Accumulate lp:{{ + main _ :- + coq.env.end-module _, + true. +}}. +Elpi Export pr. + +acc 1. +pr 1. + +(* ********************************************* *) + +Elpi Command test_data. +#[synterp] Elpi Accumulate lp:{{ + type foo int. + main-synterp _ R :- R = foo. % std.do!. +}}. +#[interp] Elpi Accumulate lp:{{ + type foo int. + main-interp _ R :- std.assert! (std.any->string R "foo") "bug". +}}. +Elpi Typecheck. +Elpi Export test_data. + +test_data. + diff --git a/theories/elpi.v b/theories/elpi.v index 1516bcfaa..3bff41c2f 100644 --- a/theories/elpi.v +++ b/theories/elpi.v @@ -5,9 +5,10 @@ Declare ML Module "coq-elpi.elpi". Elpi Document Builtins. (* Load once and forall these files in this .vo, to ease redistribution *) -Elpi Checker "coq://elpi/coq-elpi-checker.elpi". -Elpi Printer "elpi2html.elpi". (* this one is from elpi *) -Elpi Template Command "coq://elpi/elpi-command-template.elpi". +Elpi Checker "coq://elpi/coq-elpi-checker.elpi" "coq://elpi/coq-elpi-checker.elpi". + +Elpi Printer "elpi2html.elpi" "elpi2html.elpi". (* this one is from elpi *) +Elpi Template Command "coq://elpi/elpi-command-template-synterp.elpi" "coq://elpi/elpi-command-template.elpi". Elpi Template Tactic "coq://elpi/elpi-tactic-template.elpi". (* Special constant used for HOAS of holes, see coq-bultins.elpi *)