aboutsummaryrefslogtreecommitdiff
path: root/coqpp/coqpp_main.ml
diff options
context:
space:
mode:
authorGaëtan Gilbert2018-10-16 14:42:53 +0200
committerGaëtan Gilbert2018-11-02 13:26:16 +0100
commit0a478031f0213ef74c3649ea1a8d58aa89e54416 (patch)
tree19f2e2e0a71a68448e1708b9e704b7c68bf6f26f /coqpp/coqpp_main.ml
parentc44080b74f4ea1e4b7ae88dfe5a440364bed3fca (diff)
coqpp VERNAC EXTEND uses #[ att = attribute; ] syntax
I think for instance the new code in this diff is cleaner and more systematic: ~~~diff VERNAC COMMAND EXTEND VernacDeclareTacticDefinition -| [ "Ltac" ne_ltac_tacdef_body_list_sep(l, "with") ] => { +| #[ deprecation; locality; ] [ "Ltac" ne_ltac_tacdef_body_list_sep(l, "with") ] => { VtSideff (List.map (function | TacticDefinition ({CAst.v=r},_) -> r | TacticRedefinition (qid,_) -> qualid_basename qid) l), VtLater } -> { - let deprecation, locality = Attributes.(parse Notations.(deprecation ++ locality) atts) in Tacentries.register_ltac (Locality.make_module_locality locality) ?deprecation l; } END ~~~
Diffstat (limited to 'coqpp/coqpp_main.ml')
-rw-r--r--coqpp/coqpp_main.ml47
1 files changed, 45 insertions, 2 deletions
diff --git a/coqpp/coqpp_main.ml b/coqpp/coqpp_main.ml
index 5314806c24..7cecff9d75 100644
--- a/coqpp/coqpp_main.ml
+++ b/coqpp/coqpp_main.ml
@@ -309,9 +309,52 @@ let print_rule_classifier fmt r = match r.vernac_class with
else
fprintf fmt "Some @[(fun %a-> %a)@]" print_binders r.vernac_toks print_code f
+(* let print_atts fmt = function *)
+(* | None -> fprintf fmt "@[let () = Attributes.unsupported_attributes atts in@] " *)
+(* | Some atts -> *)
+(* let rec print_left fmt = function *)
+(* | [] -> assert false *)
+(* | [x,_] -> fprintf fmt "%s" x *)
+(* | (x,_) :: rem -> fprintf fmt "(%s, %a)" x print_left rem *)
+(* in *)
+(* let rec print_right fmt = function *)
+(* | [] -> assert false *)
+(* | [_,y] -> fprintf fmt "%s" y *)
+(* | (_,y) :: rem -> fprintf fmt "(%s ++ %a)" y print_right rem *)
+(* in *)
+(* let nota = match atts with [_] -> "" | _ -> "Attributes.Notations." in *)
+(* fprintf fmt "@[let %a = Attributes.parse %s(%a) atts in@] " *)
+(* print_left atts nota print_right atts *)
+
+let print_atts_left fmt = function
+ | None -> fprintf fmt "()"
+ | Some atts ->
+ let rec aux fmt = function
+ | [] -> assert false
+ | [x,_] -> fprintf fmt "%s" x
+ | (x,_) :: rem -> fprintf fmt "(%s, %a)" x aux rem
+ in
+ aux fmt atts
+
+let print_atts_right fmt = function
+ | None -> fprintf fmt "(Attributes.unsupported_attributes atts)"
+ | Some atts ->
+ let rec aux fmt = function
+ | [] -> assert false
+ | [_,y] -> fprintf fmt "%s" y
+ | (_,y) :: rem -> fprintf fmt "(%s ++ %a)" y aux rem
+ in
+ let nota = match atts with [_] -> "" | _ -> "Attributes.Notations." in
+ fprintf fmt "(Attributes.parse %s%a atts)" nota aux atts
+
+let print_body_fun fmt r =
+ fprintf fmt "let coqpp_body %a%a ~st = let () = %a in st in "
+ print_binders r.vernac_toks print_atts_left r.vernac_atts print_code r.vernac_body
+
let print_body fmt r =
- fprintf fmt "@[(fun %a~atts@ ~st@ -> let () = %a in st)@]"
- print_binders r.vernac_toks print_code r.vernac_body
+ fprintf fmt "@[(%afun %a~atts@ ~st@ -> coqpp_body %a%a ~st)@]"
+ print_body_fun r print_binders r.vernac_toks
+ print_binders r.vernac_toks print_atts_right r.vernac_atts
let rec print_sig fmt = function
| [] -> fprintf fmt "@[Vernacentries.TyNil@]"