diff options
| author | Gaƫtan Gilbert | 2019-05-07 14:34:06 +0200 |
|---|---|---|
| committer | Enrico Tassi | 2019-06-04 13:58:42 +0200 |
| commit | 02fda88fc065577f3f9604477db2e03eb4d5a9b4 (patch) | |
| tree | b64f74e5aa9d83dcb9adbda0ace93a00e4575203 /coqpp/coqpp_main.ml | |
| parent | b8842c3c8d6e6d9d4c19a75453fca9f94de6fa49 (diff) | |
Alternate syntax for ![]: VERNAC EXTEND Foo STATE proof
eg ![proof] becomes STATE proof
This commits still supports the old ![]
so there is redundancy:
~~~
VERNAC EXTEND Foo STATE proof
| ...
VERNAC EXTEND Foo
| ![proof] ...
~~~
with the ![] form being local to the rule and the STATE form
applying to the whole EXTEND except for the rules with a ![].
Diffstat (limited to 'coqpp/coqpp_main.ml')
| -rw-r--r-- | coqpp/coqpp_main.ml | 72 |
1 files changed, 33 insertions, 39 deletions
diff --git a/coqpp/coqpp_main.ml b/coqpp/coqpp_main.ml index 8444f3a58c..ad79599cf4 100644 --- a/coqpp/coqpp_main.ml +++ b/coqpp/coqpp_main.ml @@ -357,41 +357,35 @@ let print_atts_right fmt = function let nota = match atts with [_] -> "" | _ -> "Attributes.Notations." in fprintf fmt "(Attributes.parse %s%a atts)" nota aux atts -let print_body_wrapper fmt r = - match r.vernac_state with - | Some "proof_stack" -> - fprintf fmt "let proof = (%a) ~pstate:st.Vernacstate.proof in { st with Vernacstate.proof }" print_code r.vernac_body - - | Some "open_proof" -> - let push = "Some (Proof_global.push ~ontop:st.Vernacstate.proof pstate)" in - fprintf fmt "let pstate = (%a) in let proof = %s in { st with Vernacstate.proof }" print_code r.vernac_body push - - | Some "maybe_open_proof" -> - let push = "Proof_global.maybe_push ~ontop:st.Vernacstate.proof pstate" in - fprintf fmt "let pstate = (%a) in let proof = %s in { st with Vernacstate.proof }" print_code r.vernac_body push - - | Some "proof" -> - fprintf fmt "let proof = Vernacentries.modify_pstate ~pstate:st.Vernacstate.proof (%a) in { st with Vernacstate.proof }" print_code r.vernac_body - - | Some "proof_opt_query" -> - fprintf fmt "let () = (%a) ~pstate:(Vernacstate.pstate st) in st" print_code r.vernac_body - - | Some "proof_query" -> - fprintf fmt "let () = Vernacentries.with_pstate ~pstate:st.Vernacstate.proof (%a) in st" print_code r.vernac_body - - | None -> - fprintf fmt "let () = %a in st" print_code r.vernac_body - - | Some x -> - fatal ("unsupported state specifier: " ^ x) - -let print_body_fun fmt r = - fprintf fmt "let coqpp_body %a%a ~st = @[%a@] in " - print_binders r.vernac_toks print_atts_left r.vernac_atts print_body_wrapper r - -let print_body fmt r = +let understand_state = function + | "proof_stack" -> "VtModifyProofStack", false + | "maybe_open_proof" -> "VtMaybeOpenProof", true + | "open_proof" -> "VtOpenProof", true + | "proof" -> "VtModifyProof", false + | "proof_opt_query" -> "VtReadProofOpt", false + | "proof_query" -> "VtReadProof", false + | s -> fatal ("unsupported state specifier: " ^ s) + +let print_body_state state fmt r = + let state = match r.vernac_state with Some _ as s -> s | None -> state in + match state with + | None -> fprintf fmt "Vernacentries.VtDefault (fun () -> %a)" print_code r.vernac_body + | Some "CUSTOM" -> print_code fmt r.vernac_body + | Some state -> + let state, unit_wrap = understand_state state in + fprintf fmt "Vernacentries.%s (%s%a)" state (if unit_wrap then "fun () ->" else "") + print_code r.vernac_body + +let print_body_wrapper state fmt r = + fprintf fmt "Vernacentries.interp_functional_vernac (%a)" (print_body_state state) r + +let print_body_fun state fmt r = + fprintf fmt "let coqpp_body %a%a = @[%a@] in " + print_binders r.vernac_toks print_atts_left r.vernac_atts (print_body_wrapper state) r + +let print_body state fmt r = fprintf fmt "@[(%afun %a~atts@ ~st@ -> coqpp_body %a%a ~st)@]" - print_body_fun r print_binders r.vernac_toks + (print_body_fun state) r print_binders r.vernac_toks print_binders r.vernac_toks print_atts_right r.vernac_atts let rec print_sig fmt = function @@ -402,12 +396,12 @@ let rec print_sig fmt = function fprintf fmt "@[Vernacextend.TyNonTerminal (%a, %a)@]" print_symbol symb print_sig rem -let print_rule fmt r = +let print_rule state fmt r = fprintf fmt "Vernacextend.TyML (%b, %a, %a, %a)" - r.vernac_depr print_sig r.vernac_toks print_body r print_rule_classifier r + r.vernac_depr print_sig r.vernac_toks (print_body state) r print_rule_classifier r -let print_rules fmt rules = - print_list fmt (fun fmt r -> fprintf fmt "(%a)" print_rule r) rules +let print_rules state fmt rules = + print_list fmt (fun fmt r -> fprintf fmt "(%a)" (print_rule state) r) rules let print_classifier fmt = function | ClassifDefault -> fprintf fmt "" @@ -426,7 +420,7 @@ let print_ast fmt ext = let pr fmt () = fprintf fmt "Vernacextend.vernac_extend ~command:\"%s\" %a ?entry:%a %a" ext.vernacext_name print_classifier ext.vernacext_class - print_entry ext.vernacext_entry print_rules ext.vernacext_rules + print_entry ext.vernacext_entry (print_rules ext.vernacext_state) ext.vernacext_rules in let () = fprintf fmt "let () = @[%a@]@\n" pr () in () |
