From 4da233a9685cd193a84def037ec18a27c9225dce Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 12 Oct 2018 09:22:56 +0200 Subject: Port remaining EXTEND ml4 files to coqpp. Almost all of ml4 were removed in the process. The only remaining files are in the test-suite and probably need a bit of fiddling with coq_makefile, and there only two really remaning ml4 files containing code. --- plugins/extraction/g_extraction.ml4 | 167 --- plugins/extraction/g_extraction.mlg | 183 +++ plugins/firstorder/g_ground.ml4 | 164 -- plugins/firstorder/g_ground.mlg | 173 +++ plugins/funind/g_indfun.ml4 | 236 --- plugins/funind/g_indfun.mlg | 271 ++++ plugins/ltac/extraargs.ml4 | 310 ---- plugins/ltac/extraargs.mlg | 353 +++++ plugins/ltac/extratactics.ml4 | 1065 ------------- plugins/ltac/extratactics.mlg | 1160 +++++++++++++++ plugins/ltac/g_auto.ml4 | 226 --- plugins/ltac/g_auto.mlg | 249 ++++ plugins/ltac/g_class.ml4 | 117 -- plugins/ltac/g_class.mlg | 137 ++ plugins/ltac/g_ltac.ml4 | 528 ------- plugins/ltac/g_ltac.mlg | 560 +++++++ plugins/ltac/g_obligations.ml4 | 161 -- plugins/ltac/g_obligations.mlg | 173 +++ plugins/ltac/g_rewrite.ml4 | 297 ---- plugins/ltac/g_rewrite.mlg | 318 ++++ plugins/ltac/profile_ltac_tactics.ml4 | 74 - plugins/ltac/profile_ltac_tactics.mlg | 82 + plugins/setoid_ring/g_newring.ml4 | 127 -- plugins/setoid_ring/g_newring.mlg | 147 ++ plugins/ssr/ssrparser.ml4 | 2348 ----------------------------- plugins/ssr/ssrparser.mlg | 2628 +++++++++++++++++++++++++++++++++ plugins/ssr/ssrvernac.ml4 | 625 -------- plugins/ssr/ssrvernac.mlg | 675 +++++++++ plugins/ssrmatching/g_ssrmatching.ml4 | 100 -- plugins/ssrmatching/g_ssrmatching.mlg | 120 ++ plugins/syntax/g_numeral.ml4 | 38 - plugins/syntax/g_numeral.mlg | 42 + 32 files changed, 7271 insertions(+), 6583 deletions(-) delete mode 100644 plugins/extraction/g_extraction.ml4 create mode 100644 plugins/extraction/g_extraction.mlg delete mode 100644 plugins/firstorder/g_ground.ml4 create mode 100644 plugins/firstorder/g_ground.mlg delete mode 100644 plugins/funind/g_indfun.ml4 create mode 100644 plugins/funind/g_indfun.mlg delete mode 100644 plugins/ltac/extraargs.ml4 create mode 100644 plugins/ltac/extraargs.mlg delete mode 100644 plugins/ltac/extratactics.ml4 create mode 100644 plugins/ltac/extratactics.mlg delete mode 100644 plugins/ltac/g_auto.ml4 create mode 100644 plugins/ltac/g_auto.mlg delete mode 100644 plugins/ltac/g_class.ml4 create mode 100644 plugins/ltac/g_class.mlg delete mode 100644 plugins/ltac/g_ltac.ml4 create mode 100644 plugins/ltac/g_ltac.mlg delete mode 100644 plugins/ltac/g_obligations.ml4 create mode 100644 plugins/ltac/g_obligations.mlg delete mode 100644 plugins/ltac/g_rewrite.ml4 create mode 100644 plugins/ltac/g_rewrite.mlg delete mode 100644 plugins/ltac/profile_ltac_tactics.ml4 create mode 100644 plugins/ltac/profile_ltac_tactics.mlg delete mode 100644 plugins/setoid_ring/g_newring.ml4 create mode 100644 plugins/setoid_ring/g_newring.mlg delete mode 100644 plugins/ssr/ssrparser.ml4 create mode 100644 plugins/ssr/ssrparser.mlg delete mode 100644 plugins/ssr/ssrvernac.ml4 create mode 100644 plugins/ssr/ssrvernac.mlg delete mode 100644 plugins/ssrmatching/g_ssrmatching.ml4 create mode 100644 plugins/ssrmatching/g_ssrmatching.mlg delete mode 100644 plugins/syntax/g_numeral.ml4 create mode 100644 plugins/syntax/g_numeral.mlg (limited to 'plugins') diff --git a/plugins/extraction/g_extraction.ml4 b/plugins/extraction/g_extraction.ml4 deleted file mode 100644 index 370d3c248f..0000000000 --- a/plugins/extraction/g_extraction.ml4 +++ /dev/null @@ -1,167 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* [ id ] -| [ string(s) ] -> [ s ] -END - -let pr_int_or_id _ _ _ = function - | ArgInt i -> int i - | ArgId id -> Id.print id - -ARGUMENT EXTEND int_or_id - PRINTED BY pr_int_or_id -| [ preident(id) ] -> [ ArgId (Id.of_string id) ] -| [ integer(i) ] -> [ ArgInt i ] -END - -let pr_language = function - | Ocaml -> str "OCaml" - | Haskell -> str "Haskell" - | Scheme -> str "Scheme" - | JSON -> str "JSON" - -let warn_deprecated_ocaml_spelling = - CWarnings.create ~name:"deprecated-ocaml-spelling" ~category:"deprecated" - (fun () -> - strbrk ("The spelling \"OCaml\" should be used instead of \"Ocaml\".")) - -VERNAC ARGUMENT EXTEND language -PRINTED BY pr_language -| [ "Ocaml" ] -> [ let _ = warn_deprecated_ocaml_spelling () in Ocaml ] -| [ "OCaml" ] -> [ Ocaml ] -| [ "Haskell" ] -> [ Haskell ] -| [ "Scheme" ] -> [ Scheme ] -| [ "JSON" ] -> [ JSON ] -END - -(* Extraction commands *) - -VERNAC COMMAND EXTEND Extraction CLASSIFIED AS QUERY -(* Extraction in the Coq toplevel *) -| [ "Extraction" global(x) ] -> [ simple_extraction x ] -| [ "Recursive" "Extraction" ne_global_list(l) ] -> [ full_extraction None l ] - -(* Monolithic extraction to a file *) -| [ "Extraction" string(f) ne_global_list(l) ] - -> [ full_extraction (Some f) l ] - -(* Extraction to a temporary file and OCaml compilation *) -| [ "Extraction" "TestCompile" ne_global_list(l) ] - -> [ extract_and_compile l ] -END - -VERNAC COMMAND EXTEND SeparateExtraction CLASSIFIED AS QUERY -(* Same, with content splitted in several files *) -| [ "Separate" "Extraction" ne_global_list(l) ] - -> [ separate_extraction l ] -END - -(* Modular extraction (one Coq library = one ML module) *) -VERNAC COMMAND EXTEND ExtractionLibrary CLASSIFIED AS QUERY -| [ "Extraction" "Library" ident(m) ] - -> [ extraction_library false m ] -END - -VERNAC COMMAND EXTEND RecursiveExtractionLibrary CLASSIFIED AS QUERY -| [ "Recursive" "Extraction" "Library" ident(m) ] - -> [ extraction_library true m ] -END - -(* Target Language *) -VERNAC COMMAND EXTEND ExtractionLanguage CLASSIFIED AS SIDEFF -| [ "Extraction" "Language" language(l) ] - -> [ extraction_language l ] -END - -VERNAC COMMAND EXTEND ExtractionInline CLASSIFIED AS SIDEFF -(* Custom inlining directives *) -| [ "Extraction" "Inline" ne_global_list(l) ] - -> [ extraction_inline true l ] -END - -VERNAC COMMAND EXTEND ExtractionNoInline CLASSIFIED AS SIDEFF -| [ "Extraction" "NoInline" ne_global_list(l) ] - -> [ extraction_inline false l ] -END - -VERNAC COMMAND EXTEND PrintExtractionInline CLASSIFIED AS QUERY -| [ "Print" "Extraction" "Inline" ] - -> [Feedback. msg_info (print_extraction_inline ()) ] -END - -VERNAC COMMAND EXTEND ResetExtractionInline CLASSIFIED AS SIDEFF -| [ "Reset" "Extraction" "Inline" ] - -> [ reset_extraction_inline () ] -END - -VERNAC COMMAND EXTEND ExtractionImplicit CLASSIFIED AS SIDEFF -(* Custom implicit arguments of some csts/inds/constructors *) -| [ "Extraction" "Implicit" global(r) "[" int_or_id_list(l) "]" ] - -> [ extraction_implicit r l ] -END - -VERNAC COMMAND EXTEND ExtractionBlacklist CLASSIFIED AS SIDEFF -(* Force Extraction to not use some filenames *) -| [ "Extraction" "Blacklist" ne_ident_list(l) ] - -> [ extraction_blacklist l ] -END - -VERNAC COMMAND EXTEND PrintExtractionBlacklist CLASSIFIED AS QUERY -| [ "Print" "Extraction" "Blacklist" ] - -> [ Feedback.msg_info (print_extraction_blacklist ()) ] -END - -VERNAC COMMAND EXTEND ResetExtractionBlacklist CLASSIFIED AS SIDEFF -| [ "Reset" "Extraction" "Blacklist" ] - -> [ reset_extraction_blacklist () ] -END - - -(* Overriding of a Coq object by an ML one *) -VERNAC COMMAND EXTEND ExtractionConstant CLASSIFIED AS SIDEFF -| [ "Extract" "Constant" global(x) string_list(idl) "=>" mlname(y) ] - -> [ extract_constant_inline false x idl y ] -END - -VERNAC COMMAND EXTEND ExtractionInlinedConstant CLASSIFIED AS SIDEFF -| [ "Extract" "Inlined" "Constant" global(x) "=>" mlname(y) ] - -> [ extract_constant_inline true x [] y ] -END - -VERNAC COMMAND EXTEND ExtractionInductive CLASSIFIED AS SIDEFF -| [ "Extract" "Inductive" global(x) "=>" - mlname(id) "[" mlname_list(idl) "]" string_opt(o) ] - -> [ extract_inductive x id idl o ] -END -(* Show the extraction of the current proof *) - -VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY -| [ "Show" "Extraction" ] - -> [ show_extraction () ] -END diff --git a/plugins/extraction/g_extraction.mlg b/plugins/extraction/g_extraction.mlg new file mode 100644 index 0000000000..1445dffefa --- /dev/null +++ b/plugins/extraction/g_extraction.mlg @@ -0,0 +1,183 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* { id } +| [ string(s) ] -> { s } +END + +{ + +let pr_int_or_id _ _ _ = function + | ArgInt i -> int i + | ArgId id -> Id.print id + +} + +ARGUMENT EXTEND int_or_id + PRINTED BY { pr_int_or_id } +| [ preident(id) ] -> { ArgId (Id.of_string id) } +| [ integer(i) ] -> { ArgInt i } +END + +{ + +let pr_language = function + | Ocaml -> str "OCaml" + | Haskell -> str "Haskell" + | Scheme -> str "Scheme" + | JSON -> str "JSON" + +let warn_deprecated_ocaml_spelling = + CWarnings.create ~name:"deprecated-ocaml-spelling" ~category:"deprecated" + (fun () -> + strbrk ("The spelling \"OCaml\" should be used instead of \"Ocaml\".")) + +} + +VERNAC ARGUMENT EXTEND language +PRINTED BY { pr_language } +| [ "Ocaml" ] -> { let _ = warn_deprecated_ocaml_spelling () in Ocaml } +| [ "OCaml" ] -> { Ocaml } +| [ "Haskell" ] -> { Haskell } +| [ "Scheme" ] -> { Scheme } +| [ "JSON" ] -> { JSON } +END + +(* Extraction commands *) + +VERNAC COMMAND EXTEND Extraction CLASSIFIED AS QUERY +(* Extraction in the Coq toplevel *) +| [ "Extraction" global(x) ] -> { simple_extraction x } +| [ "Recursive" "Extraction" ne_global_list(l) ] -> { full_extraction None l } + +(* Monolithic extraction to a file *) +| [ "Extraction" string(f) ne_global_list(l) ] + -> { full_extraction (Some f) l } + +(* Extraction to a temporary file and OCaml compilation *) +| [ "Extraction" "TestCompile" ne_global_list(l) ] + -> { extract_and_compile l } +END + +VERNAC COMMAND EXTEND SeparateExtraction CLASSIFIED AS QUERY +(* Same, with content splitted in several files *) +| [ "Separate" "Extraction" ne_global_list(l) ] + -> { separate_extraction l } +END + +(* Modular extraction (one Coq library = one ML module) *) +VERNAC COMMAND EXTEND ExtractionLibrary CLASSIFIED AS QUERY +| [ "Extraction" "Library" ident(m) ] + -> { extraction_library false m } +END + +VERNAC COMMAND EXTEND RecursiveExtractionLibrary CLASSIFIED AS QUERY +| [ "Recursive" "Extraction" "Library" ident(m) ] + -> { extraction_library true m } +END + +(* Target Language *) +VERNAC COMMAND EXTEND ExtractionLanguage CLASSIFIED AS SIDEFF +| [ "Extraction" "Language" language(l) ] + -> { extraction_language l } +END + +VERNAC COMMAND EXTEND ExtractionInline CLASSIFIED AS SIDEFF +(* Custom inlining directives *) +| [ "Extraction" "Inline" ne_global_list(l) ] + -> { extraction_inline true l } +END + +VERNAC COMMAND EXTEND ExtractionNoInline CLASSIFIED AS SIDEFF +| [ "Extraction" "NoInline" ne_global_list(l) ] + -> { extraction_inline false l } +END + +VERNAC COMMAND EXTEND PrintExtractionInline CLASSIFIED AS QUERY +| [ "Print" "Extraction" "Inline" ] + -> {Feedback. msg_info (print_extraction_inline ()) } +END + +VERNAC COMMAND EXTEND ResetExtractionInline CLASSIFIED AS SIDEFF +| [ "Reset" "Extraction" "Inline" ] + -> { reset_extraction_inline () } +END + +VERNAC COMMAND EXTEND ExtractionImplicit CLASSIFIED AS SIDEFF +(* Custom implicit arguments of some csts/inds/constructors *) +| [ "Extraction" "Implicit" global(r) "[" int_or_id_list(l) "]" ] + -> { extraction_implicit r l } +END + +VERNAC COMMAND EXTEND ExtractionBlacklist CLASSIFIED AS SIDEFF +(* Force Extraction to not use some filenames *) +| [ "Extraction" "Blacklist" ne_ident_list(l) ] + -> { extraction_blacklist l } +END + +VERNAC COMMAND EXTEND PrintExtractionBlacklist CLASSIFIED AS QUERY +| [ "Print" "Extraction" "Blacklist" ] + -> { Feedback.msg_info (print_extraction_blacklist ()) } +END + +VERNAC COMMAND EXTEND ResetExtractionBlacklist CLASSIFIED AS SIDEFF +| [ "Reset" "Extraction" "Blacklist" ] + -> { reset_extraction_blacklist () } +END + + +(* Overriding of a Coq object by an ML one *) +VERNAC COMMAND EXTEND ExtractionConstant CLASSIFIED AS SIDEFF +| [ "Extract" "Constant" global(x) string_list(idl) "=>" mlname(y) ] + -> { extract_constant_inline false x idl y } +END + +VERNAC COMMAND EXTEND ExtractionInlinedConstant CLASSIFIED AS SIDEFF +| [ "Extract" "Inlined" "Constant" global(x) "=>" mlname(y) ] + -> { extract_constant_inline true x [] y } +END + +VERNAC COMMAND EXTEND ExtractionInductive CLASSIFIED AS SIDEFF +| [ "Extract" "Inductive" global(x) "=>" + mlname(id) "[" mlname_list(idl) "]" string_opt(o) ] + -> { extract_inductive x id idl o } +END +(* Show the extraction of the current proof *) + +VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY +| [ "Show" "Extraction" ] + -> { show_extraction () } +END diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 deleted file mode 100644 index db753fc672..0000000000 --- a/plugins/firstorder/g_ground.ml4 +++ /dev/null @@ -1,164 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* Some !ground_depth); - optwrite= - (function - None->ground_depth:=3 - | Some i->ground_depth:=(max i 0))} - in - declare_int_option gdopt - - -let _= - let congruence_depth=ref 100 in - let gdopt= - { optdepr=true; (* noop *) - optname="Congruence Depth"; - optkey=["Congruence";"Depth"]; - optread=(fun ()->Some !congruence_depth); - optwrite= - (function - None->congruence_depth:=0 - | Some i->congruence_depth:=(max i 0))} - in - declare_int_option gdopt - -let default_intuition_tac = - let tac _ _ = Auto.h_auto None [] None in - let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in - let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in - Tacenv.register_ml_tactic name [| tac |]; - Tacexpr.TacML (Loc.tag (entry, [])) - -let (set_default_solver, default_solver, print_default_solver) = - Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" - -VERNAC COMMAND FUNCTIONAL EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF -| [ "Set" "Firstorder" "Solver" tactic(t) ] -> [ - fun ~atts ~st -> let open Vernacinterp in - set_default_solver - (Locality.make_section_locality atts.locality) - (Tacintern.glob_tactic t); - st - ] -END - -VERNAC COMMAND EXTEND Firstorder_Print_Solver CLASSIFIED AS QUERY -| [ "Print" "Firstorder" "Solver" ] -> [ - Feedback.msg_info - (Pp.(++) (Pp.str"Firstorder solver tactic is ") (print_default_solver ())) ] -END - -let fail_solver=tclFAIL 0 (Pp.str "GTauto failed") - -let gen_ground_tac flag taco ids bases = - let backup= !qflag in - Proofview.tclOR begin - Proofview.Goal.enter begin fun gl -> - qflag:=flag; - let solver= - match taco with - Some tac-> tac - | None-> snd (default_solver ()) in - let startseq k = - Proofview.Goal.enter begin fun gl -> - let seq=empty_seq !ground_depth in - let seq, sigma = extend_with_ref_list (pf_env gl) (project gl) ids seq in - let seq, sigma = extend_with_auto_hints (pf_env gl) (project gl) bases seq in - tclTHEN (Proofview.Unsafe.tclEVARS sigma) (k seq) - end - in - let result=ground_tac solver startseq in - qflag := backup; - result - end - end - (fun (e, info) -> qflag := backup; Proofview.tclZERO ~info e) - -(* special for compatibility with Intuition - -let constant str = Coqlib.get_constr str - -let defined_connectives=lazy - [[],EvalConstRef (destConst (constant "core.not.type")); - [],EvalConstRef (destConst (constant "core.iff.type"))] - -let normalize_evaluables= - onAllHypsAndConcl - (function - None->unfold_in_concl (Lazy.force defined_connectives) - | Some id-> - unfold_in_hyp (Lazy.force defined_connectives) - (Tacexpr.InHypType id)) *) - -open Ppconstr -open Printer -let pr_firstorder_using_raw _ _ _ = Pptactic.pr_auto_using pr_qualid -let pr_firstorder_using_glob _ _ _ = Pptactic.pr_auto_using (pr_or_var (fun x -> pr_global (snd x))) -let pr_firstorder_using_typed _ _ _ = Pptactic.pr_auto_using pr_global - -let warn_deprecated_syntax = - CWarnings.create ~name:"firstorder-deprecated-syntax" ~category:"deprecated" - (fun () -> Pp.strbrk "Deprecated syntax; use \",\" as separator") - - -ARGUMENT EXTEND firstorder_using - TYPED AS reference_list - PRINTED BY pr_firstorder_using_typed - RAW_PRINTED BY pr_firstorder_using_raw - GLOB_PRINTED BY pr_firstorder_using_glob -| [ "using" reference(a) ] -> [ [a] ] -| [ "using" reference(a) "," ne_reference_list_sep(l,",") ] -> [ a::l ] -| [ "using" reference(a) reference(b) reference_list(l) ] -> [ - warn_deprecated_syntax (); - a::b::l - ] -| [ ] -> [ [] ] -END - -TACTIC EXTEND firstorder - [ "firstorder" tactic_opt(t) firstorder_using(l) ] -> - [ gen_ground_tac true (Option.map (tactic_of_value ist) t) l [] ] -| [ "firstorder" tactic_opt(t) "with" ne_preident_list(l) ] -> - [ gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l ] -| [ "firstorder" tactic_opt(t) firstorder_using(l) - "with" ne_preident_list(l') ] -> - [ gen_ground_tac true (Option.map (tactic_of_value ist) t) l l' ] -END - -TACTIC EXTEND gintuition - [ "gintuition" tactic_opt(t) ] -> - [ gen_ground_tac false (Option.map (tactic_of_value ist) t) [] [] ] -END diff --git a/plugins/firstorder/g_ground.mlg b/plugins/firstorder/g_ground.mlg new file mode 100644 index 0000000000..c41687e721 --- /dev/null +++ b/plugins/firstorder/g_ground.mlg @@ -0,0 +1,173 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Some !ground_depth); + optwrite= + (function + None->ground_depth:=3 + | Some i->ground_depth:=(max i 0))} + in + declare_int_option gdopt + + +let _= + let congruence_depth=ref 100 in + let gdopt= + { optdepr=true; (* noop *) + optname="Congruence Depth"; + optkey=["Congruence";"Depth"]; + optread=(fun ()->Some !congruence_depth); + optwrite= + (function + None->congruence_depth:=0 + | Some i->congruence_depth:=(max i 0))} + in + declare_int_option gdopt + +let default_intuition_tac = + let tac _ _ = Auto.h_auto None [] None in + let name = { Tacexpr.mltac_plugin = "ground_plugin"; mltac_tactic = "auto_with"; } in + let entry = { Tacexpr.mltac_name = name; mltac_index = 0 } in + Tacenv.register_ml_tactic name [| tac |]; + Tacexpr.TacML (Loc.tag (entry, [])) + +let (set_default_solver, default_solver, print_default_solver) = + Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" + +} + +VERNAC COMMAND EXTEND Firstorder_Set_Solver CLASSIFIED AS SIDEFF +| [ "Set" "Firstorder" "Solver" tactic(t) ] -> { + let open Vernacinterp in + set_default_solver + (Locality.make_section_locality atts.locality) + (Tacintern.glob_tactic t) + } +END + +VERNAC COMMAND EXTEND Firstorder_Print_Solver CLASSIFIED AS QUERY +| [ "Print" "Firstorder" "Solver" ] -> { + Feedback.msg_info + (Pp.(++) (Pp.str"Firstorder solver tactic is ") (print_default_solver ())) } +END + +{ + +let fail_solver=tclFAIL 0 (Pp.str "GTauto failed") + +let gen_ground_tac flag taco ids bases = + let backup= !qflag in + Proofview.tclOR begin + Proofview.Goal.enter begin fun gl -> + qflag:=flag; + let solver= + match taco with + Some tac-> tac + | None-> snd (default_solver ()) in + let startseq k = + Proofview.Goal.enter begin fun gl -> + let seq=empty_seq !ground_depth in + let seq, sigma = extend_with_ref_list (pf_env gl) (project gl) ids seq in + let seq, sigma = extend_with_auto_hints (pf_env gl) (project gl) bases seq in + tclTHEN (Proofview.Unsafe.tclEVARS sigma) (k seq) + end + in + let result=ground_tac solver startseq in + qflag := backup; + result + end + end + (fun (e, info) -> qflag := backup; Proofview.tclZERO ~info e) + +(* special for compatibility with Intuition + +let constant str = Coqlib.get_constr str + +let defined_connectives=lazy + [[],EvalConstRef (destConst (constant "core.not.type")); + [],EvalConstRef (destConst (constant "core.iff.type"))] + +let normalize_evaluables= + onAllHypsAndConcl + (function + None->unfold_in_concl (Lazy.force defined_connectives) + | Some id-> + unfold_in_hyp (Lazy.force defined_connectives) + (Tacexpr.InHypType id)) *) + +open Ppconstr +open Printer +let pr_firstorder_using_raw _ _ _ = Pptactic.pr_auto_using pr_qualid +let pr_firstorder_using_glob _ _ _ = Pptactic.pr_auto_using (pr_or_var (fun x -> pr_global (snd x))) +let pr_firstorder_using_typed _ _ _ = Pptactic.pr_auto_using pr_global + +let warn_deprecated_syntax = + CWarnings.create ~name:"firstorder-deprecated-syntax" ~category:"deprecated" + (fun () -> Pp.strbrk "Deprecated syntax; use \",\" as separator") + +} + +ARGUMENT EXTEND firstorder_using + TYPED AS reference list + PRINTED BY { pr_firstorder_using_typed } + RAW_PRINTED BY { pr_firstorder_using_raw } + GLOB_PRINTED BY { pr_firstorder_using_glob } +| [ "using" reference(a) ] -> { [a] } +| [ "using" reference(a) "," ne_reference_list_sep(l,",") ] -> { a::l } +| [ "using" reference(a) reference(b) reference_list(l) ] -> { + warn_deprecated_syntax (); + a::b::l + } +| [ ] -> { [] } +END + +TACTIC EXTEND firstorder +| [ "firstorder" tactic_opt(t) firstorder_using(l) ] -> + { gen_ground_tac true (Option.map (tactic_of_value ist) t) l [] } +| [ "firstorder" tactic_opt(t) "with" ne_preident_list(l) ] -> + { gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l } +| [ "firstorder" tactic_opt(t) firstorder_using(l) + "with" ne_preident_list(l') ] -> + { gen_ground_tac true (Option.map (tactic_of_value ist) t) l l' } +END + +TACTIC EXTEND gintuition +| [ "gintuition" tactic_opt(t) ] -> + { gen_ground_tac false (Option.map (tactic_of_value ist) t) [] [] } +END diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 deleted file mode 100644 index c7555c44eb..0000000000 --- a/plugins/funind/g_indfun.ml4 +++ /dev/null @@ -1,236 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* mt () - | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings prc prlc b) - -(* Duplication of printing functions because "'a with_bindings" is - (internally) not uniform in 'a: indeed constr_with_bindings at the - "typed" level has type "open_constr with_bindings" instead of - "constr with_bindings"; hence, its printer cannot be polymorphic in - (prc,prlc)... *) - -let pr_fun_ind_using_typed prc prlc _ opt_c = - match opt_c with - | None -> mt () - | Some b -> - let env = Global.env () in - let evd = Evd.from_env env in - let (_, b) = b env evd in - spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings prc prlc b) - - -ARGUMENT EXTEND fun_ind_using - TYPED AS constr_with_bindings option - PRINTED BY pr_fun_ind_using_typed - RAW_PRINTED BY pr_fun_ind_using - GLOB_PRINTED BY pr_fun_ind_using -| [ "using" constr_with_bindings(c) ] -> [ Some c ] -| [ ] -> [ None ] -END - - -TACTIC EXTEND newfuninv - [ "functional" "inversion" quantified_hypothesis(hyp) reference_opt(fname) ] -> - [ - Proofview.V82.tactic (Invfun.invfun hyp fname) - ] -END - -let pr_intro_as_pat _prc _ _ pat = - match pat with - | Some pat -> - spc () ++ str "as" ++ spc () ++ (* Miscprint.pr_intro_pattern prc pat *) - str"" - | None -> mt () - -let out_disjunctive = CAst.map (function - | IntroAction (IntroOrAndPattern l) -> l - | _ -> CErrors.user_err Pp.(str "Disjunctive or conjunctive intro pattern expected.")) - -ARGUMENT EXTEND with_names TYPED AS intropattern_opt PRINTED BY pr_intro_as_pat -| [ "as" simple_intropattern(ipat) ] -> [ Some ipat ] -| [] ->[ None ] -END - -let functional_induction b c x pat = - Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat)) - - -TACTIC EXTEND newfunind - ["functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> - [ - let c = match cl with - | [] -> assert false - | [c] -> c - | c::cl -> EConstr.applist(c,cl) - in - Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ] -END -(***** debug only ***) -TACTIC EXTEND snewfunind - ["soft" "functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> - [ - let c = match cl with - | [] -> assert false - | [c] -> c - | c::cl -> EConstr.applist(c,cl) - in - Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ] -END - - -let pr_constr_comma_sequence prc _ _ = prlist_with_sep pr_comma prc - -ARGUMENT EXTEND constr_comma_sequence' - TYPED AS constr_list - PRINTED BY pr_constr_comma_sequence -| [ constr(c) "," constr_comma_sequence'(l) ] -> [ c::l ] -| [ constr(c) ] -> [ [c] ] -END - -let pr_auto_using prc _prlc _prt = Pptactic.pr_auto_using prc - -ARGUMENT EXTEND auto_using' - TYPED AS constr_list - PRINTED BY pr_auto_using -| [ "using" constr_comma_sequence'(l) ] -> [ l ] -| [ ] -> [ [] ] -END - -module Gram = Pcoq.Gram -module Vernac = Pvernac.Vernac_ -module Tactic = Pltac - -type function_rec_definition_loc_argtype = (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) Loc.located - -let (wit_function_rec_definition_loc : function_rec_definition_loc_argtype Genarg.uniform_genarg_type) = - Genarg.create_arg "function_rec_definition_loc" - -let function_rec_definition_loc = - Pcoq.create_generic_entry Pcoq.utactic "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc) - -GEXTEND Gram - GLOBAL: function_rec_definition_loc ; - - function_rec_definition_loc: - [ [ g = Vernac.rec_definition -> Loc.tag ~loc:!@loc g ]] - ; - -END - -let () = - let raw_printer _ _ _ (loc,body) = Ppvernac.pr_rec_definition body in - Pptactic.declare_extra_vernac_genarg_pprule wit_function_rec_definition_loc raw_printer - -(* TASSI: n'importe quoi ! *) -VERNAC COMMAND EXTEND Function - ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")] - => [ let hard = List.exists (function - | _,((_,(_,(CMeasureRec _|CWfRec _)),_,_,_),_) -> true - | _,((_,(_,CStructRec),_,_,_),_) -> false) recsl in - match - Vernac_classifier.classify_vernac - (Vernacexpr.(VernacExpr([], VernacFixpoint(Decl_kinds.NoDischarge, List.map snd recsl)))) - with - | Vernacexpr.VtSideff ids, _ when hard -> - Vernacexpr.(VtStartProof ("Classic", GuaranteesOpacity, ids), VtLater) - | x -> x ] - -> [ do_generate_principle false (List.map snd recsl) ] -END - -let pr_fun_scheme_arg (princ_name,fun_name,s) = - Names.Id.print princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ - Libnames.pr_qualid fun_name ++ spc() ++ str "Sort " ++ - Termops.pr_sort_family s - -VERNAC ARGUMENT EXTEND fun_scheme_arg -PRINTED BY pr_fun_scheme_arg -| [ ident(princ_name) ":=" "Induction" "for" reference(fun_name) "Sort" sort_family(s) ] -> [ (princ_name,fun_name,s) ] -END - - -let warning_error names e = - let (e, _) = ExplainErr.process_vernac_interp_error (e, Exninfo.null) in - match e with - | Building_graph e -> - let names = pr_enum Libnames.pr_qualid names in - let error = if do_observe () then (spc () ++ CErrors.print e) else mt () in - warn_cannot_define_graph (names,error) - | Defining_principle e -> - let names = pr_enum Libnames.pr_qualid names in - let error = if do_observe () then CErrors.print e else mt () in - warn_cannot_define_principle (names,error) - | _ -> raise e - - -VERNAC COMMAND EXTEND NewFunctionalScheme - ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ] - => [ Vernacexpr.VtSideff(List.map pi1 fas), Vernacexpr.VtLater ] - -> - [ - begin - try - Functional_principles_types.build_scheme fas - with Functional_principles_types.No_graph_found -> - begin - match fas with - | (_,fun_name,_)::_ -> - begin - begin - make_graph (Smartlocate.global_with_alias fun_name) - end - ; - try Functional_principles_types.build_scheme fas - with Functional_principles_types.No_graph_found -> - CErrors.user_err Pp.(str "Cannot generate induction principle(s)") - | e when CErrors.noncritical e -> - let names = List.map (fun (_,na,_) -> na) fas in - warning_error names e - - end - | _ -> assert false (* we can only have non empty list *) - end - | e when CErrors.noncritical e -> - let names = List.map (fun (_,na,_) -> na) fas in - warning_error names e - end - - ] -END -(***** debug only ***) - -VERNAC COMMAND EXTEND NewFunctionalCase - ["Functional" "Case" fun_scheme_arg(fas) ] - => [ Vernacexpr.VtSideff[pi1 fas], Vernacexpr.VtLater ] - -> [ Functional_principles_types.build_case_scheme fas ] -END - -(***** debug only ***) -VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY -["Generate" "graph" "for" reference(c)] -> [ make_graph (Smartlocate.global_with_alias c) ] -END diff --git a/plugins/funind/g_indfun.mlg b/plugins/funind/g_indfun.mlg new file mode 100644 index 0000000000..857215751a --- /dev/null +++ b/plugins/funind/g_indfun.mlg @@ -0,0 +1,271 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* mt () + | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings prc prlc b) + +(* Duplication of printing functions because "'a with_bindings" is + (internally) not uniform in 'a: indeed constr_with_bindings at the + "typed" level has type "open_constr with_bindings" instead of + "constr with_bindings"; hence, its printer cannot be polymorphic in + (prc,prlc)... *) + +let pr_fun_ind_using_typed prc prlc _ opt_c = + match opt_c with + | None -> mt () + | Some b -> + let env = Global.env () in + let evd = Evd.from_env env in + let (_, b) = b env evd in + spc () ++ hov 2 (str "using" ++ spc () ++ Miscprint.pr_with_bindings prc prlc b) + +} + +ARGUMENT EXTEND fun_ind_using + TYPED AS constr_with_bindings option + PRINTED BY { pr_fun_ind_using_typed } + RAW_PRINTED BY { pr_fun_ind_using } + GLOB_PRINTED BY { pr_fun_ind_using } +| [ "using" constr_with_bindings(c) ] -> { Some c } +| [ ] -> { None } +END + + +TACTIC EXTEND newfuninv +| [ "functional" "inversion" quantified_hypothesis(hyp) reference_opt(fname) ] -> + { + Proofview.V82.tactic (Invfun.invfun hyp fname) + } +END + +{ + +let pr_intro_as_pat _prc _ _ pat = + match pat with + | Some pat -> + spc () ++ str "as" ++ spc () ++ (* Miscprint.pr_intro_pattern prc pat *) + str"" + | None -> mt () + +let out_disjunctive = CAst.map (function + | IntroAction (IntroOrAndPattern l) -> l + | _ -> CErrors.user_err Pp.(str "Disjunctive or conjunctive intro pattern expected.")) + +} + +ARGUMENT EXTEND with_names TYPED AS intropattern option PRINTED BY { pr_intro_as_pat } +| [ "as" simple_intropattern(ipat) ] -> { Some ipat } +| [] -> { None } +END + +{ + +let functional_induction b c x pat = + Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat)) + +} + +TACTIC EXTEND newfunind +| ["functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> + { + let c = match cl with + | [] -> assert false + | [c] -> c + | c::cl -> EConstr.applist(c,cl) + in + Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl } +END +(***** debug only ***) +TACTIC EXTEND snewfunind +| ["soft" "functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> + { + let c = match cl with + | [] -> assert false + | [c] -> c + | c::cl -> EConstr.applist(c,cl) + in + Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl } +END + +{ + +let pr_constr_comma_sequence prc _ _ = prlist_with_sep pr_comma prc + +} + +ARGUMENT EXTEND constr_comma_sequence' + TYPED AS constr list + PRINTED BY { pr_constr_comma_sequence } +| [ constr(c) "," constr_comma_sequence'(l) ] -> { c::l } +| [ constr(c) ] -> { [c] } +END + +{ + +let pr_auto_using prc _prlc _prt = Pptactic.pr_auto_using prc + +} + +ARGUMENT EXTEND auto_using' + TYPED AS constr list + PRINTED BY { pr_auto_using } +| [ "using" constr_comma_sequence'(l) ] -> { l } +| [ ] -> { [] } +END + +{ + +module Gram = Pcoq.Gram +module Vernac = Pvernac.Vernac_ +module Tactic = Pltac + +type function_rec_definition_loc_argtype = (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) Loc.located + +let (wit_function_rec_definition_loc : function_rec_definition_loc_argtype Genarg.uniform_genarg_type) = + Genarg.create_arg "function_rec_definition_loc" + +let function_rec_definition_loc = + Pcoq.create_generic_entry Pcoq.utactic "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc) + +} + +GRAMMAR EXTEND Gram + GLOBAL: function_rec_definition_loc ; + + function_rec_definition_loc: + [ [ g = Vernac.rec_definition -> { Loc.tag ~loc g } ]] + ; + +END + +{ + +let () = + let raw_printer _ _ _ (loc,body) = Ppvernac.pr_rec_definition body in + Pptactic.declare_extra_vernac_genarg_pprule wit_function_rec_definition_loc raw_printer + +} + +(* TASSI: n'importe quoi ! *) +VERNAC COMMAND EXTEND Function +| ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")] + => { let hard = List.exists (function + | _,((_,(_,(CMeasureRec _|CWfRec _)),_,_,_),_) -> true + | _,((_,(_,CStructRec),_,_,_),_) -> false) recsl in + match + Vernac_classifier.classify_vernac + (Vernacexpr.(VernacExpr([], VernacFixpoint(Decl_kinds.NoDischarge, List.map snd recsl)))) + with + | Vernacexpr.VtSideff ids, _ when hard -> + Vernacexpr.(VtStartProof ("Classic", GuaranteesOpacity, ids), VtLater) + | x -> x } + -> { do_generate_principle false (List.map snd recsl) } +END + +{ + +let pr_fun_scheme_arg (princ_name,fun_name,s) = + Names.Id.print princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ + Libnames.pr_qualid fun_name ++ spc() ++ str "Sort " ++ + Termops.pr_sort_family s + +} + +VERNAC ARGUMENT EXTEND fun_scheme_arg +PRINTED BY { pr_fun_scheme_arg } +| [ ident(princ_name) ":=" "Induction" "for" reference(fun_name) "Sort" sort_family(s) ] -> { (princ_name,fun_name,s) } +END + +{ + +let warning_error names e = + let (e, _) = ExplainErr.process_vernac_interp_error (e, Exninfo.null) in + match e with + | Building_graph e -> + let names = pr_enum Libnames.pr_qualid names in + let error = if do_observe () then (spc () ++ CErrors.print e) else mt () in + warn_cannot_define_graph (names,error) + | Defining_principle e -> + let names = pr_enum Libnames.pr_qualid names in + let error = if do_observe () then CErrors.print e else mt () in + warn_cannot_define_principle (names,error) + | _ -> raise e + +} + +VERNAC COMMAND EXTEND NewFunctionalScheme +| ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ] + => { Vernacexpr.VtSideff(List.map pi1 fas), Vernacexpr.VtLater } + -> + { + begin + try + Functional_principles_types.build_scheme fas + with Functional_principles_types.No_graph_found -> + begin + match fas with + | (_,fun_name,_)::_ -> + begin + begin + make_graph (Smartlocate.global_with_alias fun_name) + end + ; + try Functional_principles_types.build_scheme fas + with Functional_principles_types.No_graph_found -> + CErrors.user_err Pp.(str "Cannot generate induction principle(s)") + | e when CErrors.noncritical e -> + let names = List.map (fun (_,na,_) -> na) fas in + warning_error names e + + end + | _ -> assert false (* we can only have non empty list *) + end + | e when CErrors.noncritical e -> + let names = List.map (fun (_,na,_) -> na) fas in + warning_error names e + end + + } +END +(***** debug only ***) + +VERNAC COMMAND EXTEND NewFunctionalCase +| ["Functional" "Case" fun_scheme_arg(fas) ] + => { Vernacexpr.VtSideff[pi1 fas], Vernacexpr.VtLater } + -> { Functional_principles_types.build_case_scheme fas } +END + +(***** debug only ***) +VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY +| ["Generate" "graph" "for" reference(c)] -> { make_graph (Smartlocate.global_with_alias c) } +END diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4 deleted file mode 100644 index 4de27e8138..0000000000 --- a/plugins/ltac/extraargs.ml4 +++ /dev/null @@ -1,310 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* Metasyntax.add_token_obj "<-"; - Metasyntax.add_token_obj "->") - "ltac_plugin" - -let pr_orient _prc _prlc _prt = function - | true -> Pp.mt () - | false -> Pp.str " <-" - -ARGUMENT EXTEND orient TYPED AS bool PRINTED BY pr_orient -| [ "->" ] -> [ true ] -| [ "<-" ] -> [ false ] -| [ ] -> [ true ] -END - -let pr_int _ _ _ i = Pp.int i - -let _natural = Pcoq.Prim.natural - -ARGUMENT EXTEND natural TYPED AS int PRINTED BY pr_int -| [ _natural(i) ] -> [ i ] -END - -let pr_orient = pr_orient () () () - - -let pr_int_list = Pp.pr_sequence Pp.int -let pr_int_list_full _prc _prlc _prt l = pr_int_list l - -let pr_occurrences _prc _prlc _prt l = - match l with - | ArgArg x -> pr_int_list x - | ArgVar { CAst.loc = loc; v=id } -> Id.print id - -let occurrences_of = function - | [] -> NoOccurrences - | n::_ as nl when n < 0 -> AllOccurrencesBut (List.map abs nl) - | nl -> - if List.exists (fun n -> n < 0) nl then - CErrors.user_err Pp.(str "Illegal negative occurrence number."); - OnlyOccurrences nl - -let coerce_to_int v = match Value.to_int v with - | None -> raise (CannotCoerceTo "an integer") - | Some n -> n - -let int_list_of_VList v = match Value.to_list v with -| Some l -> List.map (fun n -> coerce_to_int n) l -| _ -> raise (CannotCoerceTo "an integer") - -let interp_occs ist gl l = - match l with - | ArgArg x -> x - | ArgVar ({ CAst.v = id } as locid) -> - (try int_list_of_VList (Id.Map.find id ist.lfun) - with Not_found | CannotCoerceTo _ -> [interp_int ist locid]) -let interp_occs ist gl l = - Tacmach.project gl , interp_occs ist gl l - -let glob_occs ist l = l - -let subst_occs evm l = l - -ARGUMENT EXTEND occurrences - TYPED AS int list - PRINTED BY pr_int_list_full - - INTERPRETED BY interp_occs - GLOBALIZED BY glob_occs - SUBSTITUTED BY subst_occs - - RAW_PRINTED BY pr_occurrences - GLOB_PRINTED BY pr_occurrences - -| [ ne_integer_list(l) ] -> [ ArgArg l ] -| [ var(id) ] -> [ ArgVar id ] -END - -let pr_occurrences = pr_occurrences () () () - -let pr_gen prc _prlc _prtac c = prc c - -let pr_globc _prc _prlc _prtac (_,glob) = - let _, env = Pfedit.get_current_context () in - Printer.pr_glob_constr_env env glob - -let interp_glob ist gl (t,_) = Tacmach.project gl , (ist,t) - -let glob_glob = Tacintern.intern_constr - -let pr_lconstr _ prc _ c = prc c - -let subst_glob = Tacsubst.subst_glob_constr_and_expr - -ARGUMENT EXTEND glob - PRINTED BY pr_globc - - INTERPRETED BY interp_glob - GLOBALIZED BY glob_glob - SUBSTITUTED BY subst_glob - - RAW_PRINTED BY pr_gen - GLOB_PRINTED BY pr_gen - [ constr(c) ] -> [ c ] -END - -let l_constr = Pcoq.Constr.lconstr - -ARGUMENT EXTEND lconstr - TYPED AS constr - PRINTED BY pr_lconstr - [ l_constr(c) ] -> [ c ] -END - -ARGUMENT EXTEND lglob - TYPED AS glob - PRINTED BY pr_globc - - INTERPRETED BY interp_glob - GLOBALIZED BY glob_glob - SUBSTITUTED BY subst_glob - - RAW_PRINTED BY pr_gen - GLOB_PRINTED BY pr_gen - [ lconstr(c) ] -> [ c ] -END - -let interp_casted_constr ist gl c = - interp_constr_gen (Pretyping.OfType (pf_concl gl)) ist (pf_env gl) (project gl) c - -ARGUMENT EXTEND casted_constr - TYPED AS constr - PRINTED BY pr_gen - INTERPRETED BY interp_casted_constr - [ constr(c) ] -> [ c ] -END - -type 'id gen_place= ('id * hyp_location_flag,unit) location - -type loc_place = lident gen_place -type place = Id.t gen_place - -let pr_gen_place pr_id = function - ConclLocation () -> Pp.mt () - | HypLocation (id,InHyp) -> str "in " ++ pr_id id - | HypLocation (id,InHypTypeOnly) -> - str "in (type of " ++ pr_id id ++ str ")" - | HypLocation (id,InHypValueOnly) -> - str "in (value of " ++ pr_id id ++ str ")" - -let pr_loc_place _ _ _ = pr_gen_place (fun { CAst.v = id } -> Id.print id) -let pr_place _ _ _ = pr_gen_place Id.print -let pr_hloc = pr_loc_place () () () - -let intern_place ist = function - ConclLocation () -> ConclLocation () - | HypLocation (id,hl) -> HypLocation (Tacintern.intern_hyp ist id,hl) - -let interp_place ist env sigma = function - ConclLocation () -> ConclLocation () - | HypLocation (id,hl) -> HypLocation (Tacinterp.interp_hyp ist env sigma id,hl) - -let interp_place ist gl p = - Tacmach.project gl , interp_place ist (Tacmach.pf_env gl) (Tacmach.project gl) p - -let subst_place subst pl = pl - -let warn_deprecated_instantiate_syntax = - CWarnings.create ~name:"deprecated-instantiate-syntax" ~category:"deprecated" - (fun (v,v',id) -> - let s = Id.to_string id in - Pp.strbrk - ("Syntax \"in (" ^ v ^ " of " ^ s ^ ")\" is deprecated; use \"in (" ^ v' ^ " of " ^ s ^ ")\".") - ) - -ARGUMENT EXTEND hloc - PRINTED BY pr_place - INTERPRETED BY interp_place - GLOBALIZED BY intern_place - SUBSTITUTED BY subst_place - RAW_PRINTED BY pr_loc_place - GLOB_PRINTED BY pr_loc_place - [ ] -> - [ ConclLocation () ] - | [ "in" "|-" "*" ] -> - [ ConclLocation () ] -| [ "in" ident(id) ] -> - [ HypLocation ((CAst.make id),InHyp) ] -| [ "in" "(" "Type" "of" ident(id) ")" ] -> - [ warn_deprecated_instantiate_syntax ("Type","type",id); - HypLocation ((CAst.make id),InHypTypeOnly) ] -| [ "in" "(" "Value" "of" ident(id) ")" ] -> - [ warn_deprecated_instantiate_syntax ("Value","value",id); - HypLocation ((CAst.make id),InHypValueOnly) ] -| [ "in" "(" "type" "of" ident(id) ")" ] -> - [ HypLocation ((CAst.make id),InHypTypeOnly) ] -| [ "in" "(" "value" "of" ident(id) ")" ] -> - [ HypLocation ((CAst.make id),InHypValueOnly) ] - - END - -let pr_rename _ _ _ (n, m) = Id.print n ++ str " into " ++ Id.print m - -ARGUMENT EXTEND rename - TYPED AS ident * ident - PRINTED BY pr_rename -| [ ident(n) "into" ident(m) ] -> [ (n, m) ] -END - -(* Julien: Mise en commun des differentes version de replace with in by *) - -let pr_by_arg_tac _prc _prlc prtac opt_c = - match opt_c with - | None -> mt () - | Some t -> hov 2 (str "by" ++ spc () ++ prtac (3,Notation_gram.E) t) - -ARGUMENT EXTEND by_arg_tac - TYPED AS tactic_opt - PRINTED BY pr_by_arg_tac -| [ "by" tactic3(c) ] -> [ Some c ] -| [ ] -> [ None ] -END - -let pr_by_arg_tac prtac opt_c = pr_by_arg_tac () () prtac opt_c - -let pr_in_clause _ _ _ cl = Pptactic.pr_in_clause Ppconstr.pr_lident cl -let pr_in_top_clause _ _ _ cl = Pptactic.pr_in_clause Id.print cl -let in_clause' = Pltac.in_clause - -ARGUMENT EXTEND in_clause - TYPED AS clause_dft_concl - PRINTED BY pr_in_top_clause - RAW_PRINTED BY pr_in_clause - GLOB_PRINTED BY pr_in_clause -| [ in_clause'(cl) ] -> [ cl ] -END - -let local_test_lpar_id_colon = - let err () = raise Stream.Failure in - Pcoq.Gram.Entry.of_parser "lpar_id_colon" - (fun strm -> - match Util.stream_nth 0 strm with - | Tok.KEYWORD "(" -> - (match Util.stream_nth 1 strm with - | Tok.IDENT _ -> - (match Util.stream_nth 2 strm with - | Tok.KEYWORD ":" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) - -let pr_lpar_id_colon _ _ _ _ = mt () - -ARGUMENT EXTEND test_lpar_id_colon TYPED AS unit PRINTED BY pr_lpar_id_colon -| [ local_test_lpar_id_colon(x) ] -> [ () ] -END diff --git a/plugins/ltac/extraargs.mlg b/plugins/ltac/extraargs.mlg new file mode 100644 index 0000000000..c4c4e51ecc --- /dev/null +++ b/plugins/ltac/extraargs.mlg @@ -0,0 +1,353 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* Metasyntax.add_token_obj "<-"; + Metasyntax.add_token_obj "->") + "ltac_plugin" + +let pr_orient _prc _prlc _prt = function + | true -> Pp.mt () + | false -> Pp.str " <-" + +} + +ARGUMENT EXTEND orient TYPED AS bool PRINTED BY { pr_orient } +| [ "->" ] -> { true } +| [ "<-" ] -> { false } +| [ ] -> { true } +END + +{ + +let pr_int _ _ _ i = Pp.int i + +let _natural = Pcoq.Prim.natural + +} + +ARGUMENT EXTEND natural TYPED AS int PRINTED BY { pr_int } +| [ _natural(i) ] -> { i } +END + +{ + +let pr_orient = pr_orient () () () + +let pr_int_list = Pp.pr_sequence Pp.int +let pr_int_list_full _prc _prlc _prt l = pr_int_list l + +let pr_occurrences _prc _prlc _prt l = + match l with + | ArgArg x -> pr_int_list x + | ArgVar { CAst.loc = loc; v=id } -> Id.print id + +let occurrences_of = function + | [] -> NoOccurrences + | n::_ as nl when n < 0 -> AllOccurrencesBut (List.map abs nl) + | nl -> + if List.exists (fun n -> n < 0) nl then + CErrors.user_err Pp.(str "Illegal negative occurrence number."); + OnlyOccurrences nl + +let coerce_to_int v = match Value.to_int v with + | None -> raise (CannotCoerceTo "an integer") + | Some n -> n + +let int_list_of_VList v = match Value.to_list v with +| Some l -> List.map (fun n -> coerce_to_int n) l +| _ -> raise (CannotCoerceTo "an integer") + +let interp_occs ist gl l = + match l with + | ArgArg x -> x + | ArgVar ({ CAst.v = id } as locid) -> + (try int_list_of_VList (Id.Map.find id ist.lfun) + with Not_found | CannotCoerceTo _ -> [interp_int ist locid]) +let interp_occs ist gl l = + Tacmach.project gl , interp_occs ist gl l + +let glob_occs ist l = l + +let subst_occs evm l = l + +} + +ARGUMENT EXTEND occurrences + TYPED AS int list + PRINTED BY { pr_int_list_full } + + INTERPRETED BY { interp_occs } + GLOBALIZED BY { glob_occs } + SUBSTITUTED BY { subst_occs } + + RAW_PRINTED BY { pr_occurrences } + GLOB_PRINTED BY { pr_occurrences } + +| [ ne_integer_list(l) ] -> { ArgArg l } +| [ var(id) ] -> { ArgVar id } +END + +{ + +let pr_occurrences = pr_occurrences () () () + +let pr_gen prc _prlc _prtac c = prc c + +let pr_globc _prc _prlc _prtac (_,glob) = + let _, env = Pfedit.get_current_context () in + Printer.pr_glob_constr_env env glob + +let interp_glob ist gl (t,_) = Tacmach.project gl , (ist,t) + +let glob_glob = Tacintern.intern_constr + +let pr_lconstr _ prc _ c = prc c + +let subst_glob = Tacsubst.subst_glob_constr_and_expr + +} + +ARGUMENT EXTEND glob + PRINTED BY { pr_globc } + + INTERPRETED BY { interp_glob } + GLOBALIZED BY { glob_glob } + SUBSTITUTED BY { subst_glob } + + RAW_PRINTED BY { pr_gen } + GLOB_PRINTED BY { pr_gen } +| [ constr(c) ] -> { c } +END + +{ + +let l_constr = Pcoq.Constr.lconstr + +} + +ARGUMENT EXTEND lconstr + TYPED AS constr + PRINTED BY { pr_lconstr } +| [ l_constr(c) ] -> { c } +END + +ARGUMENT EXTEND lglob + TYPED AS glob + PRINTED BY { pr_globc } + + INTERPRETED BY { interp_glob } + GLOBALIZED BY { glob_glob } + SUBSTITUTED BY { subst_glob } + + RAW_PRINTED BY { pr_gen } + GLOB_PRINTED BY { pr_gen } +| [ lconstr(c) ] -> { c } +END + +{ + +let interp_casted_constr ist gl c = + interp_constr_gen (Pretyping.OfType (pf_concl gl)) ist (pf_env gl) (project gl) c + +} + +ARGUMENT EXTEND casted_constr + TYPED AS constr + PRINTED BY { pr_gen } + INTERPRETED BY { interp_casted_constr } +| [ constr(c) ] -> { c } +END + +{ + +type 'id gen_place= ('id * hyp_location_flag,unit) location + +type loc_place = lident gen_place +type place = Id.t gen_place + +let pr_gen_place pr_id = function + ConclLocation () -> Pp.mt () + | HypLocation (id,InHyp) -> str "in " ++ pr_id id + | HypLocation (id,InHypTypeOnly) -> + str "in (type of " ++ pr_id id ++ str ")" + | HypLocation (id,InHypValueOnly) -> + str "in (value of " ++ pr_id id ++ str ")" + +let pr_loc_place _ _ _ = pr_gen_place (fun { CAst.v = id } -> Id.print id) +let pr_place _ _ _ = pr_gen_place Id.print +let pr_hloc = pr_loc_place () () () + +let intern_place ist = function + ConclLocation () -> ConclLocation () + | HypLocation (id,hl) -> HypLocation (Tacintern.intern_hyp ist id,hl) + +let interp_place ist env sigma = function + ConclLocation () -> ConclLocation () + | HypLocation (id,hl) -> HypLocation (Tacinterp.interp_hyp ist env sigma id,hl) + +let interp_place ist gl p = + Tacmach.project gl , interp_place ist (Tacmach.pf_env gl) (Tacmach.project gl) p + +let subst_place subst pl = pl + +let warn_deprecated_instantiate_syntax = + CWarnings.create ~name:"deprecated-instantiate-syntax" ~category:"deprecated" + (fun (v,v',id) -> + let s = Id.to_string id in + Pp.strbrk + ("Syntax \"in (" ^ v ^ " of " ^ s ^ ")\" is deprecated; use \"in (" ^ v' ^ " of " ^ s ^ ")\".") + ) + +} + +ARGUMENT EXTEND hloc + PRINTED BY { pr_place } + INTERPRETED BY { interp_place } + GLOBALIZED BY { intern_place } + SUBSTITUTED BY { subst_place } + RAW_PRINTED BY { pr_loc_place } + GLOB_PRINTED BY { pr_loc_place } +| [ ] -> + { ConclLocation () } + | [ "in" "|-" "*" ] -> + { ConclLocation () } +| [ "in" ident(id) ] -> + { HypLocation ((CAst.make id),InHyp) } +| [ "in" "(" "Type" "of" ident(id) ")" ] -> + { warn_deprecated_instantiate_syntax ("Type","type",id); + HypLocation ((CAst.make id),InHypTypeOnly) } +| [ "in" "(" "Value" "of" ident(id) ")" ] -> + { warn_deprecated_instantiate_syntax ("Value","value",id); + HypLocation ((CAst.make id),InHypValueOnly) } +| [ "in" "(" "type" "of" ident(id) ")" ] -> + { HypLocation ((CAst.make id),InHypTypeOnly) } +| [ "in" "(" "value" "of" ident(id) ")" ] -> + { HypLocation ((CAst.make id),InHypValueOnly) } + + END + +{ + +let pr_rename _ _ _ (n, m) = Id.print n ++ str " into " ++ Id.print m + +} + +ARGUMENT EXTEND rename + TYPED AS (ident * ident) + PRINTED BY { pr_rename } +| [ ident(n) "into" ident(m) ] -> { (n, m) } +END + +(* Julien: Mise en commun des differentes version de replace with in by *) + +{ + +let pr_by_arg_tac _prc _prlc prtac opt_c = + match opt_c with + | None -> mt () + | Some t -> hov 2 (str "by" ++ spc () ++ prtac (3,Notation_gram.E) t) + +} + +ARGUMENT EXTEND by_arg_tac + TYPED AS tactic option + PRINTED BY { pr_by_arg_tac } +| [ "by" tactic3(c) ] -> { Some c } +| [ ] -> { None } +END + +{ + +let pr_by_arg_tac prtac opt_c = pr_by_arg_tac () () prtac opt_c + +let pr_in_clause _ _ _ cl = Pptactic.pr_in_clause Ppconstr.pr_lident cl +let pr_in_top_clause _ _ _ cl = Pptactic.pr_in_clause Id.print cl +let in_clause' = Pltac.in_clause + +} + +ARGUMENT EXTEND in_clause + TYPED AS clause_dft_concl + PRINTED BY { pr_in_top_clause } + RAW_PRINTED BY { pr_in_clause } + GLOB_PRINTED BY { pr_in_clause } +| [ in_clause'(cl) ] -> { cl } +END + +{ + +let local_test_lpar_id_colon = + let err () = raise Stream.Failure in + Pcoq.Gram.Entry.of_parser "lpar_id_colon" + (fun strm -> + match Util.stream_nth 0 strm with + | Tok.KEYWORD "(" -> + (match Util.stream_nth 1 strm with + | Tok.IDENT _ -> + (match Util.stream_nth 2 strm with + | Tok.KEYWORD ":" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + +let pr_lpar_id_colon _ _ _ _ = mt () + +} + +ARGUMENT EXTEND test_lpar_id_colon TYPED AS unit PRINTED BY { pr_lpar_id_colon } +| [ local_test_lpar_id_colon(x) ] -> { () } +END diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 deleted file mode 100644 index e5b032e638..0000000000 --- a/plugins/ltac/extratactics.ml4 +++ /dev/null @@ -1,1065 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* replace_in_clause_maybe_by c1 c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) - -let replace_term ist dir_opt c cl = - with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) - -TACTIC EXTEND replace - ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] --> [ replace_in_clause_maybe_by ist c1 c2 cl tac ] -END - -TACTIC EXTEND replace_term_left - [ "replace" "->" uconstr(c) clause(cl) ] - -> [ replace_term ist (Some true) c cl ] -END - -TACTIC EXTEND replace_term_right - [ "replace" "<-" uconstr(c) clause(cl) ] - -> [ replace_term ist (Some false) c cl ] -END - -TACTIC EXTEND replace_term - [ "replace" uconstr(c) clause(cl) ] - -> [ replace_term ist None c cl ] -END - -let induction_arg_of_quantified_hyp = function - | AnonHyp n -> None,ElimOnAnonHyp n - | NamedHyp id -> None,ElimOnIdent (CAst.make id) - -(* Versions *_main must come first!! so that "1" is interpreted as a - ElimOnAnonHyp and not as a "constr", and "id" is interpreted as a - ElimOnIdent and not as "constr" *) - -let mytclWithHoles tac with_evars c = - Proofview.Goal.enter begin fun gl -> - let env = Tacmach.New.pf_env gl in - let sigma = Tacmach.New.project gl in - let sigma',c = Tactics.force_destruction_arg with_evars env sigma c in - Tacticals.New.tclWITHHOLES with_evars (tac with_evars (Some c)) sigma' - end - -let elimOnConstrWithHoles tac with_evars c = - Tacticals.New.tclDELAYEDWITHHOLES with_evars c - (fun c -> tac with_evars (Some (None,ElimOnConstr c))) - -TACTIC EXTEND simplify_eq - [ "simplify_eq" ] -> [ dEq ~keep_proofs:None false None ] -| [ "simplify_eq" destruction_arg(c) ] -> [ mytclWithHoles (dEq ~keep_proofs:None) false c ] -END -TACTIC EXTEND esimplify_eq -| [ "esimplify_eq" ] -> [ dEq ~keep_proofs:None true None ] -| [ "esimplify_eq" destruction_arg(c) ] -> [ mytclWithHoles (dEq ~keep_proofs:None) true c ] -END - -let discr_main c = elimOnConstrWithHoles discr_tac false c - -TACTIC EXTEND discriminate -| [ "discriminate" ] -> [ discr_tac false None ] -| [ "discriminate" destruction_arg(c) ] -> - [ mytclWithHoles discr_tac false c ] -END -TACTIC EXTEND ediscriminate -| [ "ediscriminate" ] -> [ discr_tac true None ] -| [ "ediscriminate" destruction_arg(c) ] -> - [ mytclWithHoles discr_tac true c ] -END - -let discrHyp id = - Proofview.tclEVARMAP >>= fun sigma -> - discr_main (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) - -let injection_main with_evars c = - elimOnConstrWithHoles (injClause None None) with_evars c - -TACTIC EXTEND injection -| [ "injection" ] -> [ injClause None None false None ] -| [ "injection" destruction_arg(c) ] -> [ mytclWithHoles (injClause None None) false c ] -END -TACTIC EXTEND einjection -| [ "einjection" ] -> [ injClause None None true None ] -| [ "einjection" destruction_arg(c) ] -> [ mytclWithHoles (injClause None None) true c ] -END -TACTIC EXTEND injection_as -| [ "injection" "as" intropattern_list(ipat)] -> - [ injClause None (Some ipat) false None ] -| [ "injection" destruction_arg(c) "as" intropattern_list(ipat)] -> - [ mytclWithHoles (injClause None (Some ipat)) false c ] -END -TACTIC EXTEND einjection_as -| [ "einjection" "as" intropattern_list(ipat)] -> - [ injClause None (Some ipat) true None ] -| [ "einjection" destruction_arg(c) "as" intropattern_list(ipat)] -> - [ mytclWithHoles (injClause None (Some ipat)) true c ] -END -TACTIC EXTEND simple_injection -| [ "simple" "injection" ] -> [ simpleInjClause None false None ] -| [ "simple" "injection" destruction_arg(c) ] -> [ mytclWithHoles (simpleInjClause None) false c ] -END - -let injHyp id = - Proofview.tclEVARMAP >>= fun sigma -> - injection_main false (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) - -TACTIC EXTEND dependent_rewrite -| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ] -| [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ] - -> [ rewriteInHyp b c id ] -END - -(** To be deprecated?, "cutrewrite (t=u) as <-" is equivalent to - "replace u with t" or "enough (t=u) as <-" and - "cutrewrite (t=u) as ->" is equivalent to "enough (t=u) as ->". *) - -TACTIC EXTEND cut_rewrite -| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b eqn ] -| [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ] - -> [ cutRewriteInHyp b eqn id ] -END - -(**********************************************************************) -(* Decompose *) - -TACTIC EXTEND decompose_sum -| [ "decompose" "sum" constr(c) ] -> [ Elim.h_decompose_or c ] -END - -TACTIC EXTEND decompose_record -| [ "decompose" "record" constr(c) ] -> [ Elim.h_decompose_and c ] -END - -(**********************************************************************) -(* Contradiction *) - -open Contradiction - -TACTIC EXTEND absurd - [ "absurd" constr(c) ] -> [ absurd c ] -END - -let onSomeWithHoles tac = function - | None -> tac None - | Some c -> Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> tac (Some c)) - -TACTIC EXTEND contradiction - [ "contradiction" constr_with_bindings_opt(c) ] -> - [ onSomeWithHoles contradiction c ] -END - -(**********************************************************************) -(* AutoRewrite *) - -open Autorewrite - -let pr_orient _prc _prlc _prt = function - | true -> Pp.mt () - | false -> Pp.str " <-" - -let pr_orient_string _prc _prlc _prt (orient, s) = - pr_orient _prc _prlc _prt orient ++ Pp.spc () ++ Pp.str s - -ARGUMENT EXTEND orient_string TYPED AS (bool * string) PRINTED BY pr_orient_string -| [ orient(r) preident(i) ] -> [ r, i ] -END - -TACTIC EXTEND autorewrite -| [ "autorewrite" "with" ne_preident_list(l) clause(cl) ] -> - [ auto_multi_rewrite l ( cl) ] -| [ "autorewrite" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> - [ - auto_multi_rewrite_with (Tacinterp.tactic_of_value ist t) l cl - ] -END - -TACTIC EXTEND autorewrite_star -| [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) ] -> - [ auto_multi_rewrite ~conds:AllMatches l cl ] -| [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> - [ auto_multi_rewrite_with ~conds:AllMatches (Tacinterp.tactic_of_value ist t) l cl ] -END - -(**********************************************************************) -(* Rewrite star *) - -let rewrite_star ist clause orient occs c (tac : Geninterp.Val.t option) = - let tac' = Option.map (fun t -> Tacinterp.tactic_of_value ist t, FirstSolved) tac in - with_delayed_uconstr ist c - (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings) true) - -TACTIC EXTEND rewrite_star -| [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) "at" occurrences(occ) by_arg_tac(tac) ] -> - [ rewrite_star ist (Some id) o (occurrences_of occ) c tac ] -| [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) "in" hyp(id) by_arg_tac(tac) ] -> - [ rewrite_star ist (Some id) o (occurrences_of occ) c tac ] -| [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) by_arg_tac(tac) ] -> - [ rewrite_star ist (Some id) o Locus.AllOccurrences c tac ] -| [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) by_arg_tac(tac) ] -> - [ rewrite_star ist None o (occurrences_of occ) c tac ] -| [ "rewrite" "*" orient(o) uconstr(c) by_arg_tac(tac) ] -> - [ rewrite_star ist None o Locus.AllOccurrences c tac ] - END - -(**********************************************************************) -(* Hint Rewrite *) - -let add_rewrite_hint ~poly bases ort t lcsr = - let env = Global.env() in - let sigma = Evd.from_env env in - let f ce = - let c, ctx = Constrintern.interp_constr env sigma ce in - let c = EConstr.to_constr sigma c in - let ctx = - let ctx = UState.context_set ctx in - if poly then ctx - else (** This is a global universe context that shouldn't be - refreshed at every use of the hint, declare it globally. *) - (Declare.declare_universe_context false ctx; - Univ.ContextSet.empty) - in - CAst.make ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in - let eqs = List.map f lcsr in - let add_hints base = add_rew_rules base eqs in - List.iter add_hints bases - -let classify_hint _ = Vernacexpr.VtSideff [], Vernacexpr.VtLater - -VERNAC COMMAND FUNCTIONAL EXTEND HintRewrite CLASSIFIED BY classify_hint - [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ":" preident_list(bl) ] -> - [ fun ~atts ~st -> add_rewrite_hint ~poly:atts.polymorphic bl o None l; st ] -| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) - ":" preident_list(bl) ] -> - [ fun ~atts ~st -> add_rewrite_hint ~poly:atts.polymorphic bl o (Some t) l; st ] -| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ] -> - [ fun ~atts ~st -> add_rewrite_hint ~poly:atts.polymorphic ["core"] o None l; st ] -| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) ] -> - [ fun ~atts ~st -> add_rewrite_hint ~poly:atts.polymorphic ["core"] o (Some t) l; st ] -END - -(**********************************************************************) -(* Refine *) - -open EConstr -open Vars - -let constr_flags () = { - Pretyping.use_typeclasses = true; - Pretyping.solve_unification_constraints = Pfedit.use_unification_heuristics (); - Pretyping.use_hook = Pfedit.solve_by_implicit_tactic (); - Pretyping.fail_evar = false; - Pretyping.expand_evars = true } - -let refine_tac ist simple with_classes c = - Proofview.Goal.enter begin fun gl -> - let concl = Proofview.Goal.concl gl in - let env = Proofview.Goal.env gl in - let flags = - { constr_flags () with Pretyping.use_typeclasses = with_classes } in - let expected_type = Pretyping.OfType concl in - let c = Tacinterp.type_uconstr ~flags ~expected_type ist c in - let update = begin fun sigma -> - c env sigma - end in - let refine = Refine.refine ~typecheck:false update in - if simple then refine - else refine <*> - Tactics.New.reduce_after_refine <*> - Proofview.shelve_unifiable - end - -TACTIC EXTEND refine -| [ "refine" uconstr(c) ] -> - [ refine_tac ist false true c ] -END - -TACTIC EXTEND simple_refine -| [ "simple" "refine" uconstr(c) ] -> - [ refine_tac ist true true c ] -END - -TACTIC EXTEND notcs_refine -| [ "notypeclasses" "refine" uconstr(c) ] -> - [ refine_tac ist false false c ] -END - -TACTIC EXTEND notcs_simple_refine -| [ "simple" "notypeclasses" "refine" uconstr(c) ] -> - [ refine_tac ist true false c ] -END - -(* Solve unification constraints using heuristics or fail if any remain *) -TACTIC EXTEND solve_constraints -[ "solve_constraints" ] -> [ Refine.solve_constraints ] -END - -(**********************************************************************) -(* Inversion lemmas (Leminv) *) - -open Inv -open Leminv - -let seff id = Vernacexpr.VtSideff [id], Vernacexpr.VtLater - -(*VERNAC ARGUMENT EXTEND sort_family -| [ "Set" ] -> [ InSet ] -| [ "Prop" ] -> [ InProp ] -| [ "Type" ] -> [ InType ] -END*) - -VERNAC COMMAND FUNCTIONAL EXTEND DeriveInversionClear -| [ "Derive" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] - => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c s false inv_clear_tac; st ] - -| [ "Derive" "Inversion_clear" ident(na) "with" constr(c) ] => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c Sorts.InProp false inv_clear_tac; st ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND DeriveInversion -| [ "Derive" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] - => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c s false inv_tac; st ] - -| [ "Derive" "Inversion" ident(na) "with" constr(c) ] => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c Sorts.InProp false inv_tac; st ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND DeriveDependentInversion -| [ "Derive" "Dependent" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] - => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c s true dinv_tac; st ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND DeriveDependentInversionClear -| [ "Derive" "Dependent" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] - => [ seff na ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - add_inversion_lemma_exn ~poly:atts.polymorphic na c s true dinv_clear_tac; st ] -END - -(**********************************************************************) -(* Subst *) - -TACTIC EXTEND subst -| [ "subst" ne_var_list(l) ] -> [ subst l ] -| [ "subst" ] -> [ subst_all () ] -END - -let simple_subst_tactic_flags = - { only_leibniz = true; rewrite_dependent_proof = false } - -TACTIC EXTEND simple_subst -| [ "simple" "subst" ] -> [ subst_all ~flags:simple_subst_tactic_flags () ] -END - -open Evar_tactics - -(**********************************************************************) -(* Evar creation *) - -(* TODO: add support for some test similar to g_constr.name_colon so that - expressions like "evar (list A)" do not raise a syntax error *) -TACTIC EXTEND evar - [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> [ let_evar (Name.Name id) typ ] -| [ "evar" constr(typ) ] -> [ let_evar Name.Anonymous typ ] -END - -TACTIC EXTEND instantiate - [ "instantiate" "(" ident(id) ":=" lglob(c) ")" ] -> - [ Tacticals.New.tclTHEN (instantiate_tac_by_name id c) Proofview.V82.nf_evar_goals ] -| [ "instantiate" "(" integer(i) ":=" lglob(c) ")" hloc(hl) ] -> - [ Tacticals.New.tclTHEN (instantiate_tac i c hl) Proofview.V82.nf_evar_goals ] -| [ "instantiate" ] -> [ Proofview.V82.nf_evar_goals ] -END - -(**********************************************************************) -(** Nijmegen "step" tactic for setoid rewriting *) - -open Tactics -open Glob_term -open Libobject -open Lib - -(* Registered lemmas are expected to be of the form - x R y -> y == z -> x R z (in the right table) - x R y -> x == z -> z R y (in the left table) -*) - -let transitivity_right_table = Summary.ref [] ~name:"transitivity-steps-r" -let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l" - -(* [step] tries to apply a rewriting lemma; then apply [tac] intended to - complete to proof of the last hypothesis (assumed to state an equality) *) - -let step left x tac = - let l = - List.map (fun lem -> - let lem = EConstr.of_constr lem in - Tacticals.New.tclTHENLAST - (apply_with_bindings (lem, ImplicitBindings [x])) - tac) - !(if left then transitivity_left_table else transitivity_right_table) - in - Tacticals.New.tclFIRST l - -(* Main function to push lemmas in persistent environment *) - -let cache_transitivity_lemma (_,(left,lem)) = - if left then - transitivity_left_table := lem :: !transitivity_left_table - else - transitivity_right_table := lem :: !transitivity_right_table - -let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref) - -let inTransitivity : bool * Constr.t -> obj = - declare_object {(default_object "TRANSITIVITY-STEPS") with - cache_function = cache_transitivity_lemma; - open_function = (fun i o -> if Int.equal i 1 then cache_transitivity_lemma o); - subst_function = subst_transitivity_lemma; - classify_function = (fun o -> Substitute o) } - -(* Main entry points *) - -let add_transitivity_lemma left lem = - let env = Global.env () in - let sigma = Evd.from_env env in - let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in - let lem' = EConstr.to_constr sigma lem' in - add_anonymous_leaf (inTransitivity (left,lem')) - -(* Vernacular syntax *) - -TACTIC EXTEND stepl -| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true c (Tacinterp.tactic_of_value ist tac) ] -| ["stepl" constr(c) ] -> [ step true c (Proofview.tclUNIT ()) ] -END - -TACTIC EXTEND stepr -| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false c (Tacinterp.tactic_of_value ist tac) ] -| ["stepr" constr(c) ] -> [ step false c (Proofview.tclUNIT ()) ] -END - -VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF -| [ "Declare" "Left" "Step" constr(t) ] -> - [ add_transitivity_lemma true t ] -END - -VERNAC COMMAND EXTEND AddStepr CLASSIFIED AS SIDEFF -| [ "Declare" "Right" "Step" constr(t) ] -> - [ add_transitivity_lemma false t ] -END - -let cache_implicit_tactic (_,tac) = match tac with - | Some tac -> Pfedit.declare_implicit_tactic (Tacinterp.eval_tactic tac) - | None -> Pfedit.clear_implicit_tactic () - -let subst_implicit_tactic (subst,tac) = - Option.map (Tacsubst.subst_tactic subst) tac - -let inImplicitTactic : glob_tactic_expr option -> obj = - declare_object {(default_object "IMPLICIT-TACTIC") with - open_function = (fun i o -> if Int.equal i 1 then cache_implicit_tactic o); - cache_function = cache_implicit_tactic; - subst_function = subst_implicit_tactic; - classify_function = (fun o -> Dispose)} - -let warn_deprecated_implicit_tactic = - CWarnings.create ~name:"deprecated-implicit-tactic" ~category:"deprecated" - (fun () -> strbrk "Implicit tactics are deprecated") - -let declare_implicit_tactic tac = - let () = warn_deprecated_implicit_tactic () in - Lib.add_anonymous_leaf (inImplicitTactic (Some (Tacintern.glob_tactic tac))) - -let clear_implicit_tactic () = - let () = warn_deprecated_implicit_tactic () in - Lib.add_anonymous_leaf (inImplicitTactic None) - -VERNAC COMMAND EXTEND ImplicitTactic CLASSIFIED AS SIDEFF -| [ "Declare" "Implicit" "Tactic" tactic(tac) ] -> [ declare_implicit_tactic tac ] -| [ "Clear" "Implicit" "Tactic" ] -> [ clear_implicit_tactic () ] -END - - - - -(**********************************************************************) -(* sozeau: abs/gen for induction on instantiated dependent inductives, using "Ford" induction as - defined by Conor McBride *) -TACTIC EXTEND generalize_eqs -| ["generalize_eqs" hyp(id) ] -> [ abstract_generalize ~generalize_vars:false id ] -END -TACTIC EXTEND dep_generalize_eqs -| ["dependent" "generalize_eqs" hyp(id) ] -> [ abstract_generalize ~generalize_vars:false ~force_dep:true id ] -END -TACTIC EXTEND generalize_eqs_vars -| ["generalize_eqs_vars" hyp(id) ] -> [ abstract_generalize ~generalize_vars:true id ] -END -TACTIC EXTEND dep_generalize_eqs_vars -| ["dependent" "generalize_eqs_vars" hyp(id) ] -> [ abstract_generalize ~force_dep:true ~generalize_vars:true id ] -END - -(** Tactic to automatically simplify hypotheses of the form [Π Δ, x_i = t_i -> T] - where [t_i] is closed w.r.t. Δ. Such hypotheses are automatically generated - during dependent induction. For internal use. *) - -TACTIC EXTEND specialize_eqs -[ "specialize_eqs" hyp(id) ] -> [ specialize_eqs id ] -END - -(**********************************************************************) -(* A tactic that considers a given occurrence of [c] in [t] and *) -(* abstract the minimal set of all the occurrences of [c] so that the *) -(* abstraction [fun x -> t[x/c]] is well-typed *) -(* *) -(* Contributed by Chung-Kil Hur (Winter 2009) *) -(**********************************************************************) - -let subst_var_with_hole occ tid t = - let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in - let locref = ref 0 in - let rec substrec x = match DAst.get x with - | GVar id -> - if Id.equal id tid - then - (decr occref; - if Int.equal !occref 0 then x - else - (incr locref; - DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ - GHole (Evar_kinds.QuestionMark { - Evar_kinds.qm_obligation=Evar_kinds.Define true; - Evar_kinds.qm_name=Anonymous; - Evar_kinds.qm_record_field=None; - }, IntroAnonymous, None))) - else x - | _ -> map_glob_constr_left_to_right substrec x in - let t' = substrec t - in - if !occref > 0 then Find_subterm.error_invalid_occurrence [occ] else t' - -let subst_hole_with_term occ tc t = - let locref = ref 0 in - let occref = ref occ in - let rec substrec c = match DAst.get c with - | GHole (Evar_kinds.QuestionMark { - Evar_kinds.qm_obligation=Evar_kinds.Define true; - Evar_kinds.qm_name=Anonymous; - Evar_kinds.qm_record_field=None; - }, IntroAnonymous, s) -> - decr occref; - if Int.equal !occref 0 then tc - else - (incr locref; - DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ - GHole (Evar_kinds.QuestionMark { - Evar_kinds.qm_obligation=Evar_kinds.Define true; - Evar_kinds.qm_name=Anonymous; - Evar_kinds.qm_record_field=None; - },IntroAnonymous,s)) - | _ -> map_glob_constr_left_to_right substrec c - in - substrec t - -open Tacmach - -let hResolve id c occ t = - Proofview.Goal.enter begin fun gl -> - let sigma = Proofview.Goal.sigma gl in - let env = Termops.clear_named_body id (Proofview.Goal.env gl) in - let concl = Proofview.Goal.concl gl in - let env_ids = Termops.vars_of_env env in - let c_raw = Detyping.detype Detyping.Now true env_ids env sigma c in - let t_raw = Detyping.detype Detyping.Now true env_ids env sigma t in - let rec resolve_hole t_hole = - try - Pretyping.understand env sigma t_hole - with - | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> - let (e, info) = CErrors.push e in - let loc_begin = Option.cata (fun l -> fst (Loc.unloc l)) 0 (Loc.get_loc info) in - resolve_hole (subst_hole_with_term loc_begin c_raw t_hole) - in - let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in - let sigma = Evd.merge_universe_context sigma ctx in - let t_constr_type = Retyping.get_type_of env sigma t_constr in - Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) - (change_concl (mkLetIn (Name.Anonymous,t_constr,t_constr_type,concl))) - end - -let hResolve_auto id c t = - let rec resolve_auto n = - try - hResolve id c n t - with - | UserError _ as e -> raise e - | e when CErrors.noncritical e -> resolve_auto (n+1) - in - resolve_auto 1 - -TACTIC EXTEND hresolve_core -| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "at" int_or_var(occ) "in" constr(t) ] -> [ hResolve id c occ t ] -| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "in" constr(t) ] -> [ hResolve_auto id c t ] -END - -(** - hget_evar -*) - -TACTIC EXTEND hget_evar -| [ "hget_evar" int_or_var(n) ] -> [ Evar_tactics.hget_evar n ] -END - -(**********************************************************************) - -(**********************************************************************) -(* A tactic that reduces one match t with ... by doing destruct t. *) -(* if t is not a variable, the tactic does *) -(* case_eq t;intros ... heq;rewrite heq in *|-. (but heq itself is *) -(* preserved). *) -(* Contributed by Julien Forest and Pierre Courtieu (july 2010) *) -(**********************************************************************) - -exception Found of unit Proofview.tactic - -let rewrite_except h = - Proofview.Goal.enter begin fun gl -> - let hyps = Tacmach.New.pf_ids_of_hyps gl in - Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else - Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) - hyps - end - - -let refl_equal () = Coqlib.lib_ref "core.eq.type" - -(* This is simply an implementation of the case_eq tactic. this code - should be replaced by a call to the tactic but I don't know how to - call it before it is defined. *) -let mkCaseEq a : unit Proofview.tactic = - Proofview.Goal.enter begin fun gl -> - let type_of_a = Tacmach.New.pf_unsafe_type_of gl a in - Tacticals.New.pf_constr_of_global (delayed_force refl_equal) >>= fun req -> - Tacticals.New.tclTHENLIST - [Tactics.generalize [(mkApp(req, [| type_of_a; a|]))]; - Proofview.Goal.enter begin fun gl -> - let concl = Proofview.Goal.concl gl in - let env = Proofview.Goal.env gl in - (** FIXME: this looks really wrong. Does anybody really use this tactic? *) - let (_, c) = Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] env (Evd.from_env env) concl in - change_concl c - end; - simplest_case a] - end - - -let case_eq_intros_rewrite x = - Proofview.Goal.enter begin fun gl -> - let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in - (* Pp.msgnl (Printer.pr_lconstr x); *) - Tacticals.New.tclTHENLIST [ - mkCaseEq x; - Proofview.Goal.enter begin fun gl -> - let concl = Proofview.Goal.concl gl in - let hyps = Tacmach.New.pf_ids_set_of_hyps gl in - let n' = nb_prod (Tacmach.New.project gl) concl in - let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in - Tacticals.New.tclTHENLIST [ - Tacticals.New.tclDO (n'-n-1) intro; - introduction h; - rewrite_except h] - end - ] - end - -let rec find_a_destructable_match sigma t = - let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in - let cl = [cl, (None, None), None], None in - let dest = TacAtom (Loc.tag @@ TacInductionDestruct(false, false, cl)) in - match EConstr.kind sigma t with - | Case (_,_,x,_) when closed0 sigma x -> - if isVar sigma x then - (* TODO check there is no rel n. *) - raise (Found (Tacinterp.eval_tactic dest)) - else - (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) - raise (Found (case_eq_intros_rewrite x)) - | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t - - -let destauto t = - Proofview.tclEVARMAP >>= fun sigma -> - try find_a_destructable_match sigma t; - Tacticals.New.tclZEROMSG (str "No destructable match found") - with Found tac -> tac - -let destauto_in id = - Proofview.Goal.enter begin fun gl -> - let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in -(* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) -(* Pp.msgnl (Printer.pr_lconstr (ctype)); *) - destauto ctype - end - -TACTIC EXTEND destauto -| [ "destauto" ] -> [ Proofview.Goal.enter begin fun gl -> destauto (Proofview.Goal.concl gl) end ] -| [ "destauto" "in" hyp(id) ] -> [ destauto_in id ] -END - -(**********************************************************************) - -(**********************************************************************) -(* A version of abstract constructing transparent terms *) -(* Introduced by Jason Gross and Benjamin Delaware in June 2016 *) -(**********************************************************************) - -TACTIC EXTEND transparent_abstract -| [ "transparent_abstract" tactic3(t) ] -> [ Proofview.Goal.enter begin fun gl -> - Tactics.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) end ] -| [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> [ Proofview.Goal.enter begin fun gl -> - Tactics.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) end ] -END - -(* ********************************************************************* *) - -TACTIC EXTEND constr_eq -| [ "constr_eq" constr(x) constr(y) ] -> [ Tactics.constr_eq ~strict:false x y ] -END - -TACTIC EXTEND constr_eq_strict -| [ "constr_eq_strict" constr(x) constr(y) ] -> [ Tactics.constr_eq ~strict:true x y ] -END - -TACTIC EXTEND constr_eq_nounivs -| [ "constr_eq_nounivs" constr(x) constr(y) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ] -END - -TACTIC EXTEND is_evar -| [ "is_evar" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Evar _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") - ] -END - -TACTIC EXTEND has_evar -| [ "has_evar" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - if Evarutil.has_undefined_evars sigma x - then Proofview.tclUNIT () - else Tacticals.New.tclFAIL 0 (str "No evars") -] -END - -TACTIC EXTEND is_hyp -| [ "is_var" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Var _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") ] -END - -TACTIC EXTEND is_fix -| [ "is_fix" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Fix _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") ] -END;; - -TACTIC EXTEND is_cofix -| [ "is_cofix" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | CoFix _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") ] -END;; - -TACTIC EXTEND is_ind -| [ "is_ind" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Ind _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") ] -END;; - -TACTIC EXTEND is_constructor -| [ "is_constructor" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Construct _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") ] -END;; - -TACTIC EXTEND is_proj -| [ "is_proj" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Proj _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") ] -END;; - -TACTIC EXTEND is_const -| [ "is_const" constr(x) ] -> [ - Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma x with - | Const _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") ] -END;; - -(* Command to grab the evars left unresolved at the end of a proof. *) -(* spiwack: I put it in extratactics because it is somewhat tied with - the semantics of the LCF-style tactics, hence with the classic tactic - mode. *) -VERNAC COMMAND EXTEND GrabEvars -[ "Grab" "Existential" "Variables" ] - => [ Vernac_classifier.classify_as_proofstep ] - -> [ Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.grab_evars p) ] -END - -(* Shelves all the goals under focus. *) -TACTIC EXTEND shelve -| [ "shelve" ] -> - [ Proofview.shelve ] -END - -(* Shelves the unifiable goals under focus, i.e. the goals which - appear in other goals under focus (the unfocused goals are not - considered). *) -TACTIC EXTEND shelve_unifiable -| [ "shelve_unifiable" ] -> - [ Proofview.shelve_unifiable ] -END - -(* Unshelves the goal shelved by the tactic. *) -TACTIC EXTEND unshelve -| [ "unshelve" tactic1(t) ] -> - [ - Proofview.with_shelf (Tacinterp.tactic_of_value ist t) >>= fun (gls, ()) -> - let gls = List.map Proofview.with_empty_state gls in - Proofview.Unsafe.tclGETGOALS >>= fun ogls -> - Proofview.Unsafe.tclSETGOALS (gls @ ogls) - ] -END - -(* Command to add every unshelved variables to the focus *) -VERNAC COMMAND EXTEND Unshelve -[ "Unshelve" ] - => [ Vernac_classifier.classify_as_proofstep ] - -> [ Proof_global.simple_with_current_proof (fun _ p -> Proof.unshelve p) ] -END - -(* Gives up on the goals under focus: the goals are considered solved, - but the proof cannot be closed until the user goes back and solve - these goals. *) -TACTIC EXTEND give_up -| [ "give_up" ] -> - [ Proofview.give_up ] -END - -(* cycles [n] goals *) -TACTIC EXTEND cycle -| [ "cycle" int_or_var(n) ] -> [ Proofview.cycle n ] -END - -(* swaps goals number [i] and [j] *) -TACTIC EXTEND swap -| [ "swap" int_or_var(i) int_or_var(j) ] -> [ Proofview.swap i j ] -END - -(* reverses the list of focused goals *) -TACTIC EXTEND revgoals -| [ "revgoals" ] -> [ Proofview.revgoals ] -END - -type cmp = - | Eq - | Lt | Le - | Gt | Ge - -type 'i test = - | Test of cmp * 'i * 'i - -let pr_cmp = function - | Eq -> Pp.str"=" - | Lt -> Pp.str"<" - | Le -> Pp.str"<=" - | Gt -> Pp.str">" - | Ge -> Pp.str">=" - -let pr_cmp' _prc _prlc _prt = pr_cmp - -let pr_test_gen f (Test(c,x,y)) = - Pp.(f x ++ pr_cmp c ++ f y) - -let pr_test = pr_test_gen (Pputils.pr_or_var Pp.int) - -let pr_test' _prc _prlc _prt = pr_test - -let pr_itest = pr_test_gen Pp.int - -let pr_itest' _prc _prlc _prt = pr_itest - - - -ARGUMENT EXTEND comparison PRINTED BY pr_cmp' -| [ "=" ] -> [ Eq ] -| [ "<" ] -> [ Lt ] -| [ "<=" ] -> [ Le ] -| [ ">" ] -> [ Gt ] -| [ ">=" ] -> [ Ge ] - END - -let interp_test ist gls = function - | Test (c,x,y) -> - project gls , - Test(c,Tacinterp.interp_int_or_var ist x,Tacinterp.interp_int_or_var ist y) - -ARGUMENT EXTEND test - PRINTED BY pr_itest' - INTERPRETED BY interp_test - RAW_PRINTED BY pr_test' - GLOB_PRINTED BY pr_test' -| [ int_or_var(x) comparison(c) int_or_var(y) ] -> [ Test(c,x,y) ] -END - -let interp_cmp = function - | Eq -> Int.equal - | Lt -> ((<):int->int->bool) - | Le -> ((<=):int->int->bool) - | Gt -> ((>):int->int->bool) - | Ge -> ((>=):int->int->bool) - -let run_test = function - | Test(c,x,y) -> interp_cmp c x y - -let guard tst = - if run_test tst then - Proofview.tclUNIT () - else - let msg = Pp.(str"Condition not satisfied:"++ws 1++(pr_itest tst)) in - Tacticals.New.tclZEROMSG msg - - -TACTIC EXTEND guard -| [ "guard" test(tst) ] -> [ guard tst ] -END - -let decompose l c = - Proofview.Goal.enter begin fun gl -> - let sigma = Tacmach.New.project gl in - let to_ind c = - if isInd sigma c then fst (destInd sigma c) - else user_err Pp.(str "not an inductive type") - in - let l = List.map to_ind l in - Elim.h_decompose l c - end - -TACTIC EXTEND decompose -| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose l c ] -END - -(** library/keys *) - -VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF -| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [ - let get_key c = - let env = Global.env () in - let evd = Evd.from_env env in - let (evd, c) = Constrintern.interp_open_constr env evd c in - let kind c = EConstr.kind evd c in - Keys.constr_key kind c - in - let k1 = get_key c in - let k2 = get_key c' in - match k1, k2 with - | Some k1, Some k2 -> Keys.declare_equiv_keys k1 k2 - | _ -> () ] -END - -VERNAC COMMAND EXTEND Print_keys CLASSIFIED AS QUERY -| [ "Print" "Equivalent" "Keys" ] -> [ Feedback.msg_info (Keys.pr_keys Printer.pr_global) ] -END - - -VERNAC COMMAND EXTEND OptimizeProof -| [ "Optimize" "Proof" ] => [ Vernac_classifier.classify_as_proofstep ] -> - [ Proof_global.compact_the_proof () ] -| [ "Optimize" "Heap" ] => [ Vernac_classifier.classify_as_proofstep ] -> - [ Gc.compact () ] -END - -(** tactic analogous to "OPTIMIZE HEAP" *) - -let tclOPTIMIZE_HEAP = - Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> Gc.compact ())) - -TACTIC EXTEND optimize_heap -| [ "optimize_heap" ] -> [ tclOPTIMIZE_HEAP ] -END diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg new file mode 100644 index 0000000000..b660865e8b --- /dev/null +++ b/plugins/ltac/extratactics.mlg @@ -0,0 +1,1160 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* replace_in_clause_maybe_by c1 c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) + +let replace_term ist dir_opt c cl = + with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) + +} + +TACTIC EXTEND replace +| ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] +-> { replace_in_clause_maybe_by ist c1 c2 cl tac } +END + +TACTIC EXTEND replace_term_left +| [ "replace" "->" uconstr(c) clause(cl) ] + -> { replace_term ist (Some true) c cl } +END + +TACTIC EXTEND replace_term_right +| [ "replace" "<-" uconstr(c) clause(cl) ] + -> { replace_term ist (Some false) c cl } +END + +TACTIC EXTEND replace_term +| [ "replace" uconstr(c) clause(cl) ] + -> { replace_term ist None c cl } +END + +{ + +let induction_arg_of_quantified_hyp = function + | AnonHyp n -> None,ElimOnAnonHyp n + | NamedHyp id -> None,ElimOnIdent (CAst.make id) + +(* Versions *_main must come first!! so that "1" is interpreted as a + ElimOnAnonHyp and not as a "constr", and "id" is interpreted as a + ElimOnIdent and not as "constr" *) + +let mytclWithHoles tac with_evars c = + Proofview.Goal.enter begin fun gl -> + let env = Tacmach.New.pf_env gl in + let sigma = Tacmach.New.project gl in + let sigma',c = Tactics.force_destruction_arg with_evars env sigma c in + Tacticals.New.tclWITHHOLES with_evars (tac with_evars (Some c)) sigma' + end + +let elimOnConstrWithHoles tac with_evars c = + Tacticals.New.tclDELAYEDWITHHOLES with_evars c + (fun c -> tac with_evars (Some (None,ElimOnConstr c))) + +} + +TACTIC EXTEND simplify_eq +| [ "simplify_eq" ] -> { dEq ~keep_proofs:None false None } +| [ "simplify_eq" destruction_arg(c) ] -> { mytclWithHoles (dEq ~keep_proofs:None) false c } +END +TACTIC EXTEND esimplify_eq +| [ "esimplify_eq" ] -> { dEq ~keep_proofs:None true None } +| [ "esimplify_eq" destruction_arg(c) ] -> { mytclWithHoles (dEq ~keep_proofs:None) true c } +END + +{ + +let discr_main c = elimOnConstrWithHoles discr_tac false c + +} + +TACTIC EXTEND discriminate +| [ "discriminate" ] -> { discr_tac false None } +| [ "discriminate" destruction_arg(c) ] -> + { mytclWithHoles discr_tac false c } +END +TACTIC EXTEND ediscriminate +| [ "ediscriminate" ] -> { discr_tac true None } +| [ "ediscriminate" destruction_arg(c) ] -> + { mytclWithHoles discr_tac true c } +END + +{ + +let discrHyp id = + Proofview.tclEVARMAP >>= fun sigma -> + discr_main (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) + +let injection_main with_evars c = + elimOnConstrWithHoles (injClause None None) with_evars c + +} + +TACTIC EXTEND injection +| [ "injection" ] -> { injClause None None false None } +| [ "injection" destruction_arg(c) ] -> { mytclWithHoles (injClause None None) false c } +END +TACTIC EXTEND einjection +| [ "einjection" ] -> { injClause None None true None } +| [ "einjection" destruction_arg(c) ] -> { mytclWithHoles (injClause None None) true c } +END +TACTIC EXTEND injection_as +| [ "injection" "as" intropattern_list(ipat)] -> + { injClause None (Some ipat) false None } +| [ "injection" destruction_arg(c) "as" intropattern_list(ipat)] -> + { mytclWithHoles (injClause None (Some ipat)) false c } +END +TACTIC EXTEND einjection_as +| [ "einjection" "as" intropattern_list(ipat)] -> + { injClause None (Some ipat) true None } +| [ "einjection" destruction_arg(c) "as" intropattern_list(ipat)] -> + { mytclWithHoles (injClause None (Some ipat)) true c } +END +TACTIC EXTEND simple_injection +| [ "simple" "injection" ] -> { simpleInjClause None false None } +| [ "simple" "injection" destruction_arg(c) ] -> { mytclWithHoles (simpleInjClause None) false c } +END + +{ + +let injHyp id = + Proofview.tclEVARMAP >>= fun sigma -> + injection_main false (fun env sigma -> (sigma, (EConstr.mkVar id, NoBindings))) + +} + +TACTIC EXTEND dependent_rewrite +| [ "dependent" "rewrite" orient(b) constr(c) ] -> { rewriteInConcl b c } +| [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ] + -> { rewriteInHyp b c id } +END + +(** To be deprecated?, "cutrewrite (t=u) as <-" is equivalent to + "replace u with t" or "enough (t=u) as <-" and + "cutrewrite (t=u) as ->" is equivalent to "enough (t=u) as ->". *) + +TACTIC EXTEND cut_rewrite +| [ "cutrewrite" orient(b) constr(eqn) ] -> { cutRewriteInConcl b eqn } +| [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ] + -> { cutRewriteInHyp b eqn id } +END + +(**********************************************************************) +(* Decompose *) + +TACTIC EXTEND decompose_sum +| [ "decompose" "sum" constr(c) ] -> { Elim.h_decompose_or c } +END + +TACTIC EXTEND decompose_record +| [ "decompose" "record" constr(c) ] -> { Elim.h_decompose_and c } +END + +(**********************************************************************) +(* Contradiction *) + +{ + +open Contradiction + +} + +TACTIC EXTEND absurd +| [ "absurd" constr(c) ] -> { absurd c } +END + +{ + +let onSomeWithHoles tac = function + | None -> tac None + | Some c -> Tacticals.New.tclDELAYEDWITHHOLES false c (fun c -> tac (Some c)) + +} + +TACTIC EXTEND contradiction +| [ "contradiction" constr_with_bindings_opt(c) ] -> + { onSomeWithHoles contradiction c } +END + +(**********************************************************************) +(* AutoRewrite *) + +{ + +open Autorewrite + +let pr_orient _prc _prlc _prt = function + | true -> Pp.mt () + | false -> Pp.str " <-" + +let pr_orient_string _prc _prlc _prt (orient, s) = + pr_orient _prc _prlc _prt orient ++ Pp.spc () ++ Pp.str s + +} + +ARGUMENT EXTEND orient_string TYPED AS (bool * string) PRINTED BY { pr_orient_string } +| [ orient(r) preident(i) ] -> { r, i } +END + +TACTIC EXTEND autorewrite +| [ "autorewrite" "with" ne_preident_list(l) clause(cl) ] -> + { auto_multi_rewrite l ( cl) } +| [ "autorewrite" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> + { + auto_multi_rewrite_with (Tacinterp.tactic_of_value ist t) l cl + } +END + +TACTIC EXTEND autorewrite_star +| [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) ] -> + { auto_multi_rewrite ~conds:AllMatches l cl } +| [ "autorewrite" "*" "with" ne_preident_list(l) clause(cl) "using" tactic(t) ] -> + { auto_multi_rewrite_with ~conds:AllMatches (Tacinterp.tactic_of_value ist t) l cl } +END + +(**********************************************************************) +(* Rewrite star *) + +{ + +let rewrite_star ist clause orient occs c (tac : Geninterp.Val.t option) = + let tac' = Option.map (fun t -> Tacinterp.tactic_of_value ist t, FirstSolved) tac in + with_delayed_uconstr ist c + (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings) true) + +} + +TACTIC EXTEND rewrite_star +| [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) "at" occurrences(occ) by_arg_tac(tac) ] -> + { rewrite_star ist (Some id) o (occurrences_of occ) c tac } +| [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) "in" hyp(id) by_arg_tac(tac) ] -> + { rewrite_star ist (Some id) o (occurrences_of occ) c tac } +| [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) by_arg_tac(tac) ] -> + { rewrite_star ist (Some id) o Locus.AllOccurrences c tac } +| [ "rewrite" "*" orient(o) uconstr(c) "at" occurrences(occ) by_arg_tac(tac) ] -> + { rewrite_star ist None o (occurrences_of occ) c tac } +| [ "rewrite" "*" orient(o) uconstr(c) by_arg_tac(tac) ] -> + { rewrite_star ist None o Locus.AllOccurrences c tac } + END + +(**********************************************************************) +(* Hint Rewrite *) + +{ + +let add_rewrite_hint ~poly bases ort t lcsr = + let env = Global.env() in + let sigma = Evd.from_env env in + let f ce = + let c, ctx = Constrintern.interp_constr env sigma ce in + let c = EConstr.to_constr sigma c in + let ctx = + let ctx = UState.context_set ctx in + if poly then ctx + else (** This is a global universe context that shouldn't be + refreshed at every use of the hint, declare it globally. *) + (Declare.declare_universe_context false ctx; + Univ.ContextSet.empty) + in + CAst.make ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in + let eqs = List.map f lcsr in + let add_hints base = add_rew_rules base eqs in + List.iter add_hints bases + +let classify_hint _ = Vernacexpr.VtSideff [], Vernacexpr.VtLater + +} + +VERNAC COMMAND EXTEND HintRewrite CLASSIFIED BY { classify_hint } +| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ":" preident_list(bl) ] -> + { add_rewrite_hint ~poly:atts.polymorphic bl o None l } +| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) + ":" preident_list(bl) ] -> + { add_rewrite_hint ~poly:atts.polymorphic bl o (Some t) l } +| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) ] -> + { add_rewrite_hint ~poly:atts.polymorphic ["core"] o None l } +| [ "Hint" "Rewrite" orient(o) ne_constr_list(l) "using" tactic(t) ] -> + { add_rewrite_hint ~poly:atts.polymorphic ["core"] o (Some t) l } +END + +(**********************************************************************) +(* Refine *) + +{ + +open EConstr +open Vars + +let constr_flags () = { + Pretyping.use_typeclasses = true; + Pretyping.solve_unification_constraints = Pfedit.use_unification_heuristics (); + Pretyping.use_hook = Pfedit.solve_by_implicit_tactic (); + Pretyping.fail_evar = false; + Pretyping.expand_evars = true } + +let refine_tac ist simple with_classes c = + Proofview.Goal.enter begin fun gl -> + let concl = Proofview.Goal.concl gl in + let env = Proofview.Goal.env gl in + let flags = + { (constr_flags ()) with Pretyping.use_typeclasses = with_classes } in + let expected_type = Pretyping.OfType concl in + let c = Tacinterp.type_uconstr ~flags ~expected_type ist c in + let update = begin fun sigma -> + c env sigma + end in + let refine = Refine.refine ~typecheck:false update in + if simple then refine + else refine <*> + Tactics.New.reduce_after_refine <*> + Proofview.shelve_unifiable + end + +} + +TACTIC EXTEND refine +| [ "refine" uconstr(c) ] -> + { refine_tac ist false true c } +END + +TACTIC EXTEND simple_refine +| [ "simple" "refine" uconstr(c) ] -> + { refine_tac ist true true c } +END + +TACTIC EXTEND notcs_refine +| [ "notypeclasses" "refine" uconstr(c) ] -> + { refine_tac ist false false c } +END + +TACTIC EXTEND notcs_simple_refine +| [ "simple" "notypeclasses" "refine" uconstr(c) ] -> + { refine_tac ist true false c } +END + +(* Solve unification constraints using heuristics or fail if any remain *) +TACTIC EXTEND solve_constraints +| [ "solve_constraints" ] -> { Refine.solve_constraints } +END + +(**********************************************************************) +(* Inversion lemmas (Leminv) *) + +{ + +open Inv +open Leminv + +let seff id = Vernacexpr.VtSideff [id], Vernacexpr.VtLater + +} + +(*VERNAC ARGUMENT EXTEND sort_family +| [ "Set" ] -> { InSet } +| [ "Prop" ] -> { InProp } +| [ "Type" ] -> { InType } +END*) + +VERNAC COMMAND EXTEND DeriveInversionClear +| [ "Derive" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] + => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c s false inv_clear_tac } + +| [ "Derive" "Inversion_clear" ident(na) "with" constr(c) ] => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c Sorts.InProp false inv_clear_tac } +END + +VERNAC COMMAND EXTEND DeriveInversion +| [ "Derive" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] + => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c s false inv_tac } + +| [ "Derive" "Inversion" ident(na) "with" constr(c) ] => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c Sorts.InProp false inv_tac } +END + +VERNAC COMMAND EXTEND DeriveDependentInversion +| [ "Derive" "Dependent" "Inversion" ident(na) "with" constr(c) "Sort" sort_family(s) ] + => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c s true dinv_tac } +END + +VERNAC COMMAND EXTEND DeriveDependentInversionClear +| [ "Derive" "Dependent" "Inversion_clear" ident(na) "with" constr(c) "Sort" sort_family(s) ] + => { seff na } + -> { + let open Vernacinterp in + add_inversion_lemma_exn ~poly:atts.polymorphic na c s true dinv_clear_tac } +END + +(**********************************************************************) +(* Subst *) + +TACTIC EXTEND subst +| [ "subst" ne_var_list(l) ] -> { subst l } +| [ "subst" ] -> { subst_all () } +END + +{ + +let simple_subst_tactic_flags = + { only_leibniz = true; rewrite_dependent_proof = false } + +} + +TACTIC EXTEND simple_subst +| [ "simple" "subst" ] -> { subst_all ~flags:simple_subst_tactic_flags () } +END + +{ + +open Evar_tactics + +} + +(**********************************************************************) +(* Evar creation *) + +(* TODO: add support for some test similar to g_constr.name_colon so that + expressions like "evar (list A)" do not raise a syntax error *) +TACTIC EXTEND evar +| [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> { let_evar (Name.Name id) typ } +| [ "evar" constr(typ) ] -> { let_evar Name.Anonymous typ } +END + +TACTIC EXTEND instantiate +| [ "instantiate" "(" ident(id) ":=" lglob(c) ")" ] -> + { Tacticals.New.tclTHEN (instantiate_tac_by_name id c) Proofview.V82.nf_evar_goals } +| [ "instantiate" "(" integer(i) ":=" lglob(c) ")" hloc(hl) ] -> + { Tacticals.New.tclTHEN (instantiate_tac i c hl) Proofview.V82.nf_evar_goals } +| [ "instantiate" ] -> { Proofview.V82.nf_evar_goals } +END + +(**********************************************************************) +(** Nijmegen "step" tactic for setoid rewriting *) + +{ + +open Tactics +open Glob_term +open Libobject +open Lib + +(* Registered lemmas are expected to be of the form + x R y -> y == z -> x R z (in the right table) + x R y -> x == z -> z R y (in the left table) +*) + +let transitivity_right_table = Summary.ref [] ~name:"transitivity-steps-r" +let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l" + +(* [step] tries to apply a rewriting lemma; then apply [tac] intended to + complete to proof of the last hypothesis (assumed to state an equality) *) + +let step left x tac = + let l = + List.map (fun lem -> + let lem = EConstr.of_constr lem in + Tacticals.New.tclTHENLAST + (apply_with_bindings (lem, ImplicitBindings [x])) + tac) + !(if left then transitivity_left_table else transitivity_right_table) + in + Tacticals.New.tclFIRST l + +(* Main function to push lemmas in persistent environment *) + +let cache_transitivity_lemma (_,(left,lem)) = + if left then + transitivity_left_table := lem :: !transitivity_left_table + else + transitivity_right_table := lem :: !transitivity_right_table + +let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref) + +let inTransitivity : bool * Constr.t -> obj = + declare_object {(default_object "TRANSITIVITY-STEPS") with + cache_function = cache_transitivity_lemma; + open_function = (fun i o -> if Int.equal i 1 then cache_transitivity_lemma o); + subst_function = subst_transitivity_lemma; + classify_function = (fun o -> Substitute o) } + +(* Main entry points *) + +let add_transitivity_lemma left lem = + let env = Global.env () in + let sigma = Evd.from_env env in + let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in + let lem' = EConstr.to_constr sigma lem' in + add_anonymous_leaf (inTransitivity (left,lem')) + +} + +(* Vernacular syntax *) + +TACTIC EXTEND stepl +| ["stepl" constr(c) "by" tactic(tac) ] -> { step true c (Tacinterp.tactic_of_value ist tac) } +| ["stepl" constr(c) ] -> { step true c (Proofview.tclUNIT ()) } +END + +TACTIC EXTEND stepr +| ["stepr" constr(c) "by" tactic(tac) ] -> { step false c (Tacinterp.tactic_of_value ist tac) } +| ["stepr" constr(c) ] -> { step false c (Proofview.tclUNIT ()) } +END + +VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF +| [ "Declare" "Left" "Step" constr(t) ] -> + { add_transitivity_lemma true t } +END + +VERNAC COMMAND EXTEND AddStepr CLASSIFIED AS SIDEFF +| [ "Declare" "Right" "Step" constr(t) ] -> + { add_transitivity_lemma false t } +END + +{ + +let cache_implicit_tactic (_,tac) = match tac with + | Some tac -> Pfedit.declare_implicit_tactic (Tacinterp.eval_tactic tac) + | None -> Pfedit.clear_implicit_tactic () + +let subst_implicit_tactic (subst,tac) = + Option.map (Tacsubst.subst_tactic subst) tac + +let inImplicitTactic : glob_tactic_expr option -> obj = + declare_object {(default_object "IMPLICIT-TACTIC") with + open_function = (fun i o -> if Int.equal i 1 then cache_implicit_tactic o); + cache_function = cache_implicit_tactic; + subst_function = subst_implicit_tactic; + classify_function = (fun o -> Dispose)} + +let warn_deprecated_implicit_tactic = + CWarnings.create ~name:"deprecated-implicit-tactic" ~category:"deprecated" + (fun () -> strbrk "Implicit tactics are deprecated") + +let declare_implicit_tactic tac = + let () = warn_deprecated_implicit_tactic () in + Lib.add_anonymous_leaf (inImplicitTactic (Some (Tacintern.glob_tactic tac))) + +let clear_implicit_tactic () = + let () = warn_deprecated_implicit_tactic () in + Lib.add_anonymous_leaf (inImplicitTactic None) + +} + +VERNAC COMMAND EXTEND ImplicitTactic CLASSIFIED AS SIDEFF +| [ "Declare" "Implicit" "Tactic" tactic(tac) ] -> { declare_implicit_tactic tac } +| [ "Clear" "Implicit" "Tactic" ] -> { clear_implicit_tactic () } +END + + + + +(**********************************************************************) +(* sozeau: abs/gen for induction on instantiated dependent inductives, using "Ford" induction as + defined by Conor McBride *) +TACTIC EXTEND generalize_eqs +| ["generalize_eqs" hyp(id) ] -> { abstract_generalize ~generalize_vars:false id } +END +TACTIC EXTEND dep_generalize_eqs +| ["dependent" "generalize_eqs" hyp(id) ] -> { abstract_generalize ~generalize_vars:false ~force_dep:true id } +END +TACTIC EXTEND generalize_eqs_vars +| ["generalize_eqs_vars" hyp(id) ] -> { abstract_generalize ~generalize_vars:true id } +END +TACTIC EXTEND dep_generalize_eqs_vars +| ["dependent" "generalize_eqs_vars" hyp(id) ] -> { abstract_generalize ~force_dep:true ~generalize_vars:true id } +END + +(** Tactic to automatically simplify hypotheses of the form [Π Δ, x_i = t_i -> T] + where [t_i] is closed w.r.t. Δ. Such hypotheses are automatically generated + during dependent induction. For internal use. *) + +TACTIC EXTEND specialize_eqs +| [ "specialize_eqs" hyp(id) ] -> { specialize_eqs id } +END + +(**********************************************************************) +(* A tactic that considers a given occurrence of [c] in [t] and *) +(* abstract the minimal set of all the occurrences of [c] so that the *) +(* abstraction [fun x -> t[x/c]] is well-typed *) +(* *) +(* Contributed by Chung-Kil Hur (Winter 2009) *) +(**********************************************************************) + +{ + +let subst_var_with_hole occ tid t = + let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in + let locref = ref 0 in + let rec substrec x = match DAst.get x with + | GVar id -> + if Id.equal id tid + then + (decr occref; + if Int.equal !occref 0 then x + else + (incr locref; + DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ + GHole (Evar_kinds.QuestionMark { + Evar_kinds.qm_obligation=Evar_kinds.Define true; + Evar_kinds.qm_name=Anonymous; + Evar_kinds.qm_record_field=None; + }, IntroAnonymous, None))) + else x + | _ -> map_glob_constr_left_to_right substrec x in + let t' = substrec t + in + if !occref > 0 then Find_subterm.error_invalid_occurrence [occ] else t' + +let subst_hole_with_term occ tc t = + let locref = ref 0 in + let occref = ref occ in + let rec substrec c = match DAst.get c with + | GHole (Evar_kinds.QuestionMark { + Evar_kinds.qm_obligation=Evar_kinds.Define true; + Evar_kinds.qm_name=Anonymous; + Evar_kinds.qm_record_field=None; + }, IntroAnonymous, s) -> + decr occref; + if Int.equal !occref 0 then tc + else + (incr locref; + DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ + GHole (Evar_kinds.QuestionMark { + Evar_kinds.qm_obligation=Evar_kinds.Define true; + Evar_kinds.qm_name=Anonymous; + Evar_kinds.qm_record_field=None; + },IntroAnonymous,s)) + | _ -> map_glob_constr_left_to_right substrec c + in + substrec t + +open Tacmach + +let hResolve id c occ t = + Proofview.Goal.enter begin fun gl -> + let sigma = Proofview.Goal.sigma gl in + let env = Termops.clear_named_body id (Proofview.Goal.env gl) in + let concl = Proofview.Goal.concl gl in + let env_ids = Termops.vars_of_env env in + let c_raw = Detyping.detype Detyping.Now true env_ids env sigma c in + let t_raw = Detyping.detype Detyping.Now true env_ids env sigma t in + let rec resolve_hole t_hole = + try + Pretyping.understand env sigma t_hole + with + | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> + let (e, info) = CErrors.push e in + let loc_begin = Option.cata (fun l -> fst (Loc.unloc l)) 0 (Loc.get_loc info) in + resolve_hole (subst_hole_with_term loc_begin c_raw t_hole) + in + let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in + let sigma = Evd.merge_universe_context sigma ctx in + let t_constr_type = Retyping.get_type_of env sigma t_constr in + Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) + (change_concl (mkLetIn (Name.Anonymous,t_constr,t_constr_type,concl))) + end + +let hResolve_auto id c t = + let rec resolve_auto n = + try + hResolve id c n t + with + | UserError _ as e -> raise e + | e when CErrors.noncritical e -> resolve_auto (n+1) + in + resolve_auto 1 + +} + +TACTIC EXTEND hresolve_core +| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "at" int_or_var(occ) "in" constr(t) ] -> { hResolve id c occ t } +| [ "hresolve_core" "(" ident(id) ":=" constr(c) ")" "in" constr(t) ] -> { hResolve_auto id c t } +END + +(** + hget_evar +*) + +TACTIC EXTEND hget_evar +| [ "hget_evar" int_or_var(n) ] -> { Evar_tactics.hget_evar n } +END + +(**********************************************************************) + +(**********************************************************************) +(* A tactic that reduces one match t with ... by doing destruct t. *) +(* if t is not a variable, the tactic does *) +(* case_eq t;intros ... heq;rewrite heq in *|-. (but heq itself is *) +(* preserved). *) +(* Contributed by Julien Forest and Pierre Courtieu (july 2010) *) +(**********************************************************************) + +{ + +exception Found of unit Proofview.tactic + +let rewrite_except h = + Proofview.Goal.enter begin fun gl -> + let hyps = Tacmach.New.pf_ids_of_hyps gl in + Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else + Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) + hyps + end + + +let refl_equal () = Coqlib.lib_ref "core.eq.type" + +(* This is simply an implementation of the case_eq tactic. this code + should be replaced by a call to the tactic but I don't know how to + call it before it is defined. *) +let mkCaseEq a : unit Proofview.tactic = + Proofview.Goal.enter begin fun gl -> + let type_of_a = Tacmach.New.pf_unsafe_type_of gl a in + Tacticals.New.pf_constr_of_global (delayed_force refl_equal) >>= fun req -> + Tacticals.New.tclTHENLIST + [Tactics.generalize [(mkApp(req, [| type_of_a; a|]))]; + Proofview.Goal.enter begin fun gl -> + let concl = Proofview.Goal.concl gl in + let env = Proofview.Goal.env gl in + (** FIXME: this looks really wrong. Does anybody really use this tactic? *) + let (_, c) = Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] env (Evd.from_env env) concl in + change_concl c + end; + simplest_case a] + end + + +let case_eq_intros_rewrite x = + Proofview.Goal.enter begin fun gl -> + let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in + (* Pp.msgnl (Printer.pr_lconstr x); *) + Tacticals.New.tclTHENLIST [ + mkCaseEq x; + Proofview.Goal.enter begin fun gl -> + let concl = Proofview.Goal.concl gl in + let hyps = Tacmach.New.pf_ids_set_of_hyps gl in + let n' = nb_prod (Tacmach.New.project gl) concl in + let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in + Tacticals.New.tclTHENLIST [ + Tacticals.New.tclDO (n'-n-1) intro; + introduction h; + rewrite_except h] + end + ] + end + +let rec find_a_destructable_match sigma t = + let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in + let cl = [cl, (None, None), None], None in + let dest = TacAtom (Loc.tag @@ TacInductionDestruct(false, false, cl)) in + match EConstr.kind sigma t with + | Case (_,_,x,_) when closed0 sigma x -> + if isVar sigma x then + (* TODO check there is no rel n. *) + raise (Found (Tacinterp.eval_tactic dest)) + else + (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) + raise (Found (case_eq_intros_rewrite x)) + | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t + + +let destauto t = + Proofview.tclEVARMAP >>= fun sigma -> + try find_a_destructable_match sigma t; + Tacticals.New.tclZEROMSG (str "No destructable match found") + with Found tac -> tac + +let destauto_in id = + Proofview.Goal.enter begin fun gl -> + let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in +(* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) +(* Pp.msgnl (Printer.pr_lconstr (ctype)); *) + destauto ctype + end + +} + +TACTIC EXTEND destauto +| [ "destauto" ] -> { Proofview.Goal.enter begin fun gl -> destauto (Proofview.Goal.concl gl) end } +| [ "destauto" "in" hyp(id) ] -> { destauto_in id } +END + +(**********************************************************************) + +(**********************************************************************) +(* A version of abstract constructing transparent terms *) +(* Introduced by Jason Gross and Benjamin Delaware in June 2016 *) +(**********************************************************************) + +TACTIC EXTEND transparent_abstract +| [ "transparent_abstract" tactic3(t) ] -> { Proofview.Goal.enter begin fun gl -> + Tactics.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) end } +| [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> { Proofview.Goal.enter begin fun gl -> + Tactics.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) end } +END + +(* ********************************************************************* *) + +TACTIC EXTEND constr_eq +| [ "constr_eq" constr(x) constr(y) ] -> { Tactics.constr_eq ~strict:false x y } +END + +TACTIC EXTEND constr_eq_strict +| [ "constr_eq_strict" constr(x) constr(y) ] -> { Tactics.constr_eq ~strict:true x y } +END + +TACTIC EXTEND constr_eq_nounivs +| [ "constr_eq_nounivs" constr(x) constr(y) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") } +END + +TACTIC EXTEND is_evar +| [ "is_evar" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Evar _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") + } +END + +TACTIC EXTEND has_evar +| [ "has_evar" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + if Evarutil.has_undefined_evars sigma x + then Proofview.tclUNIT () + else Tacticals.New.tclFAIL 0 (str "No evars") +} +END + +TACTIC EXTEND is_hyp +| [ "is_var" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Var _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") } +END + +TACTIC EXTEND is_fix +| [ "is_fix" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Fix _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") } +END + +TACTIC EXTEND is_cofix +| [ "is_cofix" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | CoFix _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") } +END + +TACTIC EXTEND is_ind +| [ "is_ind" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Ind _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") } +END + +TACTIC EXTEND is_constructor +| [ "is_constructor" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Construct _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") } +END + +TACTIC EXTEND is_proj +| [ "is_proj" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Proj _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") } +END + +TACTIC EXTEND is_const +| [ "is_const" constr(x) ] -> { + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma x with + | Const _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") } +END + +(* Command to grab the evars left unresolved at the end of a proof. *) +(* spiwack: I put it in extratactics because it is somewhat tied with + the semantics of the LCF-style tactics, hence with the classic tactic + mode. *) +VERNAC COMMAND EXTEND GrabEvars +| [ "Grab" "Existential" "Variables" ] + => { Vernac_classifier.classify_as_proofstep } + -> { Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.grab_evars p) } +END + +(* Shelves all the goals under focus. *) +TACTIC EXTEND shelve +| [ "shelve" ] -> + { Proofview.shelve } +END + +(* Shelves the unifiable goals under focus, i.e. the goals which + appear in other goals under focus (the unfocused goals are not + considered). *) +TACTIC EXTEND shelve_unifiable +| [ "shelve_unifiable" ] -> + { Proofview.shelve_unifiable } +END + +(* Unshelves the goal shelved by the tactic. *) +TACTIC EXTEND unshelve +| [ "unshelve" tactic1(t) ] -> + { + Proofview.with_shelf (Tacinterp.tactic_of_value ist t) >>= fun (gls, ()) -> + let gls = List.map Proofview.with_empty_state gls in + Proofview.Unsafe.tclGETGOALS >>= fun ogls -> + Proofview.Unsafe.tclSETGOALS (gls @ ogls) + } +END + +(* Command to add every unshelved variables to the focus *) +VERNAC COMMAND EXTEND Unshelve +| [ "Unshelve" ] + => { Vernac_classifier.classify_as_proofstep } + -> { Proof_global.simple_with_current_proof (fun _ p -> Proof.unshelve p) } +END + +(* Gives up on the goals under focus: the goals are considered solved, + but the proof cannot be closed until the user goes back and solve + these goals. *) +TACTIC EXTEND give_up +| [ "give_up" ] -> + { Proofview.give_up } +END + +(* cycles [n] goals *) +TACTIC EXTEND cycle +| [ "cycle" int_or_var(n) ] -> { Proofview.cycle n } +END + +(* swaps goals number [i] and [j] *) +TACTIC EXTEND swap +| [ "swap" int_or_var(i) int_or_var(j) ] -> { Proofview.swap i j } +END + +(* reverses the list of focused goals *) +TACTIC EXTEND revgoals +| [ "revgoals" ] -> { Proofview.revgoals } +END + +{ + +type cmp = + | Eq + | Lt | Le + | Gt | Ge + +type 'i test = + | Test of cmp * 'i * 'i + +let pr_cmp = function + | Eq -> Pp.str"=" + | Lt -> Pp.str"<" + | Le -> Pp.str"<=" + | Gt -> Pp.str">" + | Ge -> Pp.str">=" + +let pr_cmp' _prc _prlc _prt = pr_cmp + +let pr_test_gen f (Test(c,x,y)) = + Pp.(f x ++ pr_cmp c ++ f y) + +let pr_test = pr_test_gen (Pputils.pr_or_var Pp.int) + +let pr_test' _prc _prlc _prt = pr_test + +let pr_itest = pr_test_gen Pp.int + +let pr_itest' _prc _prlc _prt = pr_itest + +} + +ARGUMENT EXTEND comparison PRINTED BY { pr_cmp' } +| [ "=" ] -> { Eq } +| [ "<" ] -> { Lt } +| [ "<=" ] -> { Le } +| [ ">" ] -> { Gt } +| [ ">=" ] -> { Ge } + END + +{ + +let interp_test ist gls = function + | Test (c,x,y) -> + project gls , + Test(c,Tacinterp.interp_int_or_var ist x,Tacinterp.interp_int_or_var ist y) + +} + +ARGUMENT EXTEND test + PRINTED BY { pr_itest' } + INTERPRETED BY { interp_test } + RAW_PRINTED BY { pr_test' } + GLOB_PRINTED BY { pr_test' } +| [ int_or_var(x) comparison(c) int_or_var(y) ] -> { Test(c,x,y) } +END + +{ + +let interp_cmp = function + | Eq -> Int.equal + | Lt -> ((<):int->int->bool) + | Le -> ((<=):int->int->bool) + | Gt -> ((>):int->int->bool) + | Ge -> ((>=):int->int->bool) + +let run_test = function + | Test(c,x,y) -> interp_cmp c x y + +let guard tst = + if run_test tst then + Proofview.tclUNIT () + else + let msg = Pp.(str"Condition not satisfied:"++ws 1++(pr_itest tst)) in + Tacticals.New.tclZEROMSG msg + +} + +TACTIC EXTEND guard +| [ "guard" test(tst) ] -> { guard tst } +END + +{ + +let decompose l c = + Proofview.Goal.enter begin fun gl -> + let sigma = Tacmach.New.project gl in + let to_ind c = + if isInd sigma c then fst (destInd sigma c) + else user_err Pp.(str "not an inductive type") + in + let l = List.map to_ind l in + Elim.h_decompose l c + end + +} + +TACTIC EXTEND decompose +| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> { decompose l c } +END + +(** library/keys *) + +VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF +| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> { + let get_key c = + let env = Global.env () in + let evd = Evd.from_env env in + let (evd, c) = Constrintern.interp_open_constr env evd c in + let kind c = EConstr.kind evd c in + Keys.constr_key kind c + in + let k1 = get_key c in + let k2 = get_key c' in + match k1, k2 with + | Some k1, Some k2 -> Keys.declare_equiv_keys k1 k2 + | _ -> () } +END + +VERNAC COMMAND EXTEND Print_keys CLASSIFIED AS QUERY +| [ "Print" "Equivalent" "Keys" ] -> { Feedback.msg_info (Keys.pr_keys Printer.pr_global) } +END + + +VERNAC COMMAND EXTEND OptimizeProof +| [ "Optimize" "Proof" ] => { Vernac_classifier.classify_as_proofstep } -> + { Proof_global.compact_the_proof () } +| [ "Optimize" "Heap" ] => { Vernac_classifier.classify_as_proofstep } -> + { Gc.compact () } +END + +(** tactic analogous to "OPTIMIZE HEAP" *) + +{ + +let tclOPTIMIZE_HEAP = + Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> Gc.compact ())) + +} + +TACTIC EXTEND optimize_heap +| [ "optimize_heap" ] -> { tclOPTIMIZE_HEAP } +END diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4 deleted file mode 100644 index 6913543c9a..0000000000 --- a/plugins/ltac/g_auto.ml4 +++ /dev/null @@ -1,226 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* [ Eauto.e_assumption ] -END - -TACTIC EXTEND eexact -| [ "eexact" constr(c) ] -> [ Eauto.e_give_exact c ] -END - -let pr_hintbases _prc _prlc _prt = Pptactic.pr_hintbases - -ARGUMENT EXTEND hintbases - TYPED AS preident_list_opt - PRINTED BY pr_hintbases -| [ "with" "*" ] -> [ None ] -| [ "with" ne_preident_list(l) ] -> [ Some l ] -| [ ] -> [ Some [] ] -END - -let eval_uconstrs ist cs = - let flags = { - Pretyping.use_typeclasses = false; - solve_unification_constraints = true; - use_hook = Pfedit.solve_by_implicit_tactic (); - fail_evar = false; - expand_evars = true - } in - let map c env sigma = c env sigma in - List.map (fun c -> map (Tacinterp.type_uconstr ~flags ist c)) cs - -let pr_auto_using_raw _ _ _ = Pptactic.pr_auto_using Ppconstr.pr_constr_expr -let pr_auto_using_glob _ _ _ = Pptactic.pr_auto_using (fun (c,_) -> - let _, env = Pfedit.get_current_context () in - Printer.pr_glob_constr_env env c) -let pr_auto_using _ _ _ = Pptactic.pr_auto_using - (let sigma, env = Pfedit.get_current_context () in - Printer.pr_closed_glob_env env sigma) - -ARGUMENT EXTEND auto_using - TYPED AS uconstr_list - PRINTED BY pr_auto_using - RAW_PRINTED BY pr_auto_using_raw - GLOB_PRINTED BY pr_auto_using_glob -| [ "using" ne_uconstr_list_sep(l, ",") ] -> [ l ] -| [ ] -> [ [] ] -END - -(** Auto *) - -TACTIC EXTEND trivial -| [ "trivial" auto_using(lems) hintbases(db) ] -> - [ Auto.h_trivial (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND info_trivial -| [ "info_trivial" auto_using(lems) hintbases(db) ] -> - [ Auto.h_trivial ~debug:Info (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND debug_trivial -| [ "debug" "trivial" auto_using(lems) hintbases(db) ] -> - [ Auto.h_trivial ~debug:Debug (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND auto -| [ "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> - [ Auto.h_auto n (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND info_auto -| [ "info_auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> - [ Auto.h_auto ~debug:Info n (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND debug_auto -| [ "debug" "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> - [ Auto.h_auto ~debug:Debug n (eval_uconstrs ist lems) db ] -END - -(** Eauto *) - -TACTIC EXTEND prolog -| [ "prolog" "[" uconstr_list(l) "]" int_or_var(n) ] -> - [ Eauto.prolog_tac (eval_uconstrs ist l) n ] -END - -let make_depth n = snd (Eauto.make_dimension n None) - -TACTIC EXTEND eauto -| [ "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) - hintbases(db) ] -> - [ Eauto.gen_eauto (Eauto.make_dimension n p) (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND new_eauto -| [ "new" "auto" int_or_var_opt(n) auto_using(lems) - hintbases(db) ] -> - [ match db with - | None -> Auto.new_full_auto (make_depth n) (eval_uconstrs ist lems) - | Some l -> Auto.new_auto (make_depth n) (eval_uconstrs ist lems) l ] -END - -TACTIC EXTEND debug_eauto -| [ "debug" "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) - hintbases(db) ] -> - [ Eauto.gen_eauto ~debug:Debug (Eauto.make_dimension n p) (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND info_eauto -| [ "info_eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) - hintbases(db) ] -> - [ Eauto.gen_eauto ~debug:Info (Eauto.make_dimension n p) (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND dfs_eauto -| [ "dfs" "eauto" int_or_var_opt(p) auto_using(lems) - hintbases(db) ] -> - [ Eauto.gen_eauto (Eauto.make_dimension p None) (eval_uconstrs ist lems) db ] -END - -TACTIC EXTEND autounfold -| [ "autounfold" hintbases(db) clause_dft_concl(cl) ] -> [ Eauto.autounfold_tac db cl ] -END - -TACTIC EXTEND autounfold_one -| [ "autounfold_one" hintbases(db) "in" hyp(id) ] -> - [ Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) (Some (id, Locus.InHyp)) ] -| [ "autounfold_one" hintbases(db) ] -> - [ Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) None ] - END - -TACTIC EXTEND unify -| ["unify" constr(x) constr(y) ] -> [ Tactics.unify x y ] -| ["unify" constr(x) constr(y) "with" preident(base) ] -> [ - let table = try Some (Hints.searchtable_map base) with Not_found -> None in - match table with - | None -> - let msg = str "Hint table " ++ str base ++ str " not found" in - Tacticals.New.tclZEROMSG msg - | Some t -> - let state = Hints.Hint_db.transparent_state t in - Tactics.unify ~state x y - ] -END - - -TACTIC EXTEND convert_concl_no_check -| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check x DEFAULTcast ] -END - -let pr_pre_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Libnames.pr_qualid -let pr_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Printer.pr_global -let glob_hints_path_atom ist = Hints.glob_hints_path_atom - -ARGUMENT EXTEND hints_path_atom - PRINTED BY pr_hints_path_atom - - GLOBALIZED BY glob_hints_path_atom - - RAW_PRINTED BY pr_pre_hints_path_atom - GLOB_PRINTED BY pr_hints_path_atom -| [ ne_global_list(g) ] -> [ Hints.PathHints g ] -| [ "_" ] -> [ Hints.PathAny ] -END - -let pr_hints_path prc prx pry c = Hints.pp_hints_path c -let pr_pre_hints_path prc prx pry c = Hints.pp_hints_path_gen Libnames.pr_qualid c -let glob_hints_path ist = Hints.glob_hints_path - -ARGUMENT EXTEND hints_path -PRINTED BY pr_hints_path - -GLOBALIZED BY glob_hints_path -RAW_PRINTED BY pr_pre_hints_path -GLOB_PRINTED BY pr_hints_path - -| [ "(" hints_path(p) ")" ] -> [ p ] -| [ hints_path(p) "*" ] -> [ Hints.PathStar p ] -| [ "emp" ] -> [ Hints.PathEmpty ] -| [ "eps" ] -> [ Hints.PathEpsilon ] -| [ hints_path(p) "|" hints_path(q) ] -> [ Hints.PathOr (p, q) ] -| [ hints_path_atom(a) ] -> [ Hints.PathAtom a ] -| [ hints_path(p) hints_path(q) ] -> [ Hints.PathSeq (p, q) ] -END - -ARGUMENT EXTEND opthints - TYPED AS preident_list_opt - PRINTED BY pr_hintbases -| [ ":" ne_preident_list(l) ] -> [ Some l ] -| [ ] -> [ None ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND HintCut CLASSIFIED AS SIDEFF -| [ "Hint" "Cut" "[" hints_path(p) "]" opthints(dbnames) ] -> [ - fun ~atts ~st -> begin - let open Vernacinterp in - let entry = Hints.HintsCutEntry (Hints.glob_hints_path p) in - Hints.add_hints ~local:(Locality.make_section_locality atts.locality) - (match dbnames with None -> ["core"] | Some l -> l) entry; - st - end - ] -END - diff --git a/plugins/ltac/g_auto.mlg b/plugins/ltac/g_auto.mlg new file mode 100644 index 0000000000..c07b653f3a --- /dev/null +++ b/plugins/ltac/g_auto.mlg @@ -0,0 +1,249 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* { Eauto.e_assumption } +END + +TACTIC EXTEND eexact +| [ "eexact" constr(c) ] -> { Eauto.e_give_exact c } +END + +{ + +let pr_hintbases _prc _prlc _prt = Pptactic.pr_hintbases + +} + +ARGUMENT EXTEND hintbases + TYPED AS preident list option + PRINTED BY { pr_hintbases } +| [ "with" "*" ] -> { None } +| [ "with" ne_preident_list(l) ] -> { Some l } +| [ ] -> { Some [] } +END + +{ + +let eval_uconstrs ist cs = + let flags = { + Pretyping.use_typeclasses = false; + solve_unification_constraints = true; + use_hook = Pfedit.solve_by_implicit_tactic (); + fail_evar = false; + expand_evars = true + } in + let map c env sigma = c env sigma in + List.map (fun c -> map (Tacinterp.type_uconstr ~flags ist c)) cs + +let pr_auto_using_raw _ _ _ = Pptactic.pr_auto_using Ppconstr.pr_constr_expr +let pr_auto_using_glob _ _ _ = Pptactic.pr_auto_using (fun (c,_) -> + let _, env = Pfedit.get_current_context () in + Printer.pr_glob_constr_env env c) +let pr_auto_using _ _ _ = Pptactic.pr_auto_using + (let sigma, env = Pfedit.get_current_context () in + Printer.pr_closed_glob_env env sigma) + +} + +ARGUMENT EXTEND auto_using + TYPED AS uconstr list + PRINTED BY { pr_auto_using } + RAW_PRINTED BY { pr_auto_using_raw } + GLOB_PRINTED BY { pr_auto_using_glob } +| [ "using" ne_uconstr_list_sep(l, ",") ] -> { l } +| [ ] -> { [] } +END + +(** Auto *) + +TACTIC EXTEND trivial +| [ "trivial" auto_using(lems) hintbases(db) ] -> + { Auto.h_trivial (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND info_trivial +| [ "info_trivial" auto_using(lems) hintbases(db) ] -> + { Auto.h_trivial ~debug:Info (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND debug_trivial +| [ "debug" "trivial" auto_using(lems) hintbases(db) ] -> + { Auto.h_trivial ~debug:Debug (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND auto +| [ "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> + { Auto.h_auto n (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND info_auto +| [ "info_auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> + { Auto.h_auto ~debug:Info n (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND debug_auto +| [ "debug" "auto" int_or_var_opt(n) auto_using(lems) hintbases(db) ] -> + { Auto.h_auto ~debug:Debug n (eval_uconstrs ist lems) db } +END + +(** Eauto *) + +TACTIC EXTEND prolog +| [ "prolog" "[" uconstr_list(l) "]" int_or_var(n) ] -> + { Eauto.prolog_tac (eval_uconstrs ist l) n } +END + +{ + +let make_depth n = snd (Eauto.make_dimension n None) + +} + +TACTIC EXTEND eauto +| [ "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) + hintbases(db) ] -> + { Eauto.gen_eauto (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND new_eauto +| [ "new" "auto" int_or_var_opt(n) auto_using(lems) + hintbases(db) ] -> + { match db with + | None -> Auto.new_full_auto (make_depth n) (eval_uconstrs ist lems) + | Some l -> Auto.new_auto (make_depth n) (eval_uconstrs ist lems) l } +END + +TACTIC EXTEND debug_eauto +| [ "debug" "eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) + hintbases(db) ] -> + { Eauto.gen_eauto ~debug:Debug (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND info_eauto +| [ "info_eauto" int_or_var_opt(n) int_or_var_opt(p) auto_using(lems) + hintbases(db) ] -> + { Eauto.gen_eauto ~debug:Info (Eauto.make_dimension n p) (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND dfs_eauto +| [ "dfs" "eauto" int_or_var_opt(p) auto_using(lems) + hintbases(db) ] -> + { Eauto.gen_eauto (Eauto.make_dimension p None) (eval_uconstrs ist lems) db } +END + +TACTIC EXTEND autounfold +| [ "autounfold" hintbases(db) clause_dft_concl(cl) ] -> { Eauto.autounfold_tac db cl } +END + +TACTIC EXTEND autounfold_one +| [ "autounfold_one" hintbases(db) "in" hyp(id) ] -> + { Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) (Some (id, Locus.InHyp)) } +| [ "autounfold_one" hintbases(db) ] -> + { Eauto.autounfold_one (match db with None -> ["core"] | Some x -> "core"::x) None } + END + +TACTIC EXTEND unify +| ["unify" constr(x) constr(y) ] -> { Tactics.unify x y } +| ["unify" constr(x) constr(y) "with" preident(base) ] -> { + let table = try Some (Hints.searchtable_map base) with Not_found -> None in + match table with + | None -> + let msg = str "Hint table " ++ str base ++ str " not found" in + Tacticals.New.tclZEROMSG msg + | Some t -> + let state = Hints.Hint_db.transparent_state t in + Tactics.unify ~state x y + } +END + + +TACTIC EXTEND convert_concl_no_check +| ["convert_concl_no_check" constr(x) ] -> { Tactics.convert_concl_no_check x DEFAULTcast } +END + +{ + +let pr_pre_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Libnames.pr_qualid +let pr_hints_path_atom _ _ _ = Hints.pp_hints_path_atom Printer.pr_global +let glob_hints_path_atom ist = Hints.glob_hints_path_atom + +} + +ARGUMENT EXTEND hints_path_atom + PRINTED BY { pr_hints_path_atom } + + GLOBALIZED BY { glob_hints_path_atom } + + RAW_PRINTED BY { pr_pre_hints_path_atom } + GLOB_PRINTED BY { pr_hints_path_atom } +| [ ne_global_list(g) ] -> { Hints.PathHints g } +| [ "_" ] -> { Hints.PathAny } +END + +{ + +let pr_hints_path prc prx pry c = Hints.pp_hints_path c +let pr_pre_hints_path prc prx pry c = Hints.pp_hints_path_gen Libnames.pr_qualid c +let glob_hints_path ist = Hints.glob_hints_path + +} + +ARGUMENT EXTEND hints_path +PRINTED BY { pr_hints_path } + +GLOBALIZED BY { glob_hints_path } +RAW_PRINTED BY { pr_pre_hints_path } +GLOB_PRINTED BY { pr_hints_path } + +| [ "(" hints_path(p) ")" ] -> { p } +| [ hints_path(p) "*" ] -> { Hints.PathStar p } +| [ "emp" ] -> { Hints.PathEmpty } +| [ "eps" ] -> { Hints.PathEpsilon } +| [ hints_path(p) "|" hints_path(q) ] -> { Hints.PathOr (p, q) } +| [ hints_path_atom(a) ] -> { Hints.PathAtom a } +| [ hints_path(p) hints_path(q) ] -> { Hints.PathSeq (p, q) } +END + +ARGUMENT EXTEND opthints + TYPED AS preident list option + PRINTED BY { pr_hintbases } +| [ ":" ne_preident_list(l) ] -> { Some l } +| [ ] -> { None } +END + +VERNAC COMMAND EXTEND HintCut CLASSIFIED AS SIDEFF +| [ "Hint" "Cut" "[" hints_path(p) "]" opthints(dbnames) ] -> { + let open Vernacinterp in + let entry = Hints.HintsCutEntry (Hints.glob_hints_path p) in + Hints.add_hints ~local:(Locality.make_section_locality atts.locality) + (match dbnames with None -> ["core"] | Some l -> l) entry; + } +END + diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4 deleted file mode 100644 index 265368833b..0000000000 --- a/plugins/ltac/g_class.ml4 +++ /dev/null @@ -1,117 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* - let gr = Smartlocate.global_with_alias r in - let ev = Tacred.evaluable_of_global_reference (Global.env ()) gr in - Classes.set_typeclass_transparency ev (Locality.make_section_locality None) b) cl - -VERNAC COMMAND EXTEND Typeclasses_Unfold_Settings CLASSIFIED AS SIDEFF -| [ "Typeclasses" "Transparent" reference_list(cl) ] -> [ - set_transparency cl true ] -END - -VERNAC COMMAND EXTEND Typeclasses_Rigid_Settings CLASSIFIED AS SIDEFF -| [ "Typeclasses" "Opaque" reference_list(cl) ] -> [ - set_transparency cl false ] -END - -let pr_debug _prc _prlc _prt b = - if b then Pp.str "debug" else Pp.mt() - -ARGUMENT EXTEND debug TYPED AS bool PRINTED BY pr_debug -| [ "debug" ] -> [ true ] -| [ ] -> [ false ] -END - -let pr_search_strategy _prc _prlc _prt = function - | Some Dfs -> Pp.str "dfs" - | Some Bfs -> Pp.str "bfs" - | None -> Pp.mt () - -ARGUMENT EXTEND eauto_search_strategy PRINTED BY pr_search_strategy -| [ "(bfs)" ] -> [ Some Bfs ] -| [ "(dfs)" ] -> [ Some Dfs ] -| [ ] -> [ None ] -END - -(* true = All transparent, false = Opaque if possible *) - -VERNAC COMMAND EXTEND Typeclasses_Settings CLASSIFIED AS SIDEFF - | [ "Typeclasses" "eauto" ":=" debug(d) eauto_search_strategy(s) int_opt(depth) ] -> [ - set_typeclasses_debug d; - Option.iter set_typeclasses_strategy s; - set_typeclasses_depth depth - ] -END - -(** Compatibility: typeclasses eauto has 8.5 and 8.6 modes *) -TACTIC EXTEND typeclasses_eauto - | [ "typeclasses" "eauto" "bfs" int_or_var_opt(d) "with" ne_preident_list(l) ] -> - [ typeclasses_eauto ~strategy:Bfs ~depth:d l ] - | [ "typeclasses" "eauto" int_or_var_opt(d) "with" ne_preident_list(l) ] -> - [ typeclasses_eauto ~depth:d l ] - | [ "typeclasses" "eauto" int_or_var_opt(d) ] -> [ - typeclasses_eauto ~only_classes:true ~depth:d [Hints.typeclasses_db] ] -END - -TACTIC EXTEND head_of_constr - [ "head_of_constr" ident(h) constr(c) ] -> [ head_of_constr h c ] -END - -TACTIC EXTEND not_evar - [ "not_evar" constr(ty) ] -> [ not_evar ty ] -END - -TACTIC EXTEND is_ground - [ "is_ground" constr(ty) ] -> [ is_ground ty ] -END - -TACTIC EXTEND autoapply - [ "autoapply" constr(c) "using" preident(i) ] -> [ autoapply c i ] -END - -(** TODO: DEPRECATE *) -(* A progress test that allows to see if the evars have changed *) -open Constr -open Proofview.Notations - -let rec eq_constr_mod_evars sigma x y = - let open EConstr in - match EConstr.kind sigma x, EConstr.kind sigma y with - | Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true - | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y - -let progress_evars t = - Proofview.Goal.enter begin fun gl -> - let concl = Proofview.Goal.concl gl in - let check = - Proofview.Goal.enter begin fun gl' -> - let sigma = Tacmach.New.project gl' in - let newconcl = Proofview.Goal.concl gl' in - if eq_constr_mod_evars sigma concl newconcl - then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)") - else Proofview.tclUNIT () - end - in t <*> check - end - -TACTIC EXTEND progress_evars - [ "progress_evars" tactic(t) ] -> [ progress_evars (Tacinterp.tactic_of_value ist t) ] -END diff --git a/plugins/ltac/g_class.mlg b/plugins/ltac/g_class.mlg new file mode 100644 index 0000000000..9ecc36bdf3 --- /dev/null +++ b/plugins/ltac/g_class.mlg @@ -0,0 +1,137 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + let gr = Smartlocate.global_with_alias r in + let ev = Tacred.evaluable_of_global_reference (Global.env ()) gr in + Classes.set_typeclass_transparency ev (Locality.make_section_locality None) b) cl + +} + +VERNAC COMMAND EXTEND Typeclasses_Unfold_Settings CLASSIFIED AS SIDEFF +| [ "Typeclasses" "Transparent" reference_list(cl) ] -> { + set_transparency cl true } +END + +VERNAC COMMAND EXTEND Typeclasses_Rigid_Settings CLASSIFIED AS SIDEFF +| [ "Typeclasses" "Opaque" reference_list(cl) ] -> { + set_transparency cl false } +END + +{ + +let pr_debug _prc _prlc _prt b = + if b then Pp.str "debug" else Pp.mt() + +} + +ARGUMENT EXTEND debug TYPED AS bool PRINTED BY { pr_debug } +| [ "debug" ] -> { true } +| [ ] -> { false } +END + +{ + +let pr_search_strategy _prc _prlc _prt = function + | Some Dfs -> Pp.str "dfs" + | Some Bfs -> Pp.str "bfs" + | None -> Pp.mt () + +} + +ARGUMENT EXTEND eauto_search_strategy PRINTED BY { pr_search_strategy } +| [ "(bfs)" ] -> { Some Bfs } +| [ "(dfs)" ] -> { Some Dfs } +| [ ] -> { None } +END + +(* true = All transparent, false = Opaque if possible *) + +VERNAC COMMAND EXTEND Typeclasses_Settings CLASSIFIED AS SIDEFF + | [ "Typeclasses" "eauto" ":=" debug(d) eauto_search_strategy(s) int_opt(depth) ] -> { + set_typeclasses_debug d; + Option.iter set_typeclasses_strategy s; + set_typeclasses_depth depth + } +END + +(** Compatibility: typeclasses eauto has 8.5 and 8.6 modes *) +TACTIC EXTEND typeclasses_eauto + | [ "typeclasses" "eauto" "bfs" int_or_var_opt(d) "with" ne_preident_list(l) ] -> + { typeclasses_eauto ~strategy:Bfs ~depth:d l } + | [ "typeclasses" "eauto" int_or_var_opt(d) "with" ne_preident_list(l) ] -> + { typeclasses_eauto ~depth:d l } + | [ "typeclasses" "eauto" int_or_var_opt(d) ] -> { + typeclasses_eauto ~only_classes:true ~depth:d [Hints.typeclasses_db] } +END + +TACTIC EXTEND head_of_constr +| [ "head_of_constr" ident(h) constr(c) ] -> { head_of_constr h c } +END + +TACTIC EXTEND not_evar +| [ "not_evar" constr(ty) ] -> { not_evar ty } +END + +TACTIC EXTEND is_ground +| [ "is_ground" constr(ty) ] -> { is_ground ty } +END + +TACTIC EXTEND autoapply +| [ "autoapply" constr(c) "using" preident(i) ] -> { autoapply c i } +END + +{ + +(** TODO: DEPRECATE *) +(* A progress test that allows to see if the evars have changed *) +open Constr +open Proofview.Notations + +let rec eq_constr_mod_evars sigma x y = + let open EConstr in + match EConstr.kind sigma x, EConstr.kind sigma y with + | Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true + | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y + +let progress_evars t = + Proofview.Goal.enter begin fun gl -> + let concl = Proofview.Goal.concl gl in + let check = + Proofview.Goal.enter begin fun gl' -> + let sigma = Tacmach.New.project gl' in + let newconcl = Proofview.Goal.concl gl' in + if eq_constr_mod_evars sigma concl newconcl + then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)") + else Proofview.tclUNIT () + end + in t <*> check + end + +} + +TACTIC EXTEND progress_evars +| [ "progress_evars" tactic(t) ] -> { progress_evars (Tacinterp.tactic_of_value ist t) } +END diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 deleted file mode 100644 index 929390b1c4..0000000000 --- a/plugins/ltac/g_ltac.ml4 +++ /dev/null @@ -1,528 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* a - | e -> Tacexp (e:raw_tactic_expr) - -let genarg_of_unit () = in_gen (rawwit Stdarg.wit_unit) () -let genarg_of_int n = in_gen (rawwit Stdarg.wit_int) n -let genarg_of_ipattern pat = in_gen (rawwit Tacarg.wit_intro_pattern) pat -let genarg_of_uconstr c = in_gen (rawwit Stdarg.wit_uconstr) c -let in_tac tac = in_gen (rawwit Tacarg.wit_ltac) tac - -let reference_to_id qid = - if Libnames.qualid_is_ident qid then - CAst.make ?loc:qid.CAst.loc @@ Libnames.qualid_basename qid - else - CErrors.user_err ?loc:qid.CAst.loc - (str "This expression should be a simple identifier.") - -let tactic_mode = Entry.create "vernac:tactic_command" - -let new_entry name = - let e = Entry.create name in - e - -let toplevel_selector = new_entry "vernac:toplevel_selector" -let tacdef_body = new_entry "tactic:tacdef_body" - -(* Registers the Classic Proof Mode (which uses [tactic_mode] as a parser for - proof editing and changes nothing else). Then sets it as the default proof mode. *) -let _ = - let mode = { - Proof_global.name = "Classic"; - set = (fun () -> Pvernac.set_command_entry tactic_mode); - reset = (fun () -> Pvernac.(set_command_entry noedit_mode)); - } in - Proof_global.register_proof_mode mode - -(* Hack to parse "[ id" without dropping [ *) -let test_bracket_ident = - Gram.Entry.of_parser "test_bracket_ident" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "[" -> - (match stream_nth 1 strm with - | IDENT _ -> () - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - -(* Tactics grammar rules *) - -let hint = G_proofs.hint - -GEXTEND Gram - GLOBAL: tactic tacdef_body tactic_expr binder_tactic tactic_arg command hint - tactic_mode constr_may_eval constr_eval toplevel_selector - operconstr; - - tactic_then_last: - [ [ "|"; lta = LIST0 OPT tactic_expr SEP "|" -> - Array.map (function None -> TacId [] | Some t -> t) (Array.of_list lta) - | -> [||] - ] ] - ; - tactic_then_gen: - [ [ ta = tactic_expr; "|"; (first,last) = tactic_then_gen -> (ta::first, last) - | ta = tactic_expr; ".."; l = tactic_then_last -> ([], Some (ta, l)) - | ".."; l = tactic_then_last -> ([], Some (TacId [], l)) - | ta = tactic_expr -> ([ta], None) - | "|"; (first,last) = tactic_then_gen -> (TacId [] :: first, last) - | -> ([TacId []], None) - ] ] - ; - tactic_then_locality: (* [true] for the local variant [TacThens] and [false] - for [TacExtend] *) - [ [ "[" ; l = OPT">" -> if Option.is_empty l then true else false ] ] - ; - tactic_expr: - [ "5" RIGHTA - [ te = binder_tactic -> te ] - | "4" LEFTA - [ ta0 = tactic_expr; ";"; ta1 = binder_tactic -> TacThen (ta0, ta1) - | ta0 = tactic_expr; ";"; ta1 = tactic_expr -> TacThen (ta0,ta1) - | ta0 = tactic_expr; ";"; l = tactic_then_locality; (first,tail) = tactic_then_gen; "]" -> - match l , tail with - | false , Some (t,last) -> TacThen (ta0,TacExtendTac (Array.of_list first, t, last)) - | true , Some (t,last) -> TacThens3parts (ta0, Array.of_list first, t, last) - | false , None -> TacThen (ta0,TacDispatch first) - | true , None -> TacThens (ta0,first) ] - | "3" RIGHTA - [ IDENT "try"; ta = tactic_expr -> TacTry ta - | IDENT "do"; n = int_or_var; ta = tactic_expr -> TacDo (n,ta) - | IDENT "timeout"; n = int_or_var; ta = tactic_expr -> TacTimeout (n,ta) - | IDENT "time"; s = OPT string; ta = tactic_expr -> TacTime (s,ta) - | IDENT "repeat"; ta = tactic_expr -> TacRepeat ta - | IDENT "progress"; ta = tactic_expr -> TacProgress ta - | IDENT "once"; ta = tactic_expr -> TacOnce ta - | IDENT "exactly_once"; ta = tactic_expr -> TacExactlyOnce ta - | IDENT "infoH"; ta = tactic_expr -> TacShowHyps ta -(*To do: put Abstract in Refiner*) - | IDENT "abstract"; tc = NEXT -> TacAbstract (tc,None) - | IDENT "abstract"; tc = NEXT; "using"; s = ident -> - TacAbstract (tc,Some s) - | sel = selector; ta = tactic_expr -> TacSelect (sel, ta) ] -(*End of To do*) - | "2" RIGHTA - [ ta0 = tactic_expr; "+"; ta1 = binder_tactic -> TacOr (ta0,ta1) - | ta0 = tactic_expr; "+"; ta1 = tactic_expr -> TacOr (ta0,ta1) - | IDENT "tryif" ; ta = tactic_expr ; - "then" ; tat = tactic_expr ; - "else" ; tae = tactic_expr -> TacIfThenCatch(ta,tat,tae) - | ta0 = tactic_expr; "||"; ta1 = binder_tactic -> TacOrelse (ta0,ta1) - | ta0 = tactic_expr; "||"; ta1 = tactic_expr -> TacOrelse (ta0,ta1) ] - | "1" RIGHTA - [ b = match_key; IDENT "goal"; "with"; mrl = match_context_list; "end" -> - TacMatchGoal (b,false,mrl) - | b = match_key; IDENT "reverse"; IDENT "goal"; "with"; - mrl = match_context_list; "end" -> - TacMatchGoal (b,true,mrl) - | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> - TacMatch (b,c,mrl) - | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacFirst l - | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacSolve l - | IDENT "idtac"; l = LIST0 message_token -> TacId l - | g=failkw; n = [ n = int_or_var -> n | -> fail_default_value ]; - l = LIST0 message_token -> TacFail (g,n,l) - | st = simple_tactic -> st - | a = tactic_arg -> TacArg(Loc.tag ~loc:!@loc a) - | r = reference; la = LIST0 tactic_arg_compat -> - TacArg(Loc.tag ~loc:!@loc @@ TacCall (Loc.tag ~loc:!@loc (r,la))) ] - | "0" - [ "("; a = tactic_expr; ")" -> a - | "["; ">"; (tf,tail) = tactic_then_gen; "]" -> - begin match tail with - | Some (t,tl) -> TacExtendTac(Array.of_list tf,t,tl) - | None -> TacDispatch tf - end - | a = tactic_atom -> TacArg (Loc.tag ~loc:!@loc a) ] ] - ; - failkw: - [ [ IDENT "fail" -> TacLocal | IDENT "gfail" -> TacGlobal ] ] - ; - (* binder_tactic: level 5 of tactic_expr *) - binder_tactic: - [ RIGHTA - [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr LEVEL "5" -> - TacFun (it,body) - | "let"; isrec = [IDENT "rec" -> true | -> false]; - llc = LIST1 let_clause SEP "with"; "in"; - body = tactic_expr LEVEL "5" -> TacLetIn (isrec,llc,body) - | IDENT "info"; tc = tactic_expr LEVEL "5" -> TacInfo tc ] ] - ; - (* Tactic arguments to the right of an application *) - tactic_arg_compat: - [ [ a = tactic_arg -> a - | c = Constr.constr -> (match c with { CAst.v = CRef (r,None) } -> Reference r | c -> ConstrMayEval (ConstrTerm c)) - (* Unambiguous entries: tolerated w/o "ltac:" modifier *) - | "()" -> TacGeneric (genarg_of_unit ()) ] ] - ; - (* Can be used as argument and at toplevel in tactic expressions. *) - tactic_arg: - [ [ c = constr_eval -> ConstrMayEval c - | IDENT "fresh"; l = LIST0 fresh_id -> TacFreshId l - | IDENT "type_term"; c=uconstr -> TacPretype c - | IDENT "numgoals" -> TacNumgoals ] ] - ; - (* If a qualid is given, use its short name. TODO: have the shortest - non ambiguous name where dots are replaced by "_"? Probably too - verbose most of the time. *) - fresh_id: - [ [ s = STRING -> Locus.ArgArg s (*| id = ident -> Locus.ArgVar (!@loc,id)*) - | qid = qualid -> Locus.ArgVar (CAst.make ~loc:!@loc @@ Libnames.qualid_basename qid) ] ] - ; - constr_eval: - [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> - ConstrEval (rtc,c) - | IDENT "context"; id = identref; "["; c = Constr.lconstr; "]" -> - ConstrContext (id,c) - | IDENT "type"; IDENT "of"; c = Constr.constr -> - ConstrTypeOf c ] ] - ; - constr_may_eval: (* For extensions *) - [ [ c = constr_eval -> c - | c = Constr.constr -> ConstrTerm c ] ] - ; - tactic_atom: - [ [ n = integer -> TacGeneric (genarg_of_int n) - | r = reference -> TacCall (Loc.tag ~loc:!@loc (r,[])) - | "()" -> TacGeneric (genarg_of_unit ()) ] ] - ; - match_key: - [ [ "match" -> Once - | "lazymatch" -> Select - | "multimatch" -> General ] ] - ; - input_fun: - [ [ "_" -> Name.Anonymous - | l = ident -> Name.Name l ] ] - ; - let_clause: - [ [ idr = identref; ":="; te = tactic_expr -> - (CAst.map (fun id -> Name id) idr, arg_of_expr te) - | na = ["_" -> CAst.make ~loc:!@loc Anonymous]; ":="; te = tactic_expr -> - (na, arg_of_expr te) - | idr = identref; args = LIST1 input_fun; ":="; te = tactic_expr -> - (CAst.map (fun id -> Name id) idr, arg_of_expr (TacFun(args,te))) ] ] - ; - match_pattern: - [ [ IDENT "context"; oid = OPT Constr.ident; - "["; pc = Constr.lconstr_pattern; "]" -> - Subterm (oid, pc) - | pc = Constr.lconstr_pattern -> Term pc ] ] - ; - match_hyps: - [ [ na = name; ":"; mp = match_pattern -> Hyp (na, mp) - | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> Def (na, mpv, mpt) - | na = name; ":="; mpv = match_pattern -> - let t, ty = - match mpv with - | Term t -> (match t with - | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) - | _ -> mpv, None) - | _ -> mpv, None - in Def (na, t, Option.default (Term (CAst.make @@ CHole (None, IntroAnonymous, None))) ty) - ] ] - ; - match_context_rule: - [ [ largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; - "=>"; te = tactic_expr -> Pat (largs, mp, te) - | "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; - "]"; "=>"; te = tactic_expr -> Pat (largs, mp, te) - | "_"; "=>"; te = tactic_expr -> All te ] ] - ; - match_context_list: - [ [ mrl = LIST1 match_context_rule SEP "|" -> mrl - | "|"; mrl = LIST1 match_context_rule SEP "|" -> mrl ] ] - ; - match_rule: - [ [ mp = match_pattern; "=>"; te = tactic_expr -> Pat ([],mp,te) - | "_"; "=>"; te = tactic_expr -> All te ] ] - ; - match_list: - [ [ mrl = LIST1 match_rule SEP "|" -> mrl - | "|"; mrl = LIST1 match_rule SEP "|" -> mrl ] ] - ; - message_token: - [ [ id = identref -> MsgIdent id - | s = STRING -> MsgString s - | n = integer -> MsgInt n ] ] - ; - - ltac_def_kind: - [ [ ":=" -> false - | "::=" -> true ] ] - ; - - (* Definitions for tactics *) - tacdef_body: - [ [ name = Constr.global; it=LIST1 input_fun; - redef = ltac_def_kind; body = tactic_expr -> - if redef then Tacexpr.TacticRedefinition (name, TacFun (it, body)) - else - let id = reference_to_id name in - Tacexpr.TacticDefinition (id, TacFun (it, body)) - | name = Constr.global; redef = ltac_def_kind; - body = tactic_expr -> - if redef then Tacexpr.TacticRedefinition (name, body) - else - let id = reference_to_id name in - Tacexpr.TacticDefinition (id, body) - ] ] - ; - tactic: - [ [ tac = tactic_expr -> tac ] ] - ; - - range_selector: - [ [ n = natural ; "-" ; m = natural -> (n, m) - | n = natural -> (n, n) ] ] - ; - (* We unfold a range selectors list once so that we can make a special case - * for a unique SelectNth selector. *) - range_selector_or_nth: - [ [ n = natural ; "-" ; m = natural; - l = OPT [","; l = LIST1 range_selector SEP "," -> l] -> - Goal_select.SelectList ((n, m) :: Option.default [] l) - | n = natural; - l = OPT [","; l = LIST1 range_selector SEP "," -> l] -> - let open Goal_select in - Option.cata (fun l -> SelectList ((n, n) :: l)) (SelectNth n) l ] ] - ; - selector_body: - [ [ l = range_selector_or_nth -> l - | test_bracket_ident; "["; id = ident; "]" -> Goal_select.SelectId id ] ] - ; - selector: - [ [ IDENT "only"; sel = selector_body; ":" -> sel ] ] - ; - toplevel_selector: - [ [ sel = selector_body; ":" -> sel - | "!"; ":" -> Goal_select.SelectAlreadyFocused - | IDENT "all"; ":" -> Goal_select.SelectAll ] ] - ; - tactic_mode: - [ [ g = OPT toplevel_selector; tac = G_vernac.query_command -> tac g - | g = OPT toplevel_selector; "{" -> Vernacexpr.VernacSubproof g ] ] - ; - command: - [ [ IDENT "Proof"; "with"; ta = Pltac.tactic; - l = OPT [ "using"; l = G_vernac.section_subset_expr -> l ] -> - Vernacexpr.VernacProof (Some (in_tac ta), l) - | IDENT "Proof"; "using"; l = G_vernac.section_subset_expr; - ta = OPT [ "with"; ta = Pltac.tactic -> in_tac ta ] -> - Vernacexpr.VernacProof (ta,Some l) ] ] - ; - hint: - [ [ IDENT "Extern"; n = natural; c = OPT Constr.constr_pattern ; "=>"; - tac = Pltac.tactic -> - Hints.HintsExtern (n,c, in_tac tac) ] ] - ; - operconstr: LEVEL "0" - [ [ IDENT "ltac"; ":"; "("; tac = Pltac.tactic_expr; ")" -> - let arg = Genarg.in_gen (Genarg.rawwit Tacarg.wit_tactic) tac in - CAst.make ~loc:!@loc @@ CHole (None, IntroAnonymous, Some arg) ] ] - ; - END - -open Stdarg -open Tacarg -open Vernacexpr -open Vernac_classifier -open Goptions -open Libnames - -let print_info_trace = ref None - -let _ = declare_int_option { - optdepr = false; - optname = "print info trace"; - optkey = ["Info" ; "Level"]; - optread = (fun () -> !print_info_trace); - optwrite = fun n -> print_info_trace := n; -} - -let vernac_solve n info tcom b = - let open Goal_select in - let status = Proof_global.with_current_proof (fun etac p -> - let with_end_tac = if b then Some etac else None in - let global = match n with SelectAll | SelectList _ -> true | _ -> false in - let info = Option.append info !print_info_trace in - let (p,status) = - Pfedit.solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p - in - (* in case a strict subtree was completed, - go back to the top of the prooftree *) - let p = Proof.maximal_unfocus Vernacentries.command_focus p in - p,status) in - if not status then Feedback.feedback Feedback.AddedAxiom - -let pr_ltac_selector s = Pptactic.pr_goal_selector ~toplevel:true s - -VERNAC ARGUMENT EXTEND ltac_selector PRINTED BY pr_ltac_selector -| [ toplevel_selector(s) ] -> [ s ] -END - -let pr_ltac_info n = str "Info" ++ spc () ++ int n - -VERNAC ARGUMENT EXTEND ltac_info PRINTED BY pr_ltac_info -| [ "Info" natural(n) ] -> [ n ] -END - -let pr_ltac_use_default b = - if b then (* Bug: a space is inserted before "..." *) str ".." else mt () - -VERNAC ARGUMENT EXTEND ltac_use_default PRINTED BY pr_ltac_use_default -| [ "." ] -> [ false ] -| [ "..." ] -> [ true ] -END - -let is_anonymous_abstract = function - | TacAbstract (_,None) -> true - | TacSolve [TacAbstract (_,None)] -> true - | _ -> false -let rm_abstract = function - | TacAbstract (t,_) -> t - | TacSolve [TacAbstract (t,_)] -> TacSolve [t] - | x -> x -let is_explicit_terminator = function TacSolve _ -> true | _ -> false - -VERNAC tactic_mode EXTEND VernacSolve -| [ - ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => - [ classify_as_proofstep ] -> [ - let g = Option.default (Goal_select.get_default_goal_selector ()) g in - vernac_solve g n t def - ] -| [ - "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => - [ - let anon_abstracting_tac = is_anonymous_abstract t in - let solving_tac = is_explicit_terminator t in - let parallel = `Yes (solving_tac,anon_abstracting_tac) in - let pbr = if solving_tac then Some "par" else None in - VtProofStep{ parallel = parallel; proof_block_detection = pbr }, - VtLater - ] -> [ - let t = rm_abstract t in - vernac_solve Goal_select.SelectAll n t def - ] -END - -let pr_ltac_tactic_level n = str "(at level " ++ int n ++ str ")" - -VERNAC ARGUMENT EXTEND ltac_tactic_level PRINTED BY pr_ltac_tactic_level -| [ "(" "at" "level" natural(n) ")" ] -> [ n ] -END - -VERNAC ARGUMENT EXTEND ltac_production_sep -| [ "," string(sep) ] -> [ sep ] -END - -let pr_ltac_production_item = function -| Tacentries.TacTerm s -> quote (str s) -| Tacentries.TacNonTerm (_, ((arg, None), None)) -> str arg -| Tacentries.TacNonTerm (_, ((arg, Some _), None)) -> assert false -| Tacentries.TacNonTerm (_, ((arg, sep), Some id)) -> - let sep = match sep with - | None -> mt () - | Some sep -> str "," ++ spc () ++ quote (str sep) - in - str arg ++ str "(" ++ Id.print id ++ sep ++ str ")" - -VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item -| [ string(s) ] -> [ Tacentries.TacTerm s ] -| [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> - [ Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, sep), Some p)) ] -| [ ident(nt) ] -> - [ Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, None), None)) ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND VernacTacticNotation -| [ "Tactic" "Notation" ltac_tactic_level_opt(n) ne_ltac_production_item_list(r) ":=" tactic(e) ] => - [ VtSideff [], VtNow ] -> - [ fun ~atts ~st -> let open Vernacinterp in - let n = Option.default 0 n in - let deprecation = atts.deprecated in - Tacentries.add_tactic_notation (Locality.make_module_locality atts.locality) n ?deprecation r e; - st - ] -END - -VERNAC COMMAND EXTEND VernacPrintLtac CLASSIFIED AS QUERY -| [ "Print" "Ltac" reference(r) ] -> - [ Feedback.msg_notice (Tacintern.print_ltac r) ] -END - -VERNAC COMMAND EXTEND VernacLocateLtac CLASSIFIED AS QUERY -| [ "Locate" "Ltac" reference(r) ] -> - [ Tacentries.print_located_tactic r ] -END - -let pr_ltac_ref = Libnames.pr_qualid - -let pr_tacdef_body tacdef_body = - let id, redef, body = - match tacdef_body with - | TacticDefinition ({CAst.v=id}, body) -> Id.print id, false, body - | TacticRedefinition (id, body) -> pr_ltac_ref id, true, body - in - let idl, body = - match body with - | Tacexpr.TacFun (idl,b) -> idl,b - | _ -> [], body in - id ++ - prlist (function Name.Anonymous -> str " _" - | Name.Name id -> spc () ++ Id.print id) idl - ++ (if redef then str" ::=" else str" :=") ++ brk(1,1) - ++ Pptactic.pr_raw_tactic body - -VERNAC ARGUMENT EXTEND ltac_tacdef_body -PRINTED BY pr_tacdef_body -| [ tacdef_body(t) ] -> [ t ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND VernacDeclareTacticDefinition -| [ "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 - ] -> [ fun ~atts ~st -> let open Vernacinterp in - let deprecation = atts.deprecated in - Tacentries.register_ltac (Locality.make_module_locality atts.locality) ?deprecation l; - st - ] -END - -VERNAC COMMAND EXTEND VernacPrintLtacs CLASSIFIED AS QUERY -| [ "Print" "Ltac" "Signatures" ] -> [ Tacentries.print_ltacs () ] -END diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg new file mode 100644 index 0000000000..d62f985350 --- /dev/null +++ b/plugins/ltac/g_ltac.mlg @@ -0,0 +1,560 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* a + | e -> Tacexp (e:raw_tactic_expr) + +let genarg_of_unit () = in_gen (rawwit Stdarg.wit_unit) () +let genarg_of_int n = in_gen (rawwit Stdarg.wit_int) n +let genarg_of_ipattern pat = in_gen (rawwit Tacarg.wit_intro_pattern) pat +let genarg_of_uconstr c = in_gen (rawwit Stdarg.wit_uconstr) c +let in_tac tac = in_gen (rawwit Tacarg.wit_ltac) tac + +let reference_to_id qid = + if Libnames.qualid_is_ident qid then + CAst.make ?loc:qid.CAst.loc @@ Libnames.qualid_basename qid + else + CErrors.user_err ?loc:qid.CAst.loc + (str "This expression should be a simple identifier.") + +let tactic_mode = Entry.create "vernac:tactic_command" + +let new_entry name = + let e = Entry.create name in + e + +let toplevel_selector = new_entry "vernac:toplevel_selector" +let tacdef_body = new_entry "tactic:tacdef_body" + +(* Registers the Classic Proof Mode (which uses [tactic_mode] as a parser for + proof editing and changes nothing else). Then sets it as the default proof mode. *) +let _ = + let mode = { + Proof_global.name = "Classic"; + set = (fun () -> Pvernac.set_command_entry tactic_mode); + reset = (fun () -> Pvernac.(set_command_entry noedit_mode)); + } in + Proof_global.register_proof_mode mode + +(* Hack to parse "[ id" without dropping [ *) +let test_bracket_ident = + Gram.Entry.of_parser "test_bracket_ident" + (fun strm -> + match stream_nth 0 strm with + | KEYWORD "[" -> + (match stream_nth 1 strm with + | IDENT _ -> () + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + +(* Tactics grammar rules *) + +let hint = G_proofs.hint + +} + +GRAMMAR EXTEND Gram + GLOBAL: tactic tacdef_body tactic_expr binder_tactic tactic_arg command hint + tactic_mode constr_may_eval constr_eval toplevel_selector + operconstr; + + tactic_then_last: + [ [ "|"; lta = LIST0 (OPT tactic_expr) SEP "|" -> + { Array.map (function None -> TacId [] | Some t -> t) (Array.of_list lta) } + | -> { [||] } + ] ] + ; + tactic_then_gen: + [ [ ta = tactic_expr; "|"; tg = tactic_then_gen -> { let (first,last) = tg in (ta::first, last) } + | ta = tactic_expr; ".."; l = tactic_then_last -> { ([], Some (ta, l)) } + | ".."; l = tactic_then_last -> { ([], Some (TacId [], l)) } + | ta = tactic_expr -> { ([ta], None) } + | "|"; tg = tactic_then_gen -> { let (first,last) = tg in (TacId [] :: first, last) } + | -> { ([TacId []], None) } + ] ] + ; + tactic_then_locality: (* [true] for the local variant [TacThens] and [false] + for [TacExtend] *) + [ [ "[" ; l = OPT">" -> { if Option.is_empty l then true else false } ] ] + ; + tactic_expr: + [ "5" RIGHTA + [ te = binder_tactic -> { te } ] + | "4" LEFTA + [ ta0 = tactic_expr; ";"; ta1 = binder_tactic -> { TacThen (ta0, ta1) } + | ta0 = tactic_expr; ";"; ta1 = tactic_expr -> { TacThen (ta0,ta1) } + | ta0 = tactic_expr; ";"; l = tactic_then_locality; tg = tactic_then_gen; "]" -> { + let (first,tail) = tg in + match l , tail with + | false , Some (t,last) -> TacThen (ta0,TacExtendTac (Array.of_list first, t, last)) + | true , Some (t,last) -> TacThens3parts (ta0, Array.of_list first, t, last) + | false , None -> TacThen (ta0,TacDispatch first) + | true , None -> TacThens (ta0,first) } ] + | "3" RIGHTA + [ IDENT "try"; ta = tactic_expr -> { TacTry ta } + | IDENT "do"; n = int_or_var; ta = tactic_expr -> { TacDo (n,ta) } + | IDENT "timeout"; n = int_or_var; ta = tactic_expr -> { TacTimeout (n,ta) } + | IDENT "time"; s = OPT string; ta = tactic_expr -> { TacTime (s,ta) } + | IDENT "repeat"; ta = tactic_expr -> { TacRepeat ta } + | IDENT "progress"; ta = tactic_expr -> { TacProgress ta } + | IDENT "once"; ta = tactic_expr -> { TacOnce ta } + | IDENT "exactly_once"; ta = tactic_expr -> { TacExactlyOnce ta } + | IDENT "infoH"; ta = tactic_expr -> { TacShowHyps ta } +(*To do: put Abstract in Refiner*) + | IDENT "abstract"; tc = NEXT -> { TacAbstract (tc,None) } + | IDENT "abstract"; tc = NEXT; "using"; s = ident -> + { TacAbstract (tc,Some s) } + | sel = selector; ta = tactic_expr -> { TacSelect (sel, ta) } ] +(*End of To do*) + | "2" RIGHTA + [ ta0 = tactic_expr; "+"; ta1 = binder_tactic -> { TacOr (ta0,ta1) } + | ta0 = tactic_expr; "+"; ta1 = tactic_expr -> { TacOr (ta0,ta1) } + | IDENT "tryif" ; ta = tactic_expr ; + "then" ; tat = tactic_expr ; + "else" ; tae = tactic_expr -> { TacIfThenCatch(ta,tat,tae) } + | ta0 = tactic_expr; "||"; ta1 = binder_tactic -> { TacOrelse (ta0,ta1) } + | ta0 = tactic_expr; "||"; ta1 = tactic_expr -> { TacOrelse (ta0,ta1) } ] + | "1" RIGHTA + [ b = match_key; IDENT "goal"; "with"; mrl = match_context_list; "end" -> + { TacMatchGoal (b,false,mrl) } + | b = match_key; IDENT "reverse"; IDENT "goal"; "with"; + mrl = match_context_list; "end" -> + { TacMatchGoal (b,true,mrl) } + | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> + { TacMatch (b,c,mrl) } + | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> + { TacFirst l } + | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> + { TacSolve l } + | IDENT "idtac"; l = LIST0 message_token -> { TacId l } + | g=failkw; n = [ n = int_or_var -> { n } | -> { fail_default_value } ]; + l = LIST0 message_token -> { TacFail (g,n,l) } + | st = simple_tactic -> { st } + | a = tactic_arg -> { TacArg(Loc.tag ~loc a) } + | r = reference; la = LIST0 tactic_arg_compat -> + { TacArg(Loc.tag ~loc @@ TacCall (Loc.tag ~loc (r,la))) } ] + | "0" + [ "("; a = tactic_expr; ")" -> { a } + | "["; ">"; tg = tactic_then_gen; "]" -> { + let (tf,tail) = tg in + begin match tail with + | Some (t,tl) -> TacExtendTac(Array.of_list tf,t,tl) + | None -> TacDispatch tf + end } + | a = tactic_atom -> { TacArg (Loc.tag ~loc a) } ] ] + ; + failkw: + [ [ IDENT "fail" -> { TacLocal } | IDENT "gfail" -> { TacGlobal } ] ] + ; + (* binder_tactic: level 5 of tactic_expr *) + binder_tactic: + [ RIGHTA + [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr LEVEL "5" -> + { TacFun (it,body) } + | "let"; isrec = [IDENT "rec" -> { true } | -> { false } ]; + llc = LIST1 let_clause SEP "with"; "in"; + body = tactic_expr LEVEL "5" -> { TacLetIn (isrec,llc,body) } + | IDENT "info"; tc = tactic_expr LEVEL "5" -> { TacInfo tc } ] ] + ; + (* Tactic arguments to the right of an application *) + tactic_arg_compat: + [ [ a = tactic_arg -> { a } + | c = Constr.constr -> { (match c with { CAst.v = CRef (r,None) } -> Reference r | c -> ConstrMayEval (ConstrTerm c)) } + (* Unambiguous entries: tolerated w/o "ltac:" modifier *) + | "()" -> { TacGeneric (genarg_of_unit ()) } ] ] + ; + (* Can be used as argument and at toplevel in tactic expressions. *) + tactic_arg: + [ [ c = constr_eval -> { ConstrMayEval c } + | IDENT "fresh"; l = LIST0 fresh_id -> { TacFreshId l } + | IDENT "type_term"; c=uconstr -> { TacPretype c } + | IDENT "numgoals" -> { TacNumgoals } ] ] + ; + (* If a qualid is given, use its short name. TODO: have the shortest + non ambiguous name where dots are replaced by "_"? Probably too + verbose most of the time. *) + fresh_id: + [ [ s = STRING -> { Locus.ArgArg s (*| id = ident -> Locus.ArgVar (!@loc,id)*) } + | qid = qualid -> { Locus.ArgVar (CAst.make ~loc @@ Libnames.qualid_basename qid) } ] ] + ; + constr_eval: + [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> + { ConstrEval (rtc,c) } + | IDENT "context"; id = identref; "["; c = Constr.lconstr; "]" -> + { ConstrContext (id,c) } + | IDENT "type"; IDENT "of"; c = Constr.constr -> + { ConstrTypeOf c } ] ] + ; + constr_may_eval: (* For extensions *) + [ [ c = constr_eval -> { c } + | c = Constr.constr -> { ConstrTerm c } ] ] + ; + tactic_atom: + [ [ n = integer -> { TacGeneric (genarg_of_int n) } + | r = reference -> { TacCall (Loc.tag ~loc (r,[])) } + | "()" -> { TacGeneric (genarg_of_unit ()) } ] ] + ; + match_key: + [ [ "match" -> { Once } + | "lazymatch" -> { Select } + | "multimatch" -> { General } ] ] + ; + input_fun: + [ [ "_" -> { Name.Anonymous } + | l = ident -> { Name.Name l } ] ] + ; + let_clause: + [ [ idr = identref; ":="; te = tactic_expr -> + { (CAst.map (fun id -> Name id) idr, arg_of_expr te) } + | na = ["_" -> { CAst.make ~loc Anonymous } ]; ":="; te = tactic_expr -> + { (na, arg_of_expr te) } + | idr = identref; args = LIST1 input_fun; ":="; te = tactic_expr -> + { (CAst.map (fun id -> Name id) idr, arg_of_expr (TacFun(args,te))) } ] ] + ; + match_pattern: + [ [ IDENT "context"; oid = OPT Constr.ident; + "["; pc = Constr.lconstr_pattern; "]" -> + { Subterm (oid, pc) } + | pc = Constr.lconstr_pattern -> { Term pc } ] ] + ; + match_hyps: + [ [ na = name; ":"; mp = match_pattern -> { Hyp (na, mp) } + | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> { Def (na, mpv, mpt) } + | na = name; ":="; mpv = match_pattern -> + { let t, ty = + match mpv with + | Term t -> (match t with + | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) + | _ -> mpv, None) + | _ -> mpv, None + in Def (na, t, Option.default (Term (CAst.make @@ CHole (None, IntroAnonymous, None))) ty) } + ] ] + ; + match_context_rule: + [ [ largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; + "=>"; te = tactic_expr -> { Pat (largs, mp, te) } + | "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; + "]"; "=>"; te = tactic_expr -> { Pat (largs, mp, te) } + | "_"; "=>"; te = tactic_expr -> { All te } ] ] + ; + match_context_list: + [ [ mrl = LIST1 match_context_rule SEP "|" -> { mrl } + | "|"; mrl = LIST1 match_context_rule SEP "|" -> { mrl } ] ] + ; + match_rule: + [ [ mp = match_pattern; "=>"; te = tactic_expr -> { Pat ([],mp,te) } + | "_"; "=>"; te = tactic_expr -> { All te } ] ] + ; + match_list: + [ [ mrl = LIST1 match_rule SEP "|" -> { mrl } + | "|"; mrl = LIST1 match_rule SEP "|" -> { mrl } ] ] + ; + message_token: + [ [ id = identref -> { MsgIdent id } + | s = STRING -> { MsgString s } + | n = integer -> { MsgInt n } ] ] + ; + + ltac_def_kind: + [ [ ":=" -> { false } + | "::=" -> { true } ] ] + ; + + (* Definitions for tactics *) + tacdef_body: + [ [ name = Constr.global; it=LIST1 input_fun; + redef = ltac_def_kind; body = tactic_expr -> + { if redef then Tacexpr.TacticRedefinition (name, TacFun (it, body)) + else + let id = reference_to_id name in + Tacexpr.TacticDefinition (id, TacFun (it, body)) } + | name = Constr.global; redef = ltac_def_kind; + body = tactic_expr -> + { if redef then Tacexpr.TacticRedefinition (name, body) + else + let id = reference_to_id name in + Tacexpr.TacticDefinition (id, body) } + ] ] + ; + tactic: + [ [ tac = tactic_expr -> { tac } ] ] + ; + + range_selector: + [ [ n = natural ; "-" ; m = natural -> { (n, m) } + | n = natural -> { (n, n) } ] ] + ; + (* We unfold a range selectors list once so that we can make a special case + * for a unique SelectNth selector. *) + range_selector_or_nth: + [ [ n = natural ; "-" ; m = natural; + l = OPT [","; l = LIST1 range_selector SEP "," -> { l } ] -> + { Goal_select.SelectList ((n, m) :: Option.default [] l) } + | n = natural; + l = OPT [","; l = LIST1 range_selector SEP "," -> { l } ] -> + { let open Goal_select in + Option.cata (fun l -> SelectList ((n, n) :: l)) (SelectNth n) l } ] ] + ; + selector_body: + [ [ l = range_selector_or_nth -> { l } + | test_bracket_ident; "["; id = ident; "]" -> { Goal_select.SelectId id } ] ] + ; + selector: + [ [ IDENT "only"; sel = selector_body; ":" -> { sel } ] ] + ; + toplevel_selector: + [ [ sel = selector_body; ":" -> { sel } + | "!"; ":" -> { Goal_select.SelectAlreadyFocused } + | IDENT "all"; ":" -> { Goal_select.SelectAll } ] ] + ; + tactic_mode: + [ [ g = OPT toplevel_selector; tac = G_vernac.query_command -> { tac g } + | g = OPT toplevel_selector; "{" -> { Vernacexpr.VernacSubproof g } ] ] + ; + command: + [ [ IDENT "Proof"; "with"; ta = Pltac.tactic; + l = OPT [ "using"; l = G_vernac.section_subset_expr -> { l } ] -> + { Vernacexpr.VernacProof (Some (in_tac ta), l) } + | IDENT "Proof"; "using"; l = G_vernac.section_subset_expr; + ta = OPT [ "with"; ta = Pltac.tactic -> { in_tac ta } ] -> + { Vernacexpr.VernacProof (ta,Some l) } ] ] + ; + hint: + [ [ IDENT "Extern"; n = natural; c = OPT Constr.constr_pattern ; "=>"; + tac = Pltac.tactic -> + { Hints.HintsExtern (n,c, in_tac tac) } ] ] + ; + operconstr: LEVEL "0" + [ [ IDENT "ltac"; ":"; "("; tac = Pltac.tactic_expr; ")" -> + { let arg = Genarg.in_gen (Genarg.rawwit Tacarg.wit_tactic) tac in + CAst.make ~loc @@ CHole (None, IntroAnonymous, Some arg) } ] ] + ; + END + +{ + +open Stdarg +open Tacarg +open Vernacexpr +open Vernac_classifier +open Goptions +open Libnames + +let print_info_trace = ref None + +let _ = declare_int_option { + optdepr = false; + optname = "print info trace"; + optkey = ["Info" ; "Level"]; + optread = (fun () -> !print_info_trace); + optwrite = fun n -> print_info_trace := n; +} + +let vernac_solve n info tcom b = + let open Goal_select in + let status = Proof_global.with_current_proof (fun etac p -> + let with_end_tac = if b then Some etac else None in + let global = match n with SelectAll | SelectList _ -> true | _ -> false in + let info = Option.append info !print_info_trace in + let (p,status) = + Pfedit.solve n info (Tacinterp.hide_interp global tcom None) ?with_end_tac p + in + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + let p = Proof.maximal_unfocus Vernacentries.command_focus p in + p,status) in + if not status then Feedback.feedback Feedback.AddedAxiom + +let pr_ltac_selector s = Pptactic.pr_goal_selector ~toplevel:true s + +} + +VERNAC ARGUMENT EXTEND ltac_selector PRINTED BY { pr_ltac_selector } +| [ toplevel_selector(s) ] -> { s } +END + +{ + +let pr_ltac_info n = str "Info" ++ spc () ++ int n + +} + +VERNAC ARGUMENT EXTEND ltac_info PRINTED BY { pr_ltac_info } +| [ "Info" natural(n) ] -> { n } +END + +{ + +let pr_ltac_use_default b = + if b then (* Bug: a space is inserted before "..." *) str ".." else mt () + +} + +VERNAC ARGUMENT EXTEND ltac_use_default PRINTED BY { pr_ltac_use_default } +| [ "." ] -> { false } +| [ "..." ] -> { true } +END + +{ + +let is_anonymous_abstract = function + | TacAbstract (_,None) -> true + | TacSolve [TacAbstract (_,None)] -> true + | _ -> false +let rm_abstract = function + | TacAbstract (t,_) -> t + | TacSolve [TacAbstract (t,_)] -> TacSolve [t] + | x -> x +let is_explicit_terminator = function TacSolve _ -> true | _ -> false + +} + +VERNAC { tactic_mode } EXTEND VernacSolve +| [ ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => + { classify_as_proofstep } -> { + let g = Option.default (Goal_select.get_default_goal_selector ()) g in + vernac_solve g n t def + } +| [ "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] => + { + let anon_abstracting_tac = is_anonymous_abstract t in + let solving_tac = is_explicit_terminator t in + let parallel = `Yes (solving_tac,anon_abstracting_tac) in + let pbr = if solving_tac then Some "par" else None in + VtProofStep{ parallel = parallel; proof_block_detection = pbr }, + VtLater + } -> { + let t = rm_abstract t in + vernac_solve Goal_select.SelectAll n t def + } +END + +{ + +let pr_ltac_tactic_level n = str "(at level " ++ int n ++ str ")" + +} + +VERNAC ARGUMENT EXTEND ltac_tactic_level PRINTED BY { pr_ltac_tactic_level } +| [ "(" "at" "level" natural(n) ")" ] -> { n } +END + +VERNAC ARGUMENT EXTEND ltac_production_sep +| [ "," string(sep) ] -> { sep } +END + +{ + +let pr_ltac_production_item = function +| Tacentries.TacTerm s -> quote (str s) +| Tacentries.TacNonTerm (_, ((arg, None), None)) -> str arg +| Tacentries.TacNonTerm (_, ((arg, Some _), None)) -> assert false +| Tacentries.TacNonTerm (_, ((arg, sep), Some id)) -> + let sep = match sep with + | None -> mt () + | Some sep -> str "," ++ spc () ++ quote (str sep) + in + str arg ++ str "(" ++ Id.print id ++ sep ++ str ")" + +} + +VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY { pr_ltac_production_item } +| [ string(s) ] -> { Tacentries.TacTerm s } +| [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> + { Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, sep), Some p)) } +| [ ident(nt) ] -> + { Tacentries.TacNonTerm (Loc.tag ~loc ((Id.to_string nt, None), None)) } +END + +VERNAC COMMAND EXTEND VernacTacticNotation +| [ "Tactic" "Notation" ltac_tactic_level_opt(n) ne_ltac_production_item_list(r) ":=" tactic(e) ] => + { VtSideff [], VtNow } -> + { let open Vernacinterp in + let n = Option.default 0 n in + let deprecation = atts.deprecated in + Tacentries.add_tactic_notation (Locality.make_module_locality atts.locality) n ?deprecation r e; + } +END + +VERNAC COMMAND EXTEND VernacPrintLtac CLASSIFIED AS QUERY +| [ "Print" "Ltac" reference(r) ] -> + { Feedback.msg_notice (Tacintern.print_ltac r) } +END + +VERNAC COMMAND EXTEND VernacLocateLtac CLASSIFIED AS QUERY +| [ "Locate" "Ltac" reference(r) ] -> + { Tacentries.print_located_tactic r } +END + +{ + +let pr_ltac_ref = Libnames.pr_qualid + +let pr_tacdef_body tacdef_body = + let id, redef, body = + match tacdef_body with + | TacticDefinition ({CAst.v=id}, body) -> Id.print id, false, body + | TacticRedefinition (id, body) -> pr_ltac_ref id, true, body + in + let idl, body = + match body with + | Tacexpr.TacFun (idl,b) -> idl,b + | _ -> [], body in + id ++ + prlist (function Name.Anonymous -> str " _" + | Name.Name id -> spc () ++ Id.print id) idl + ++ (if redef then str" ::=" else str" :=") ++ brk(1,1) + ++ Pptactic.pr_raw_tactic body + +} + +VERNAC ARGUMENT EXTEND ltac_tacdef_body +PRINTED BY { pr_tacdef_body } +| [ tacdef_body(t) ] -> { t } +END + +VERNAC COMMAND EXTEND VernacDeclareTacticDefinition +| [ "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 open Vernacinterp in + let deprecation = atts.deprecated in + Tacentries.register_ltac (Locality.make_module_locality atts.locality) ?deprecation l; + } +END + +VERNAC COMMAND EXTEND VernacPrintLtacs CLASSIFIED AS QUERY +| [ "Print" "Ltac" "Signatures" ] -> { Tacentries.print_ltacs () } +END diff --git a/plugins/ltac/g_obligations.ml4 b/plugins/ltac/g_obligations.ml4 deleted file mode 100644 index 1f56244c75..0000000000 --- a/plugins/ltac/g_obligations.ml4 +++ /dev/null @@ -1,161 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* - snd (get_default_tactic ()) - end in - Obligations.default_tactic := tac - -let with_tac f tac = - let env = Genintern.empty_glob_sign (Global.env ()) in - let tac = match tac with - | None -> None - | Some tac -> - let tac = Genarg.in_gen (Genarg.rawwit wit_ltac) tac in - let _, tac = Genintern.generic_intern env tac in - Some tac - in - f tac - -(* We define new entries for programs, with the use of this module - * Subtac. These entries are named Subtac. - *) - -module Gram = Pcoq.Gram -module Tactic = Pltac - -open Pcoq - -let sigref loc = mkRefC (Libnames.qualid_of_string ~loc "Coq.Init.Specif.sig") - -type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type - -let wit_withtac : Tacexpr.raw_tactic_expr option Genarg.uniform_genarg_type = - Genarg.create_arg "withtac" - -let withtac = Pcoq.create_generic_entry Pcoq.utactic "withtac" (Genarg.rawwit wit_withtac) - -GEXTEND Gram - GLOBAL: withtac; - - withtac: - [ [ "with"; t = Tactic.tactic -> Some t - | -> None ] ] - ; - - Constr.closed_binder: - [[ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" -> - let typ = mkAppC (sigref !@loc, [mkLambdaC ([id], default_binder_kind, t, c)]) in - [CLocalAssum ([id], default_binder_kind, typ)] - ] ]; - - END - -open Obligations - -let obligation obl tac = with_tac (fun t -> Obligations.obligation obl t) tac -let next_obligation obl tac = with_tac (fun t -> Obligations.next_obligation obl t) tac - -let classify_obbl _ = Vernacexpr.(VtStartProof ("Classic",Doesn'tGuaranteeOpacity,[]), VtLater) - -VERNAC COMMAND EXTEND Obligations CLASSIFIED BY classify_obbl -| [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] -> - [ obligation (num, Some name, Some t) tac ] -| [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] -> - [ obligation (num, Some name, None) tac ] -| [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] -> - [ obligation (num, None, Some t) tac ] -| [ "Obligation" integer(num) withtac(tac) ] -> - [ obligation (num, None, None) tac ] -| [ "Next" "Obligation" "of" ident(name) withtac(tac) ] -> - [ next_obligation (Some name) tac ] -| [ "Next" "Obligation" withtac(tac) ] -> [ next_obligation None tac ] -END - -VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF -| [ "Solve" "Obligation" integer(num) "of" ident(name) "with" tactic(t) ] -> - [ try_solve_obligation num (Some name) (Some (Tacinterp.interp t)) ] -| [ "Solve" "Obligation" integer(num) "with" tactic(t) ] -> - [ try_solve_obligation num None (Some (Tacinterp.interp t)) ] -END - -VERNAC COMMAND EXTEND Solve_Obligations CLASSIFIED AS SIDEFF -| [ "Solve" "Obligations" "of" ident(name) "with" tactic(t) ] -> - [ try_solve_obligations (Some name) (Some (Tacinterp.interp t)) ] -| [ "Solve" "Obligations" "with" tactic(t) ] -> - [ try_solve_obligations None (Some (Tacinterp.interp t)) ] -| [ "Solve" "Obligations" ] -> - [ try_solve_obligations None None ] -END - -VERNAC COMMAND EXTEND Solve_All_Obligations CLASSIFIED AS SIDEFF -| [ "Solve" "All" "Obligations" "with" tactic(t) ] -> - [ solve_all_obligations (Some (Tacinterp.interp t)) ] -| [ "Solve" "All" "Obligations" ] -> - [ solve_all_obligations None ] -END - -VERNAC COMMAND EXTEND Admit_Obligations CLASSIFIED AS SIDEFF -| [ "Admit" "Obligations" "of" ident(name) ] -> [ admit_obligations (Some name) ] -| [ "Admit" "Obligations" ] -> [ admit_obligations None ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND Set_Solver CLASSIFIED AS SIDEFF -| [ "Obligation" "Tactic" ":=" tactic(t) ] -> [ - fun ~atts ~st -> begin - let open Vernacinterp in - set_default_tactic - (Locality.make_section_locality atts.locality) - (Tacintern.glob_tactic t); - st - end] -END - -open Pp - -VERNAC COMMAND EXTEND Show_Solver CLASSIFIED AS QUERY -| [ "Show" "Obligation" "Tactic" ] -> [ - Feedback.msg_info (str"Program obligation tactic is " ++ print_default_tactic ()) ] -END - -VERNAC COMMAND EXTEND Show_Obligations CLASSIFIED AS QUERY -| [ "Obligations" "of" ident(name) ] -> [ show_obligations (Some name) ] -| [ "Obligations" ] -> [ show_obligations None ] -END - -VERNAC COMMAND EXTEND Show_Preterm CLASSIFIED AS QUERY -| [ "Preterm" "of" ident(name) ] -> [ Feedback.msg_info (show_term (Some name)) ] -| [ "Preterm" ] -> [ Feedback.msg_info (show_term None) ] -END - -open Pp - -(* Declare a printer for the content of Program tactics *) -let () = - let printer _ _ _ = function - | None -> mt () - | Some tac -> str "with" ++ spc () ++ Pptactic.pr_raw_tactic tac - in - Pptactic.declare_extra_vernac_genarg_pprule wit_withtac printer diff --git a/plugins/ltac/g_obligations.mlg b/plugins/ltac/g_obligations.mlg new file mode 100644 index 0000000000..26f2b08d3a --- /dev/null +++ b/plugins/ltac/g_obligations.mlg @@ -0,0 +1,173 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + snd (get_default_tactic ()) + end in + Obligations.default_tactic := tac + +let with_tac f tac = + let env = Genintern.empty_glob_sign (Global.env ()) in + let tac = match tac with + | None -> None + | Some tac -> + let tac = Genarg.in_gen (Genarg.rawwit wit_ltac) tac in + let _, tac = Genintern.generic_intern env tac in + Some tac + in + f tac + +(* We define new entries for programs, with the use of this module + * Subtac. These entries are named Subtac. + *) + +module Gram = Pcoq.Gram +module Tactic = Pltac + +open Pcoq + +let sigref loc = mkRefC (Libnames.qualid_of_string ~loc "Coq.Init.Specif.sig") + +type 'a withtac_argtype = (Tacexpr.raw_tactic_expr option, 'a) Genarg.abstract_argument_type + +let wit_withtac : Tacexpr.raw_tactic_expr option Genarg.uniform_genarg_type = + Genarg.create_arg "withtac" + +let withtac = Pcoq.create_generic_entry Pcoq.utactic "withtac" (Genarg.rawwit wit_withtac) + +} + +GRAMMAR EXTEND Gram + GLOBAL: withtac; + + withtac: + [ [ "with"; t = Tactic.tactic -> { Some t } + | -> { None } ] ] + ; + + Constr.closed_binder: + [[ "("; id=Prim.name; ":"; t=Constr.lconstr; "|"; c=Constr.lconstr; ")" -> { + let typ = mkAppC (sigref loc, [mkLambdaC ([id], default_binder_kind, t, c)]) in + [CLocalAssum ([id], default_binder_kind, typ)] } + ] ]; + + END + +{ + +open Obligations + +let obligation obl tac = with_tac (fun t -> Obligations.obligation obl t) tac +let next_obligation obl tac = with_tac (fun t -> Obligations.next_obligation obl t) tac + +let classify_obbl _ = Vernacexpr.(VtStartProof ("Classic",Doesn'tGuaranteeOpacity,[]), VtLater) + +} + +VERNAC COMMAND EXTEND Obligations CLASSIFIED BY { classify_obbl } +| [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] -> + { obligation (num, Some name, Some t) tac } +| [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] -> + { obligation (num, Some name, None) tac } +| [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] -> + { obligation (num, None, Some t) tac } +| [ "Obligation" integer(num) withtac(tac) ] -> + { obligation (num, None, None) tac } +| [ "Next" "Obligation" "of" ident(name) withtac(tac) ] -> + { next_obligation (Some name) tac } +| [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac } +END + +VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF +| [ "Solve" "Obligation" integer(num) "of" ident(name) "with" tactic(t) ] -> + { try_solve_obligation num (Some name) (Some (Tacinterp.interp t)) } +| [ "Solve" "Obligation" integer(num) "with" tactic(t) ] -> + { try_solve_obligation num None (Some (Tacinterp.interp t)) } +END + +VERNAC COMMAND EXTEND Solve_Obligations CLASSIFIED AS SIDEFF +| [ "Solve" "Obligations" "of" ident(name) "with" tactic(t) ] -> + { try_solve_obligations (Some name) (Some (Tacinterp.interp t)) } +| [ "Solve" "Obligations" "with" tactic(t) ] -> + { try_solve_obligations None (Some (Tacinterp.interp t)) } +| [ "Solve" "Obligations" ] -> + { try_solve_obligations None None } +END + +VERNAC COMMAND EXTEND Solve_All_Obligations CLASSIFIED AS SIDEFF +| [ "Solve" "All" "Obligations" "with" tactic(t) ] -> + { solve_all_obligations (Some (Tacinterp.interp t)) } +| [ "Solve" "All" "Obligations" ] -> + { solve_all_obligations None } +END + +VERNAC COMMAND EXTEND Admit_Obligations CLASSIFIED AS SIDEFF +| [ "Admit" "Obligations" "of" ident(name) ] -> { admit_obligations (Some name) } +| [ "Admit" "Obligations" ] -> { admit_obligations None } +END + +VERNAC COMMAND EXTEND Set_Solver CLASSIFIED AS SIDEFF +| [ "Obligation" "Tactic" ":=" tactic(t) ] -> { + let open Vernacinterp in + set_default_tactic + (Locality.make_section_locality atts.locality) + (Tacintern.glob_tactic t); + } +END + +{ + +open Pp + +} + +VERNAC COMMAND EXTEND Show_Solver CLASSIFIED AS QUERY +| [ "Show" "Obligation" "Tactic" ] -> { + Feedback.msg_info (str"Program obligation tactic is " ++ print_default_tactic ()) } +END + +VERNAC COMMAND EXTEND Show_Obligations CLASSIFIED AS QUERY +| [ "Obligations" "of" ident(name) ] -> { show_obligations (Some name) } +| [ "Obligations" ] -> { show_obligations None } +END + +VERNAC COMMAND EXTEND Show_Preterm CLASSIFIED AS QUERY +| [ "Preterm" "of" ident(name) ] -> { Feedback.msg_info (show_term (Some name)) } +| [ "Preterm" ] -> { Feedback.msg_info (show_term None) } +END + +{ + +(* Declare a printer for the content of Program tactics *) +let () = + let printer _ _ _ = function + | None -> mt () + | Some tac -> str "with" ++ spc () ++ Pptactic.pr_raw_tactic tac + in + Pptactic.declare_extra_vernac_genarg_pprule wit_withtac printer + +} diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 deleted file mode 100644 index f1634f1561..0000000000 --- a/plugins/ltac/g_rewrite.ml4 +++ /dev/null @@ -1,297 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* [ bl ] -END - -type raw_strategy = (constr_expr, Tacexpr.raw_red_expr) strategy_ast -type glob_strategy = (Tacexpr.glob_constr_and_expr, Tacexpr.raw_red_expr) strategy_ast - -let interp_strategy ist gl s = - let sigma = project gl in - sigma, strategy_of_ast s -let glob_strategy ist s = map_strategy (Tacintern.intern_constr ist) (fun c -> c) s -let subst_strategy s str = str - -let pr_strategy _ _ _ (s : strategy) = Pp.str "" -let pr_raw_strategy prc prlc _ (s : raw_strategy) = - let prr = Pptactic.pr_red_expr (prc, prlc, Pputils.pr_or_by_notation Libnames.pr_qualid, prc) in - Rewrite.pr_strategy prc prr s -let pr_glob_strategy prc prlc _ (s : glob_strategy) = - let prr = Pptactic.pr_red_expr - (Ppconstr.pr_constr_expr, - Ppconstr.pr_lconstr_expr, - Pputils.pr_or_by_notation Libnames.pr_qualid, - Ppconstr.pr_constr_expr) - in - Rewrite.pr_strategy prc prr s - -ARGUMENT EXTEND rewstrategy - PRINTED BY pr_strategy - - INTERPRETED BY interp_strategy - GLOBALIZED BY glob_strategy - SUBSTITUTED BY subst_strategy - - RAW_PRINTED BY pr_raw_strategy - GLOB_PRINTED BY pr_glob_strategy - - [ glob(c) ] -> [ StratConstr (c, true) ] - | [ "<-" constr(c) ] -> [ StratConstr (c, false) ] - | [ "subterms" rewstrategy(h) ] -> [ StratUnary (Subterms, h) ] - | [ "subterm" rewstrategy(h) ] -> [ StratUnary (Subterm, h) ] - | [ "innermost" rewstrategy(h) ] -> [ StratUnary(Innermost, h) ] - | [ "outermost" rewstrategy(h) ] -> [ StratUnary(Outermost, h) ] - | [ "bottomup" rewstrategy(h) ] -> [ StratUnary(Bottomup, h) ] - | [ "topdown" rewstrategy(h) ] -> [ StratUnary(Topdown, h) ] - | [ "id" ] -> [ StratId ] - | [ "fail" ] -> [ StratFail ] - | [ "refl" ] -> [ StratRefl ] - | [ "progress" rewstrategy(h) ] -> [ StratUnary (Progress, h) ] - | [ "try" rewstrategy(h) ] -> [ StratUnary (Try, h) ] - | [ "any" rewstrategy(h) ] -> [ StratUnary (Any, h) ] - | [ "repeat" rewstrategy(h) ] -> [ StratUnary (Repeat, h) ] - | [ rewstrategy(h) ";" rewstrategy(h') ] -> [ StratBinary (Compose, h, h') ] - | [ "(" rewstrategy(h) ")" ] -> [ h ] - | [ "choice" rewstrategy(h) rewstrategy(h') ] -> [ StratBinary (Choice, h, h') ] - | [ "old_hints" preident(h) ] -> [ StratHints (true, h) ] - | [ "hints" preident(h) ] -> [ StratHints (false, h) ] - | [ "terms" constr_list(h) ] -> [ StratTerms h ] - | [ "eval" red_expr(r) ] -> [ StratEval r ] - | [ "fold" constr(c) ] -> [ StratFold c ] -END - -(* By default the strategy for "rewrite_db" is top-down *) - -let db_strat db = StratUnary (Topdown, StratHints (false, db)) -let cl_rewrite_clause_db db = cl_rewrite_clause_strat (strategy_of_ast (db_strat db)) - -TACTIC EXTEND rewrite_strat -| [ "rewrite_strat" rewstrategy(s) "in" hyp(id) ] -> [ cl_rewrite_clause_strat s (Some id) ] -| [ "rewrite_strat" rewstrategy(s) ] -> [ cl_rewrite_clause_strat s None ] -| [ "rewrite_db" preident(db) "in" hyp(id) ] -> [ cl_rewrite_clause_db db (Some id) ] -| [ "rewrite_db" preident(db) ] -> [ cl_rewrite_clause_db db None ] -END - -let clsubstitute o c = - Proofview.Goal.enter begin fun gl -> - let is_tac id = match DAst.get (fst (fst (snd c))) with GVar id' when Id.equal id' id -> true | _ -> false in - let hyps = Tacmach.New.pf_ids_of_hyps gl in - Tacticals.New.tclMAP - (fun cl -> - match cl with - | Some id when is_tac id -> Tacticals.New.tclIDTAC - | _ -> cl_rewrite_clause c o AllOccurrences cl) - (None :: List.map (fun id -> Some id) hyps) - end - -TACTIC EXTEND substitute -| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> [ clsubstitute o c ] -END - - -(* Compatibility with old Setoids *) - -TACTIC EXTEND setoid_rewrite - [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) ] - -> [ cl_rewrite_clause c o AllOccurrences None ] - | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) ] -> - [ cl_rewrite_clause c o AllOccurrences (Some id) ] - | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) ] -> - [ cl_rewrite_clause c o (occurrences_of occ) None ] - | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) "in" hyp(id)] -> - [ cl_rewrite_clause c o (occurrences_of occ) (Some id) ] - | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) "at" occurrences(occ)] -> - [ cl_rewrite_clause c o (occurrences_of occ) (Some id) ] -END - -VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF - | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> - [ declare_relation a aeq n (Some lemma1) (Some lemma2) None ] - - | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "as" ident(n) ] -> - [ declare_relation a aeq n (Some lemma1) None None ] - | [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] -> - [ declare_relation a aeq n None None None ] -END - -VERNAC COMMAND EXTEND AddRelation2 CLASSIFIED AS SIDEFF - [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) - "as" ident(n) ] -> - [ declare_relation a aeq n None (Some lemma2) None ] - | [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> - [ declare_relation a aeq n None (Some lemma2) (Some lemma3) ] -END - -VERNAC COMMAND EXTEND AddRelation3 CLASSIFIED AS SIDEFF - [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> - [ declare_relation a aeq n (Some lemma1) None (Some lemma3) ] - | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) - "as" ident(n) ] -> - [ declare_relation a aeq n (Some lemma1) (Some lemma2) (Some lemma3) ] - | [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) - "as" ident(n) ] -> - [ declare_relation a aeq n None None (Some lemma3) ] -END - -type binders_argtype = local_binder_expr list - -let wit_binders = - (Genarg.create_arg "binders" : binders_argtype Genarg.uniform_genarg_type) - -let binders = Pcoq.create_generic_entry Pcoq.utactic "binders" (Genarg.rawwit wit_binders) - -let () = - let raw_printer _ _ _ l = Pp.pr_non_empty_arg Ppconstr.pr_binders l in - Pptactic.declare_extra_vernac_genarg_pprule wit_binders raw_printer - -open Pcoq - -GEXTEND Gram - GLOBAL: binders; - binders: - [ [ b = Pcoq.Constr.binders -> b ] ]; -END - -VERNAC COMMAND EXTEND AddParametricRelation CLASSIFIED AS SIDEFF - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) - "reflexivity" "proved" "by" constr(lemma1) - "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n (Some lemma1) (Some lemma2) None ] - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) - "reflexivity" "proved" "by" constr(lemma1) - "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n (Some lemma1) None None ] - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n None None None ] -END - -VERNAC COMMAND EXTEND AddParametricRelation2 CLASSIFIED AS SIDEFF - [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) - "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n None (Some lemma2) None ] - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n None (Some lemma2) (Some lemma3) ] -END - -VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF - [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n (Some lemma1) None (Some lemma3) ] - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) - "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) - "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n (Some lemma1) (Some lemma2) (Some lemma3) ] - | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) - "as" ident(n) ] -> - [ declare_relation ~binders:b a aeq n None None (Some lemma3) ] -END - -VERNAC COMMAND FUNCTIONAL EXTEND AddSetoid1 CLASSIFIED AS SIDEFF - [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> - [ fun ~atts ~st -> let open Vernacinterp in - add_setoid (not (Locality.make_section_locality atts.locality)) [] a aeq t n; - st - ] - | [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> - [ fun ~atts ~st -> let open Vernacinterp in - add_setoid (not (Locality.make_section_locality atts.locality)) binders a aeq t n; - st - ] - | [ "Add" "Morphism" constr(m) ":" ident(n) ] - (* This command may or may not open a goal *) - => [ Vernacexpr.VtUnknown, Vernacexpr.VtNow ] - -> [ fun ~atts ~st -> let open Vernacinterp in - add_morphism_infer (not (Locality.make_section_locality atts.locality)) m n; - st - ] - | [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ] - => [ Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) ] - -> [ fun ~atts ~st -> let open Vernacinterp in - add_morphism (not (Locality.make_section_locality atts.locality)) [] m s n; - st - ] - | [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m) - "with" "signature" lconstr(s) "as" ident(n) ] - => [ Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) ] - -> [ fun ~atts ~st -> let open Vernacinterp in - add_morphism (not (Locality.make_section_locality atts.locality)) binders m s n; - st - ] -END - -TACTIC EXTEND setoid_symmetry - [ "setoid_symmetry" ] -> [ setoid_symmetry ] - | [ "setoid_symmetry" "in" hyp(n) ] -> [ setoid_symmetry_in n ] -END - -TACTIC EXTEND setoid_reflexivity -[ "setoid_reflexivity" ] -> [ setoid_reflexivity ] -END - -TACTIC EXTEND setoid_transitivity - [ "setoid_transitivity" constr(t) ] -> [ setoid_transitivity (Some t) ] -| [ "setoid_etransitivity" ] -> [ setoid_transitivity None ] -END - -VERNAC COMMAND EXTEND PrintRewriteHintDb CLASSIFIED AS QUERY - [ "Print" "Rewrite" "HintDb" preident(s) ] -> - [ let sigma, env = Pfedit.get_current_context () in - Feedback.msg_notice (Autorewrite.print_rewrite_hintdb env sigma s) ] -END diff --git a/plugins/ltac/g_rewrite.mlg b/plugins/ltac/g_rewrite.mlg new file mode 100644 index 0000000000..3e47724c4c --- /dev/null +++ b/plugins/ltac/g_rewrite.mlg @@ -0,0 +1,318 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* { bl } +END + +{ + +type raw_strategy = (constr_expr, Tacexpr.raw_red_expr) strategy_ast +type glob_strategy = (Tacexpr.glob_constr_and_expr, Tacexpr.raw_red_expr) strategy_ast + +let interp_strategy ist gl s = + let sigma = project gl in + sigma, strategy_of_ast s +let glob_strategy ist s = map_strategy (Tacintern.intern_constr ist) (fun c -> c) s +let subst_strategy s str = str + +let pr_strategy _ _ _ (s : strategy) = Pp.str "" +let pr_raw_strategy prc prlc _ (s : raw_strategy) = + let prr = Pptactic.pr_red_expr (prc, prlc, Pputils.pr_or_by_notation Libnames.pr_qualid, prc) in + Rewrite.pr_strategy prc prr s +let pr_glob_strategy prc prlc _ (s : glob_strategy) = + let prr = Pptactic.pr_red_expr + (Ppconstr.pr_constr_expr, + Ppconstr.pr_lconstr_expr, + Pputils.pr_or_by_notation Libnames.pr_qualid, + Ppconstr.pr_constr_expr) + in + Rewrite.pr_strategy prc prr s + +} + +ARGUMENT EXTEND rewstrategy + PRINTED BY { pr_strategy } + + INTERPRETED BY { interp_strategy } + GLOBALIZED BY { glob_strategy } + SUBSTITUTED BY { subst_strategy } + + RAW_PRINTED BY { pr_raw_strategy } + GLOB_PRINTED BY { pr_glob_strategy } + + | [ glob(c) ] -> { StratConstr (c, true) } + | [ "<-" constr(c) ] -> { StratConstr (c, false) } + | [ "subterms" rewstrategy(h) ] -> { StratUnary (Subterms, h) } + | [ "subterm" rewstrategy(h) ] -> { StratUnary (Subterm, h) } + | [ "innermost" rewstrategy(h) ] -> { StratUnary(Innermost, h) } + | [ "outermost" rewstrategy(h) ] -> { StratUnary(Outermost, h) } + | [ "bottomup" rewstrategy(h) ] -> { StratUnary(Bottomup, h) } + | [ "topdown" rewstrategy(h) ] -> { StratUnary(Topdown, h) } + | [ "id" ] -> { StratId } + | [ "fail" ] -> { StratFail } + | [ "refl" ] -> { StratRefl } + | [ "progress" rewstrategy(h) ] -> { StratUnary (Progress, h) } + | [ "try" rewstrategy(h) ] -> { StratUnary (Try, h) } + | [ "any" rewstrategy(h) ] -> { StratUnary (Any, h) } + | [ "repeat" rewstrategy(h) ] -> { StratUnary (Repeat, h) } + | [ rewstrategy(h) ";" rewstrategy(h') ] -> { StratBinary (Compose, h, h') } + | [ "(" rewstrategy(h) ")" ] -> { h } + | [ "choice" rewstrategy(h) rewstrategy(h') ] -> { StratBinary (Choice, h, h') } + | [ "old_hints" preident(h) ] -> { StratHints (true, h) } + | [ "hints" preident(h) ] -> { StratHints (false, h) } + | [ "terms" constr_list(h) ] -> { StratTerms h } + | [ "eval" red_expr(r) ] -> { StratEval r } + | [ "fold" constr(c) ] -> { StratFold c } +END + +(* By default the strategy for "rewrite_db" is top-down *) + +{ + +let db_strat db = StratUnary (Topdown, StratHints (false, db)) +let cl_rewrite_clause_db db = cl_rewrite_clause_strat (strategy_of_ast (db_strat db)) + +} + +TACTIC EXTEND rewrite_strat +| [ "rewrite_strat" rewstrategy(s) "in" hyp(id) ] -> { cl_rewrite_clause_strat s (Some id) } +| [ "rewrite_strat" rewstrategy(s) ] -> { cl_rewrite_clause_strat s None } +| [ "rewrite_db" preident(db) "in" hyp(id) ] -> { cl_rewrite_clause_db db (Some id) } +| [ "rewrite_db" preident(db) ] -> { cl_rewrite_clause_db db None } +END + +{ + +let clsubstitute o c = + Proofview.Goal.enter begin fun gl -> + let is_tac id = match DAst.get (fst (fst (snd c))) with GVar id' when Id.equal id' id -> true | _ -> false in + let hyps = Tacmach.New.pf_ids_of_hyps gl in + Tacticals.New.tclMAP + (fun cl -> + match cl with + | Some id when is_tac id -> Tacticals.New.tclIDTAC + | _ -> cl_rewrite_clause c o AllOccurrences cl) + (None :: List.map (fun id -> Some id) hyps) + end + +} + +TACTIC EXTEND substitute +| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> { clsubstitute o c } +END + + +(* Compatibility with old Setoids *) + +TACTIC EXTEND setoid_rewrite + | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) ] + -> { cl_rewrite_clause c o AllOccurrences None } + | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) ] -> + { cl_rewrite_clause c o AllOccurrences (Some id) } + | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) ] -> + { cl_rewrite_clause c o (occurrences_of occ) None } + | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "at" occurrences(occ) "in" hyp(id)] -> + { cl_rewrite_clause c o (occurrences_of occ) (Some id) } + | [ "setoid_rewrite" orient(o) glob_constr_with_bindings(c) "in" hyp(id) "at" occurrences(occ)] -> + { cl_rewrite_clause c o (occurrences_of occ) (Some id) } +END + +VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF + | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> + { declare_relation a aeq n (Some lemma1) (Some lemma2) None } + + | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "as" ident(n) ] -> + { declare_relation a aeq n (Some lemma1) None None } + | [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] -> + { declare_relation a aeq n None None None } +END + +VERNAC COMMAND EXTEND AddRelation2 CLASSIFIED AS SIDEFF + | [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) + "as" ident(n) ] -> + { declare_relation a aeq n None (Some lemma2) None } + | [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> + { declare_relation a aeq n None (Some lemma2) (Some lemma3) } +END + +VERNAC COMMAND EXTEND AddRelation3 CLASSIFIED AS SIDEFF + | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> + { declare_relation a aeq n (Some lemma1) None (Some lemma3) } + | [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) + "as" ident(n) ] -> + { declare_relation a aeq n (Some lemma1) (Some lemma2) (Some lemma3) } + | [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) + "as" ident(n) ] -> + { declare_relation a aeq n None None (Some lemma3) } +END + +{ + +type binders_argtype = local_binder_expr list + +let wit_binders = + (Genarg.create_arg "binders" : binders_argtype Genarg.uniform_genarg_type) + +let binders = Pcoq.create_generic_entry Pcoq.utactic "binders" (Genarg.rawwit wit_binders) + +let () = + let raw_printer _ _ _ l = Pp.pr_non_empty_arg Ppconstr.pr_binders l in + Pptactic.declare_extra_vernac_genarg_pprule wit_binders raw_printer + +open Pcoq + +} + +GRAMMAR EXTEND Gram + GLOBAL: binders; + binders: + [ [ b = Pcoq.Constr.binders -> { b } ] ]; +END + +VERNAC COMMAND EXTEND AddParametricRelation CLASSIFIED AS SIDEFF + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) + "reflexivity" "proved" "by" constr(lemma1) + "symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n (Some lemma1) (Some lemma2) None } + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) + "reflexivity" "proved" "by" constr(lemma1) + "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n (Some lemma1) None None } + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n None None None } +END + +VERNAC COMMAND EXTEND AddParametricRelation2 CLASSIFIED AS SIDEFF + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) + "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n None (Some lemma2) None } + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n None (Some lemma2) (Some lemma3) } +END + +VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n (Some lemma1) None (Some lemma3) } + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1) + "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) + "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n (Some lemma1) (Some lemma2) (Some lemma3) } + | [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3) + "as" ident(n) ] -> + { declare_relation ~binders:b a aeq n None None (Some lemma3) } +END + +VERNAC COMMAND EXTEND AddSetoid1 CLASSIFIED AS SIDEFF + | [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> + { let open Vernacinterp in + add_setoid (not (Locality.make_section_locality atts.locality)) [] a aeq t n; + } + | [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] -> + { let open Vernacinterp in + add_setoid (not (Locality.make_section_locality atts.locality)) binders a aeq t n; + } + | [ "Add" "Morphism" constr(m) ":" ident(n) ] + (* This command may or may not open a goal *) + => { Vernacexpr.VtUnknown, Vernacexpr.VtNow } + -> { let open Vernacinterp in + add_morphism_infer (not (Locality.make_section_locality atts.locality)) m n; + } + | [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ] + => { Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) } + -> { let open Vernacinterp in + add_morphism (not (Locality.make_section_locality atts.locality)) [] m s n; + } + | [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m) + "with" "signature" lconstr(s) "as" ident(n) ] + => { Vernacexpr.(VtStartProof("Classic",GuaranteesOpacity,[n]), VtLater) } + -> { let open Vernacinterp in + add_morphism (not (Locality.make_section_locality atts.locality)) binders m s n; + } +END + +TACTIC EXTEND setoid_symmetry + | [ "setoid_symmetry" ] -> { setoid_symmetry } + | [ "setoid_symmetry" "in" hyp(n) ] -> { setoid_symmetry_in n } +END + +TACTIC EXTEND setoid_reflexivity +| [ "setoid_reflexivity" ] -> { setoid_reflexivity } +END + +TACTIC EXTEND setoid_transitivity +| [ "setoid_transitivity" constr(t) ] -> { setoid_transitivity (Some t) } +| [ "setoid_etransitivity" ] -> { setoid_transitivity None } +END + +VERNAC COMMAND EXTEND PrintRewriteHintDb CLASSIFIED AS QUERY +| [ "Print" "Rewrite" "HintDb" preident(s) ] -> + { let sigma, env = Pfedit.get_current_context () in + Feedback.msg_notice (Autorewrite.print_rewrite_hintdb env sigma s) } +END diff --git a/plugins/ltac/profile_ltac_tactics.ml4 b/plugins/ltac/profile_ltac_tactics.ml4 deleted file mode 100644 index 983e1578be..0000000000 --- a/plugins/ltac/profile_ltac_tactics.ml4 +++ /dev/null @@ -1,74 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* set_profiling b)) - -let tclRESET_PROFILE = - Proofview.tclLIFT (Proofview.NonLogical.make reset_profile) - -let tclSHOW_PROFILE ~cutoff = - Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results ~cutoff)) - -let tclSHOW_PROFILE_TACTIC s = - Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results_tactic s)) - -let tclRESTART_TIMER s = - Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> restart_timer s)) - -let tclFINISH_TIMING ?(prefix="Timer") (s : string option) = - Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> finish_timing ~prefix s)) - -TACTIC EXTEND start_ltac_profiling -| [ "start" "ltac" "profiling" ] -> [ tclSET_PROFILING true ] -END - -TACTIC EXTEND stop_ltac_profiling -| [ "stop" "ltac" "profiling" ] -> [ tclSET_PROFILING false ] -END - -TACTIC EXTEND reset_ltac_profile -| [ "reset" "ltac" "profile" ] -> [ tclRESET_PROFILE ] -END - -TACTIC EXTEND show_ltac_profile -| [ "show" "ltac" "profile" ] -> [ tclSHOW_PROFILE ~cutoff:!Flags.profile_ltac_cutoff ] -| [ "show" "ltac" "profile" "cutoff" int(n) ] -> [ tclSHOW_PROFILE ~cutoff:(float_of_int n) ] -| [ "show" "ltac" "profile" string(s) ] -> [ tclSHOW_PROFILE_TACTIC s ] -END - -TACTIC EXTEND restart_timer -| [ "restart_timer" string_opt(s) ] -> [ tclRESTART_TIMER s ] -END - -TACTIC EXTEND finish_timing -| [ "finish_timing" string_opt(s) ] -> [ tclFINISH_TIMING ~prefix:"Timer" s ] -| [ "finish_timing" "(" string(prefix) ")" string_opt(s) ] -> [ tclFINISH_TIMING ~prefix s ] -END - -VERNAC COMMAND EXTEND ResetLtacProfiling CLASSIFIED AS SIDEFF - [ "Reset" "Ltac" "Profile" ] -> [ reset_profile () ] -END - -VERNAC COMMAND EXTEND ShowLtacProfile CLASSIFIED AS QUERY -| [ "Show" "Ltac" "Profile" ] -> [ print_results ~cutoff:!Flags.profile_ltac_cutoff ] -| [ "Show" "Ltac" "Profile" "CutOff" int(n) ] -> [ print_results ~cutoff:(float_of_int n) ] -END - -VERNAC COMMAND EXTEND ShowLtacProfileTactic CLASSIFIED AS QUERY - [ "Show" "Ltac" "Profile" string(s) ] -> [ print_results_tactic s ] -END diff --git a/plugins/ltac/profile_ltac_tactics.mlg b/plugins/ltac/profile_ltac_tactics.mlg new file mode 100644 index 0000000000..2713819c7b --- /dev/null +++ b/plugins/ltac/profile_ltac_tactics.mlg @@ -0,0 +1,82 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* set_profiling b)) + +let tclRESET_PROFILE = + Proofview.tclLIFT (Proofview.NonLogical.make reset_profile) + +let tclSHOW_PROFILE ~cutoff = + Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results ~cutoff)) + +let tclSHOW_PROFILE_TACTIC s = + Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> print_results_tactic s)) + +let tclRESTART_TIMER s = + Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> restart_timer s)) + +let tclFINISH_TIMING ?(prefix="Timer") (s : string option) = + Proofview.tclLIFT (Proofview.NonLogical.make (fun () -> finish_timing ~prefix s)) + +} + +TACTIC EXTEND start_ltac_profiling +| [ "start" "ltac" "profiling" ] -> { tclSET_PROFILING true } +END + +TACTIC EXTEND stop_ltac_profiling +| [ "stop" "ltac" "profiling" ] -> { tclSET_PROFILING false } +END + +TACTIC EXTEND reset_ltac_profile +| [ "reset" "ltac" "profile" ] -> { tclRESET_PROFILE } +END + +TACTIC EXTEND show_ltac_profile +| [ "show" "ltac" "profile" ] -> { tclSHOW_PROFILE ~cutoff:!Flags.profile_ltac_cutoff } +| [ "show" "ltac" "profile" "cutoff" int(n) ] -> { tclSHOW_PROFILE ~cutoff:(float_of_int n) } +| [ "show" "ltac" "profile" string(s) ] -> { tclSHOW_PROFILE_TACTIC s } +END + +TACTIC EXTEND restart_timer +| [ "restart_timer" string_opt(s) ] -> { tclRESTART_TIMER s } +END + +TACTIC EXTEND finish_timing +| [ "finish_timing" string_opt(s) ] -> { tclFINISH_TIMING ~prefix:"Timer" s } +| [ "finish_timing" "(" string(prefix) ")" string_opt(s) ] -> { tclFINISH_TIMING ~prefix s } +END + +VERNAC COMMAND EXTEND ResetLtacProfiling CLASSIFIED AS SIDEFF +| [ "Reset" "Ltac" "Profile" ] -> { reset_profile () } +END + +VERNAC COMMAND EXTEND ShowLtacProfile CLASSIFIED AS QUERY +| [ "Show" "Ltac" "Profile" ] -> { print_results ~cutoff:!Flags.profile_ltac_cutoff } +| [ "Show" "Ltac" "Profile" "CutOff" int(n) ] -> { print_results ~cutoff:(float_of_int n) } +END + +VERNAC COMMAND EXTEND ShowLtacProfileTactic CLASSIFIED AS QUERY +| [ "Show" "Ltac" "Profile" string(s) ] -> { print_results_tactic s } +END diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 deleted file mode 100644 index 4ea0b30bd7..0000000000 --- a/plugins/setoid_ring/g_newring.ml4 +++ /dev/null @@ -1,127 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* - [ protect_tac_in map id ] -| [ "protect_fv" string(map) ] -> - [ protect_tac map ] -END - -open Pptactic -open Ppconstr - -let pr_ring_mod = function - | Ring_kind (Computational eq_test) -> str "decidable" ++ pr_arg pr_constr_expr eq_test - | Ring_kind Abstract -> str "abstract" - | Ring_kind (Morphism morph) -> str "morphism" ++ pr_arg pr_constr_expr morph - | Const_tac (CstTac cst_tac) -> str "constants" ++ spc () ++ str "[" ++ pr_raw_tactic cst_tac ++ str "]" - | Const_tac (Closed l) -> str "closed" ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" - | Pre_tac t -> str "preprocess" ++ spc () ++ str "[" ++ pr_raw_tactic t ++ str "]" - | Post_tac t -> str "postprocess" ++ spc () ++ str "[" ++ pr_raw_tactic t ++ str "]" - | Setoid(sth,ext) -> str "setoid" ++ pr_arg pr_constr_expr sth ++ pr_arg pr_constr_expr ext - | Pow_spec(Closed l,spec) -> str "power_tac" ++ pr_arg pr_constr_expr spec ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" - | Pow_spec(CstTac cst_tac,spec) -> str "power_tac" ++ pr_arg pr_constr_expr spec ++ spc () ++ str "[" ++ pr_raw_tactic cst_tac ++ str "]" - | Sign_spec t -> str "sign" ++ pr_arg pr_constr_expr t - | Div_spec t -> str "div" ++ pr_arg pr_constr_expr t - -VERNAC ARGUMENT EXTEND ring_mod - PRINTED BY pr_ring_mod - | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational eq_test) ] - | [ "abstract" ] -> [ Ring_kind Abstract ] - | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism morph) ] - | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ] - | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ] - | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ] - | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ] - | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ] - | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ] - | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] -> - [ Pow_spec (Closed l, pow_spec) ] - | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] -> - [ Pow_spec (CstTac cst_tac, pow_spec) ] - | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ] -END - -let pr_ring_mods l = surround (prlist_with_sep pr_comma pr_ring_mod l) - -VERNAC ARGUMENT EXTEND ring_mods - PRINTED BY pr_ring_mods - | [ "(" ne_ring_mod_list_sep(mods, ",") ")" ] -> [ mods ] -END - -VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF - | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] -> - [ let l = match l with None -> [] | Some l -> l in add_theory id t l] - | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [ - Feedback.msg_notice (strbrk "The following ring structures have been declared:"); - Spmap.iter (fun fn fi -> - let sigma, env = Pfedit.get_current_context () in - Feedback.msg_notice (hov 2 - (Ppconstr.pr_id (Libnames.basename fn)++spc()++ - str"with carrier "++ pr_constr_env env sigma fi.ring_carrier++spc()++ - str"and equivalence relation "++ pr_constr_env env sigma fi.ring_req)) - ) !from_name ] -END - -TACTIC EXTEND ring_lookup -| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> - [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t ] -END - -let pr_field_mod = function - | Ring_mod m -> pr_ring_mod m - | Inject inj -> str "completeness" ++ pr_arg pr_constr_expr inj - -VERNAC ARGUMENT EXTEND field_mod - PRINTED BY pr_field_mod - | [ ring_mod(m) ] -> [ Ring_mod m ] - | [ "completeness" constr(inj) ] -> [ Inject inj ] -END - -let pr_field_mods l = surround (prlist_with_sep pr_comma pr_field_mod l) - -VERNAC ARGUMENT EXTEND field_mods - PRINTED BY pr_field_mods - | [ "(" ne_field_mod_list_sep(mods, ",") ")" ] -> [ mods ] -END - -VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF -| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] -> - [ let l = match l with None -> [] | Some l -> l in add_field_theory id t l ] -| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [ - Feedback.msg_notice (strbrk "The following field structures have been declared:"); - Spmap.iter (fun fn fi -> - let sigma, env = Pfedit.get_current_context () in - Feedback.msg_notice (hov 2 - (Ppconstr.pr_id (Libnames.basename fn)++spc()++ - str"with carrier "++ pr_constr_env env sigma fi.field_carrier++spc()++ - str"and equivalence relation "++ pr_constr_env env sigma fi.field_req)) - ) !field_from_name ] -END - -TACTIC EXTEND field_lookup -| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> - [ let (t,l) = List.sep_last lt in field_lookup f lH l t ] -END diff --git a/plugins/setoid_ring/g_newring.mlg b/plugins/setoid_ring/g_newring.mlg new file mode 100644 index 0000000000..3ddea7eb30 --- /dev/null +++ b/plugins/setoid_ring/g_newring.mlg @@ -0,0 +1,147 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* + { protect_tac_in map id } +| [ "protect_fv" string(map) ] -> + { protect_tac map } +END + +{ + +open Pptactic +open Ppconstr + +let pr_ring_mod = function + | Ring_kind (Computational eq_test) -> str "decidable" ++ pr_arg pr_constr_expr eq_test + | Ring_kind Abstract -> str "abstract" + | Ring_kind (Morphism morph) -> str "morphism" ++ pr_arg pr_constr_expr morph + | Const_tac (CstTac cst_tac) -> str "constants" ++ spc () ++ str "[" ++ pr_raw_tactic cst_tac ++ str "]" + | Const_tac (Closed l) -> str "closed" ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" + | Pre_tac t -> str "preprocess" ++ spc () ++ str "[" ++ pr_raw_tactic t ++ str "]" + | Post_tac t -> str "postprocess" ++ spc () ++ str "[" ++ pr_raw_tactic t ++ str "]" + | Setoid(sth,ext) -> str "setoid" ++ pr_arg pr_constr_expr sth ++ pr_arg pr_constr_expr ext + | Pow_spec(Closed l,spec) -> str "power_tac" ++ pr_arg pr_constr_expr spec ++ spc () ++ str "[" ++ prlist_with_sep spc pr_qualid l ++ str "]" + | Pow_spec(CstTac cst_tac,spec) -> str "power_tac" ++ pr_arg pr_constr_expr spec ++ spc () ++ str "[" ++ pr_raw_tactic cst_tac ++ str "]" + | Sign_spec t -> str "sign" ++ pr_arg pr_constr_expr t + | Div_spec t -> str "div" ++ pr_arg pr_constr_expr t + +} + +VERNAC ARGUMENT EXTEND ring_mod + PRINTED BY { pr_ring_mod } + | [ "decidable" constr(eq_test) ] -> { Ring_kind(Computational eq_test) } + | [ "abstract" ] -> { Ring_kind Abstract } + | [ "morphism" constr(morph) ] -> { Ring_kind(Morphism morph) } + | [ "constants" "[" tactic(cst_tac) "]" ] -> { Const_tac(CstTac cst_tac) } + | [ "closed" "[" ne_global_list(l) "]" ] -> { Const_tac(Closed l) } + | [ "preprocess" "[" tactic(pre) "]" ] -> { Pre_tac pre } + | [ "postprocess" "[" tactic(post) "]" ] -> { Post_tac post } + | [ "setoid" constr(sth) constr(ext) ] -> { Setoid(sth,ext) } + | [ "sign" constr(sign_spec) ] -> { Sign_spec sign_spec } + | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] -> + { Pow_spec (Closed l, pow_spec) } + | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] -> + { Pow_spec (CstTac cst_tac, pow_spec) } + | [ "div" constr(div_spec) ] -> { Div_spec div_spec } +END + +{ + +let pr_ring_mods l = surround (prlist_with_sep pr_comma pr_ring_mod l) + +} + +VERNAC ARGUMENT EXTEND ring_mods + PRINTED BY { pr_ring_mods } + | [ "(" ne_ring_mod_list_sep(mods, ",") ")" ] -> { mods } +END + +VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF + | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] -> + { let l = match l with None -> [] | Some l -> l in add_theory id t l } + | [ "Print" "Rings" ] => {Vernac_classifier.classify_as_query} -> { + Feedback.msg_notice (strbrk "The following ring structures have been declared:"); + Spmap.iter (fun fn fi -> + let sigma, env = Pfedit.get_current_context () in + Feedback.msg_notice (hov 2 + (Ppconstr.pr_id (Libnames.basename fn)++spc()++ + str"with carrier "++ pr_constr_env env sigma fi.ring_carrier++spc()++ + str"and equivalence relation "++ pr_constr_env env sigma fi.ring_req)) + ) !from_name } +END + +TACTIC EXTEND ring_lookup +| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> + { let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t } +END + +{ + +let pr_field_mod = function + | Ring_mod m -> pr_ring_mod m + | Inject inj -> str "completeness" ++ pr_arg pr_constr_expr inj + +} + +VERNAC ARGUMENT EXTEND field_mod + PRINTED BY { pr_field_mod } + | [ ring_mod(m) ] -> { Ring_mod m } + | [ "completeness" constr(inj) ] -> { Inject inj } +END + +{ + +let pr_field_mods l = surround (prlist_with_sep pr_comma pr_field_mod l) + +} + +VERNAC ARGUMENT EXTEND field_mods + PRINTED BY { pr_field_mods } + | [ "(" ne_field_mod_list_sep(mods, ",") ")" ] -> { mods } +END + +VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF +| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] -> + { let l = match l with None -> [] | Some l -> l in add_field_theory id t l } +| [ "Print" "Fields" ] => {Vernac_classifier.classify_as_query} -> { + Feedback.msg_notice (strbrk "The following field structures have been declared:"); + Spmap.iter (fun fn fi -> + let sigma, env = Pfedit.get_current_context () in + Feedback.msg_notice (hov 2 + (Ppconstr.pr_id (Libnames.basename fn)++spc()++ + str"with carrier "++ pr_constr_env env sigma fi.field_carrier++spc()++ + str"and equivalence relation "++ pr_constr_env env sigma fi.field_req)) + ) !field_from_name } +END + +TACTIC EXTEND field_lookup +| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> + { let (t,l) = List.sep_last lt in field_lookup f lH l t } +END diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4 deleted file mode 100644 index 319f58931a..0000000000 --- a/plugins/ssr/ssrparser.ml4 +++ /dev/null @@ -1,2348 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* [ CErrors.anomaly (Pp.str "Grammar placeholder match") ] -END -GEXTEND Gram - GLOBAL: ssrtacarg; - ssrtacarg: [[ tac = tactic_expr LEVEL "5" -> tac ]]; -END - -(* Lexically closed tactic for tacticals. *) -let pr_ssrtclarg _ _ prt tac = prt tacltop tac -ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg - PRINTED BY pr_ssrtclarg -| [ ssrtacarg(tac) ] -> [ tac ] -END - -open Genarg - -(** Adding a new uninterpreted generic argument type *) -let add_genarg tag pr = - let wit = Genarg.make0 tag in - let tag = Geninterp.Val.create tag in - let glob ist x = (ist, x) in - let subst _ x = x in - let interp ist x = Ftactic.return (Geninterp.Val.Dyn (tag, x)) in - let gen_pr _ _ _ = pr in - let () = Genintern.register_intern0 wit glob in - let () = Genintern.register_subst0 wit subst in - let () = Geninterp.register_interp0 wit interp in - let () = Geninterp.register_val0 wit (Some (Geninterp.Val.Base tag)) in - Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; - wit - -(** Primitive parsing to avoid syntax conflicts with basic tactics. *) - -let accept_before_syms syms strm = - match Util.stream_nth 1 strm with - | Tok.KEYWORD sym when List.mem sym syms -> () - | _ -> raise Stream.Failure - -let accept_before_syms_or_any_id syms strm = - match Util.stream_nth 1 strm with - | Tok.KEYWORD sym when List.mem sym syms -> () - | Tok.IDENT _ -> () - | _ -> raise Stream.Failure - -let accept_before_syms_or_ids syms ids strm = - match Util.stream_nth 1 strm with - | Tok.KEYWORD sym when List.mem sym syms -> () - | Tok.IDENT id when List.mem id ids -> () - | _ -> raise Stream.Failure - -open Ssrast -let pr_id = Ppconstr.pr_id -let pr_name = function Name id -> pr_id id | Anonymous -> str "_" -let pr_spc () = str " " -let pr_list = prlist_with_sep - -(**************************** ssrhyp **************************************) - -let pr_ssrhyp _ _ _ = pr_hyp - -let wit_ssrhyprep = add_genarg "ssrhyprep" pr_hyp - -let intern_hyp ist (SsrHyp (loc, id) as hyp) = - let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_var) CAst.(make ?loc id)) in - if not_section_id id then hyp else - hyp_err ?loc "Can't clear section hypothesis " id - -open Pcoq.Prim - -ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY pr_ssrhyp - INTERPRETED BY interp_hyp - GLOBALIZED BY intern_hyp - | [ ident(id) ] -> [ SsrHyp (Loc.tag ~loc id) ] -END - - -let pr_hoi = hoik pr_hyp -let pr_ssrhoi _ _ _ = pr_hoi - -let wit_ssrhoirep = add_genarg "ssrhoirep" pr_hoi - -let intern_ssrhoi ist = function - | Hyp h -> Hyp (intern_hyp ist h) - | Id (SsrHyp (_, id)) as hyp -> - let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_ident) id) in - hyp - -let interp_ssrhoi ist gl = function - | Hyp h -> let s, h' = interp_hyp ist gl h in s, Hyp h' - | Id (SsrHyp (loc, id)) -> - let s, id' = interp_wit wit_ident ist gl id in - s, Id (SsrHyp (loc, id')) - -ARGUMENT EXTEND ssrhoi_hyp TYPED AS ssrhoirep PRINTED BY pr_ssrhoi - INTERPRETED BY interp_ssrhoi - GLOBALIZED BY intern_ssrhoi - | [ ident(id) ] -> [ Hyp (SsrHyp(Loc.tag ~loc id)) ] -END -ARGUMENT EXTEND ssrhoi_id TYPED AS ssrhoirep PRINTED BY pr_ssrhoi - INTERPRETED BY interp_ssrhoi - GLOBALIZED BY intern_ssrhoi - | [ ident(id) ] -> [ Id (SsrHyp(Loc.tag ~loc id)) ] -END - - -let pr_ssrhyps _ _ _ = pr_hyps - -ARGUMENT EXTEND ssrhyps TYPED AS ssrhyp list PRINTED BY pr_ssrhyps - INTERPRETED BY interp_hyps - | [ ssrhyp_list(hyps) ] -> [ check_hyps_uniq [] hyps; hyps ] -END - -(** Rewriting direction *) - - -let pr_rwdir = function L2R -> mt() | R2L -> str "-" - -let wit_ssrdir = add_genarg "ssrdir" pr_dir - -(** Simpl switch *) - -let pr_ssrsimpl _ _ _ = pr_simpl - -let wit_ssrsimplrep = add_genarg "ssrsimplrep" pr_simpl - -let test_ssrslashnum b1 b2 strm = - match Util.stream_nth 0 strm with - | Tok.KEYWORD "/" -> - (match Util.stream_nth 1 strm with - | Tok.INT _ when b1 -> - (match Util.stream_nth 2 strm with - | Tok.KEYWORD "=" | Tok.KEYWORD "/=" when not b2 -> () - | Tok.KEYWORD "/" -> - if not b2 then () else begin - match Util.stream_nth 3 strm with - | Tok.INT _ -> () - | _ -> raise Stream.Failure - end - | _ -> raise Stream.Failure) - | Tok.KEYWORD "/" when not b1 -> - (match Util.stream_nth 2 strm with - | Tok.KEYWORD "=" when not b2 -> () - | Tok.INT _ when b2 -> - (match Util.stream_nth 3 strm with - | Tok.KEYWORD "=" -> () - | _ -> raise Stream.Failure) - | _ when not b2 -> () - | _ -> raise Stream.Failure) - | Tok.KEYWORD "=" when not b1 && not b2 -> () - | _ -> raise Stream.Failure) - | Tok.KEYWORD "//" when not b1 -> - (match Util.stream_nth 1 strm with - | Tok.KEYWORD "=" when not b2 -> () - | Tok.INT _ when b2 -> - (match Util.stream_nth 2 strm with - | Tok.KEYWORD "=" -> () - | _ -> raise Stream.Failure) - | _ when not b2 -> () - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure - -let test_ssrslashnum10 = test_ssrslashnum true false -let test_ssrslashnum11 = test_ssrslashnum true true -let test_ssrslashnum01 = test_ssrslashnum false true -let test_ssrslashnum00 = test_ssrslashnum false false - -let negate_parser f x = - let rc = try Some (f x) with Stream.Failure -> None in - match rc with - | None -> () - | Some _ -> raise Stream.Failure - -let test_not_ssrslashnum = - Pcoq.Gram.Entry.of_parser - "test_not_ssrslashnum" (negate_parser test_ssrslashnum10) -let test_ssrslashnum00 = - Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum00 -let test_ssrslashnum10 = - Pcoq.Gram.Entry.of_parser "test_ssrslashnum10" test_ssrslashnum10 -let test_ssrslashnum11 = - Pcoq.Gram.Entry.of_parser "test_ssrslashnum11" test_ssrslashnum11 -let test_ssrslashnum01 = - Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum01 - - -ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl -| [ "//=" ] -> [ SimplCut (~-1,~-1) ] -| [ "/=" ] -> [ Simpl ~-1 ] -END - -Pcoq.(Prim.( -GEXTEND Gram - GLOBAL: ssrsimpl_ne; - ssrsimpl_ne: [ - [ test_ssrslashnum11; "/"; n = natural; "/"; m = natural; "=" -> SimplCut(n,m) - | test_ssrslashnum10; "/"; n = natural; "/" -> Cut n - | test_ssrslashnum10; "/"; n = natural; "=" -> Simpl n - | test_ssrslashnum10; "/"; n = natural; "/=" -> SimplCut (n,~-1) - | test_ssrslashnum10; "/"; n = natural; "/"; "=" -> SimplCut (n,~-1) - | test_ssrslashnum01; "//"; m = natural; "=" -> SimplCut (~-1,m) - | test_ssrslashnum00; "//" -> Cut ~-1 - ]]; - -END -)) - -ARGUMENT EXTEND ssrsimpl TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl -| [ ssrsimpl_ne(sim) ] -> [ sim ] -| [ ] -> [ Nop ] -END - - -let pr_ssrclear _ _ _ = pr_clear mt - -ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyps PRINTED BY pr_ssrclear -| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ check_hyps_uniq [] clr; clr ] -END - -ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY pr_ssrclear -| [ ssrclear_ne(clr) ] -> [ clr ] -| [ ] -> [ [] ] -END - -(** Indexes *) - -(* Since SSR indexes are always positive numbers, we use the 0 value *) -(* to encode an omitted index. We reuse the in or_var type, but we *) -(* supply our own interpretation function, which checks for non *) -(* positive values, and allows the use of constr numerals, so that *) -(* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *) - - -let pr_index = function - | ArgVar {CAst.v=id} -> pr_id id - | ArgArg n when n > 0 -> int n - | _ -> mt () -let pr_ssrindex _ _ _ = pr_index - -let noindex = ArgArg 0 - -let check_index ?loc i = - if i > 0 then i else CErrors.user_err ?loc (str"Index not positive") -let mk_index ?loc = function - | ArgArg i -> ArgArg (check_index ?loc i) - | iv -> iv - -let interp_index ist gl idx = - Tacmach.project gl, - match idx with - | ArgArg _ -> idx - | ArgVar id -> - let i = - try - let v = Id.Map.find id.CAst.v ist.Tacinterp.lfun in - begin match Tacinterp.Value.to_int v with - | Some i -> i - | None -> - begin match Tacinterp.Value.to_constr v with - | Some c -> - let rc = Detyping.detype Detyping.Now false Id.Set.empty (pf_env gl) (project gl) c in - begin match Notation.uninterp_prim_token rc with - | _, Constrexpr.Numeral (s,b) -> - let n = int_of_string s in if b then n else -n - | _ -> raise Not_found - end - | None -> raise Not_found - end end - with _ -> CErrors.user_err ?loc:id.CAst.loc (str"Index not a number") in - ArgArg (check_index ?loc:id.CAst.loc i) - -open Pltac - -ARGUMENT EXTEND ssrindex PRINTED BY pr_ssrindex - INTERPRETED BY interp_index -| [ int_or_var(i) ] -> [ mk_index ~loc i ] -END - - -(** Occurrence switch *) - -(* The standard syntax of complemented occurrence lists involves a single *) -(* initial "-", e.g., {-1 3 5}. An initial *) -(* "+" may be used to indicate positive occurrences (the default). The *) -(* "+" is optional, except if the list of occurrences starts with a *) -(* variable or is empty (to avoid confusion with a clear switch). The *) -(* empty positive switch "{+}" selects no occurrences, while the empty *) -(* negative switch "{-}" selects all occurrences explicitly; this is the *) -(* default, but "{-}" prevents the implicit clear, and can be used to *) -(* force dependent elimination -- see ndefectelimtac below. *) - - -let pr_ssrocc _ _ _ = pr_occ - -open Pcoq.Prim - -ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY pr_ssrocc -| [ natural(n) natural_list(occ) ] -> [ - Some (false, List.map (check_index ~loc) (n::occ)) ] -| [ "-" natural_list(occ) ] -> [ Some (true, occ) ] -| [ "+" natural_list(occ) ] -> [ Some (false, occ) ] -END - - -(* modality *) - - -let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt () - -let wit_ssrmmod = add_genarg "ssrmmod" pr_mmod -let ssrmmod = Pcoq.create_generic_entry Pcoq.utactic "ssrmmod" (Genarg.rawwit wit_ssrmmod);; - -GEXTEND Gram - GLOBAL: ssrmmod; - ssrmmod: [[ "!" -> Must | LEFTQMARK -> May | "?" -> May]]; -END - -(** Rewrite multiplier: !n ?n *) - -let pr_mult (n, m) = - if n > 0 && m <> Once then int n ++ pr_mmod m else pr_mmod m - -let pr_ssrmult _ _ _ = pr_mult - -ARGUMENT EXTEND ssrmult_ne TYPED AS int * ssrmmod PRINTED BY pr_ssrmult - | [ natural(n) ssrmmod(m) ] -> [ check_index ~loc n, m ] - | [ ssrmmod(m) ] -> [ notimes, m ] -END - -ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY pr_ssrmult - | [ ssrmult_ne(m) ] -> [ m ] - | [ ] -> [ nomult ] -END - -(** Discharge occ switch (combined occurrence / clear switch *) - -let pr_docc = function - | None, occ -> pr_occ occ - | Some clr, _ -> pr_clear mt clr - -let pr_ssrdocc _ _ _ = pr_docc - -ARGUMENT EXTEND ssrdocc TYPED AS ssrclear option * ssrocc PRINTED BY pr_ssrdocc -| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] -| [ "{" ssrhyp_list(clr) "}" ] -> [ mkclr clr ] -END - -(* Old kinds of terms *) - -let input_ssrtermkind strm = match Util.stream_nth 0 strm with - | Tok.KEYWORD "(" -> xInParens - | Tok.KEYWORD "@" -> xWithAt - | _ -> xNoFlag - -let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind - -(* New kinds of terms *) - -let input_term_annotation strm = - match Stream.npeek 2 strm with - | Tok.KEYWORD "(" :: Tok.KEYWORD "(" :: _ -> `DoubleParens - | Tok.KEYWORD "(" :: _ -> `Parens - | Tok.KEYWORD "@" :: _ -> `At - | _ -> `None -let term_annotation = - Gram.Entry.of_parser "term_annotation" input_term_annotation - -(* terms *) - -(** Terms parsing. ********************************************************) - -(* Because we allow wildcards in term references, we need to stage the *) -(* interpretation of terms so that it occurs at the right time during *) -(* the execution of the tactic (e.g., so that we don't report an error *) -(* for a term that isn't actually used in the execution). *) -(* The term representation tracks whether the concrete initial term *) -(* started with an opening paren, which might avoid a conflict between *) -(* the ssrreflect term syntax and Gallina notation. *) - -(* Old terms *) -let pr_ssrterm _ _ _ = pr_term -let glob_ssrterm gs = function - | k, (_, Some c) -> k, Tacintern.intern_constr gs c - | ct -> ct -let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c -let interp_ssrterm _ gl t = Tacmach.project gl, t - -open Pcoq.Constr - -ARGUMENT EXTEND ssrterm - PRINTED BY pr_ssrterm - INTERPRETED BY interp_ssrterm - GLOBALIZED BY glob_ssrterm SUBSTITUTED BY subst_ssrterm - RAW_PRINTED BY pr_ssrterm - GLOB_PRINTED BY pr_ssrterm -| [ "YouShouldNotTypeThis" constr(c) ] -> [ mk_lterm c ] -END - - -GEXTEND Gram - GLOBAL: ssrterm; - ssrterm: [[ k = ssrtermkind; c = Pcoq.Constr.constr -> mk_term k c ]]; -END - -(* New terms *) - -let pp_ast_closure_term _ _ _ = pr_ast_closure_term - -ARGUMENT EXTEND ast_closure_term - PRINTED BY pp_ast_closure_term - INTERPRETED BY interp_ast_closure_term - GLOBALIZED BY glob_ast_closure_term - SUBSTITUTED BY subst_ast_closure_term - RAW_PRINTED BY pp_ast_closure_term - GLOB_PRINTED BY pp_ast_closure_term - | [ term_annotation(a) constr(c) ] -> [ mk_ast_closure_term a c ] -END -ARGUMENT EXTEND ast_closure_lterm - PRINTED BY pp_ast_closure_term - INTERPRETED BY interp_ast_closure_term - GLOBALIZED BY glob_ast_closure_term - SUBSTITUTED BY subst_ast_closure_term - RAW_PRINTED BY pp_ast_closure_term - GLOB_PRINTED BY pp_ast_closure_term - | [ term_annotation(a) lconstr(c) ] -> [ mk_ast_closure_term a c ] -END - -(* Old Views *) - -let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c) - -let pr_ssrbwdview _ _ _ = pr_view - -ARGUMENT EXTEND ssrbwdview TYPED AS ssrterm list - PRINTED BY pr_ssrbwdview -| [ "YouShouldNotTypeThis" ] -> [ [] ] -END - -Pcoq.( -GEXTEND Gram - GLOBAL: ssrbwdview; - ssrbwdview: [ - [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> [mk_term xNoFlag c] - | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrbwdview -> - (mk_term xNoFlag c) :: w ]]; -END -) - -(* New Views *) - - -let pr_ssrfwdview _ _ _ = pr_view2 - -ARGUMENT EXTEND ssrfwdview TYPED AS ast_closure_term list - PRINTED BY pr_ssrfwdview -| [ "YouShouldNotTypeThis" ] -> [ [] ] -END - -Pcoq.( -GEXTEND Gram - GLOBAL: ssrfwdview; - ssrfwdview: [ - [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> - [mk_ast_closure_term `None c] - | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrfwdview -> - (mk_ast_closure_term `None c) :: w ]]; -END -) - -(* }}} *) - -(* ipats *) - - -let remove_loc x = x.CAst.v - -let ipat_of_intro_pattern p = Tactypes.( - let rec ipat_of_intro_pattern = function - | IntroNaming (IntroIdentifier id) -> IPatId id - | IntroAction IntroWildcard -> IPatAnon Drop - | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> - IPatCase - (List.map (List.map ipat_of_intro_pattern) - (List.map (List.map remove_loc) iorpat)) - | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> - IPatCase - [List.map ipat_of_intro_pattern (List.map remove_loc iandpat)] - | IntroNaming IntroAnonymous -> IPatAnon One - | IntroAction (IntroRewrite b) -> IPatRewrite (allocc, if b then L2R else R2L) - | IntroNaming (IntroFresh id) -> IPatAnon One - | IntroAction (IntroApplyOn _) -> (* to do *) CErrors.user_err (Pp.str "TO DO") - | IntroAction (IntroInjection ips) -> - IPatInj [List.map ipat_of_intro_pattern (List.map remove_loc ips)] - | IntroForthcoming _ -> - (* Unable to determine which kind of ipat interp_introid could - * return [HH] *) - assert false - in - ipat_of_intro_pattern p -) - -let rec map_ipat map_id map_ssrhyp map_ast_closure_term = function - | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop) as x -> x - | IPatId id -> IPatId (map_id id) - | IPatAbstractVars l -> IPatAbstractVars (List.map map_id l) - | IPatClear clr -> IPatClear (List.map map_ssrhyp clr) - | IPatCase iorpat -> IPatCase (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) - | IPatDispatch iorpat -> IPatDispatch (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) - | IPatInj iorpat -> IPatInj (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) - | IPatView (clr,v) -> IPatView (clr,List.map map_ast_closure_term v) - | IPatTac _ -> assert false (*internal usage only *) - -let wit_ssripatrep = add_genarg "ssripatrep" pr_ipat - -let pr_ssripat _ _ _ = pr_ipat -let pr_ssripats _ _ _ = pr_ipats -let pr_ssriorpat _ _ _ = pr_iorpat - -let intern_ipat ist = - map_ipat - (fun id -> id) - (intern_hyp ist) - (glob_ast_closure_term ist) - -let intern_ipats ist = List.map (intern_ipat ist) - -let interp_intro_pattern = interp_wit wit_intro_pattern - -let interp_introid ist gl id = - try IntroNaming (IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (Loc.tag id)))))) - with _ -> (snd (interp_intro_pattern ist gl (CAst.make @@ IntroNaming (IntroIdentifier id)))).CAst.v - -let get_intro_id = function - | IntroNaming (IntroIdentifier id) -> id - | _ -> assert false - -let rec add_intro_pattern_hyps ipat hyps = - let {CAst.loc=loc;v=ipat} = ipat in - match ipat with - | IntroNaming (IntroIdentifier id) -> - if not_section_id id then SsrHyp (loc, id) :: hyps else - hyp_err ?loc "Can't delete section hypothesis " id - | IntroAction IntroWildcard -> hyps - | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> - List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps - | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> - List.fold_right add_intro_pattern_hyps iandpat hyps - | IntroNaming IntroAnonymous -> [] - | IntroNaming (IntroFresh _) -> [] - | IntroAction (IntroRewrite _) -> hyps - | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps - | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps - | IntroForthcoming _ -> - (* As in ipat_of_intro_pattern, was unable to determine which kind - of ipat interp_introid could return [HH] *) assert false - -(* We interp the ipat using the standard ltac machinery for ids, since - * we have no clue what a name could be bound to (maybe another ipat) *) -let interp_ipat ist gl = - let ltacvar id = Id.Map.mem id ist.Tacinterp.lfun in - let rec interp = function - | IPatId id when ltacvar id -> - ipat_of_intro_pattern (interp_introid ist gl id) - | IPatId _ as x -> x - | IPatClear clr -> - let add_hyps (SsrHyp (loc, id) as hyp) hyps = - if not (ltacvar id) then hyp :: hyps else - add_intro_pattern_hyps CAst.(make ?loc (interp_introid ist gl id)) hyps in - let clr' = List.fold_right add_hyps clr [] in - check_hyps_uniq [] clr'; IPatClear clr' - | IPatCase(iorpat) -> - IPatCase(List.map (List.map interp) iorpat) - | IPatDispatch(iorpat) -> - IPatDispatch(List.map (List.map interp) iorpat) - | IPatInj iorpat -> IPatInj (List.map (List.map interp) iorpat) - | IPatAbstractVars l -> - IPatAbstractVars (List.map get_intro_id (List.map (interp_introid ist gl) l)) - | IPatView (clr,l) -> IPatView (clr,List.map (fun x -> snd(interp_ast_closure_term ist - gl x)) l) - | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop) as x -> x - | IPatTac _ -> assert false (*internal usage only *) - in - interp - -let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l - -let pushIPatRewrite = function - | pats :: orpat -> (IPatRewrite (allocc, L2R) :: pats) :: orpat - | [] -> [] - -let pushIPatNoop = function - | pats :: orpat -> (IPatNoop :: pats) :: orpat - | [] -> [] - -ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY pr_ssripats - INTERPRETED BY interp_ipats - GLOBALIZED BY intern_ipats - | [ "_" ] -> [ [IPatAnon Drop] ] - | [ "*" ] -> [ [IPatAnon All] ] - (* - | [ "^" "*" ] -> [ [IPatFastMode] ] - | [ "^" "_" ] -> [ [IPatSeed `Wild] ] - | [ "^_" ] -> [ [IPatSeed `Wild] ] - | [ "^" "?" ] -> [ [IPatSeed `Anon] ] - | [ "^?" ] -> [ [IPatSeed `Anon] ] - | [ "^" ident(id) ] -> [ [IPatSeed (`Id(id,`Pre))] ] - | [ "^" "~" ident(id) ] -> [ [IPatSeed (`Id(id,`Post))] ] - | [ "^~" ident(id) ] -> [ [IPatSeed (`Id(id,`Post))] ] - *) - | [ ident(id) ] -> [ [IPatId id] ] - | [ "?" ] -> [ [IPatAnon One] ] -(* TODO | [ "+" ] -> [ [IPatAnon Temporary] ] *) - | [ ssrsimpl_ne(sim) ] -> [ [IPatSimpl sim] ] - | [ ssrdocc(occ) "->" ] -> [ match occ with - | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") - | None, occ -> [IPatRewrite (occ, L2R)] - | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, L2R)]] - | [ ssrdocc(occ) "<-" ] -> [ match occ with - | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") - | None, occ -> [IPatRewrite (occ, R2L)] - | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, R2L)]] - | [ ssrdocc(occ) ssrfwdview(v) ] -> [ match occ with - | Some [], _ -> [IPatView (true,v)] - | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl;IPatView (false,v)] - | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here") ] - | [ ssrdocc(occ) ] -> [ match occ with - | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl] - | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here")] - | [ "->" ] -> [ [IPatRewrite (allocc, L2R)] ] - | [ "<-" ] -> [ [IPatRewrite (allocc, R2L)] ] - | [ "-" ] -> [ [IPatNoop] ] - | [ "-/" "=" ] -> [ [IPatNoop;IPatSimpl(Simpl ~-1)] ] - | [ "-/=" ] -> [ [IPatNoop;IPatSimpl(Simpl ~-1)] ] - | [ "-/" "/" ] -> [ [IPatNoop;IPatSimpl(Cut ~-1)] ] - | [ "-//" ] -> [ [IPatNoop;IPatSimpl(Cut ~-1)] ] - | [ "-/" integer(n) "/" ] -> [ [IPatNoop;IPatSimpl(Cut n)] ] - | [ "-/" "/=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ] - | [ "-//" "=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ] - | [ "-//=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] ] - | [ "-/" integer(n) "/=" ] -> [ [IPatNoop;IPatSimpl(SimplCut (n,~-1))] ] - | [ "-/" integer(n) "/" integer (m) "=" ] -> - [ [IPatNoop;IPatSimpl(SimplCut(n,m))] ] - | [ ssrfwdview(v) ] -> [ [IPatView (false,v)] ] - | [ "[" ":" ident_list(idl) "]" ] -> [ [IPatAbstractVars idl] ] - | [ "[:" ident_list(idl) "]" ] -> [ [IPatAbstractVars idl] ] -END - -ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY pr_ssripats - | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] - | [ ] -> [ [] ] -END - -ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY pr_ssriorpat -| [ ssripats(pats) "|" ssriorpat(orpat) ] -> [ pats :: orpat ] -| [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> [ pats :: pushIPatRewrite orpat ] -| [ ssripats(pats) "|-" ssriorpat(orpat) ] -> [ pats :: pushIPatNoop orpat ] -| [ ssripats(pats) "|->" ssriorpat(orpat) ] -> [ pats :: pushIPatRewrite orpat ] -| [ ssripats(pats) "||" ssriorpat(orpat) ] -> [ pats :: [] :: orpat ] -| [ ssripats(pats) "|||" ssriorpat(orpat) ] -> [ pats :: [] :: [] :: orpat ] -| [ ssripats(pats) "||||" ssriorpat(orpat) ] -> [ [pats; []; []; []] @ orpat ] -| [ ssripats(pats) ] -> [ [pats] ] -END - -let reject_ssrhid strm = - match Util.stream_nth 0 strm with - | Tok.KEYWORD "[" -> - (match Util.stream_nth 1 strm with - | Tok.KEYWORD ":" -> raise Stream.Failure - | _ -> ()) - | _ -> () - -let test_nohidden = Pcoq.Gram.Entry.of_parser "test_ssrhid" reject_ssrhid - -ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY pr_ssripat - | [ "YouShouldNotTypeThis" ssriorpat(x) ] -> [ IPatCase(x) ] -END - -Pcoq.( -GEXTEND Gram - GLOBAL: ssrcpat; - ssrcpat: [ - [ test_nohidden; "["; iorpat = ssriorpat; "]" -> -(* check_no_inner_seed !@loc false iorpat; - IPatCase (understand_case_type iorpat) *) - IPatCase iorpat -(* - | test_nohidden; "("; iorpat = ssriorpat; ")" -> -(* check_no_inner_seed !@loc false iorpat; - IPatCase (understand_case_type iorpat) *) - IPatDispatch iorpat -*) - | test_nohidden; "[="; iorpat = ssriorpat; "]" -> -(* check_no_inner_seed !@loc false iorpat; *) - IPatInj iorpat ]]; -END -);; - -Pcoq.( -GEXTEND Gram - GLOBAL: ssripat; - ssripat: [[ pat = ssrcpat -> [pat] ]]; -END -) - -ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY pr_ssripats - | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] - END - -(* subsets of patterns *) - -(* TODO: review what this function does, it looks suspicious *) -let check_ssrhpats loc w_binders ipats = - let err_loc s = CErrors.user_err ~loc ~hdr:"ssreflect" s in - let clr, ipats = - let rec aux clr = function - | IPatClear cl :: tl -> aux (clr @ cl) tl -(* | IPatSimpl (cl, sim) :: tl -> clr @ cl, IPatSimpl ([], sim) :: tl *) - | tl -> clr, tl - in aux [] ipats in - let simpl, ipats = - match List.rev ipats with - | IPatSimpl _ as s :: tl -> [s], List.rev tl - | _ -> [], ipats in - if simpl <> [] && not w_binders then - err_loc (str "No s-item allowed here: " ++ pr_ipats simpl); - let ipat, binders = - let rec loop ipat = function - | [] -> ipat, [] - | ( IPatId _| IPatAnon _| IPatCase _ | IPatDispatch _ | IPatRewrite _ as i) :: tl -> - if w_binders then - if simpl <> [] && tl <> [] then - err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) - else if not (List.for_all (function IPatId _ -> true | _ -> false) tl) - then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) - else ipat @ [i], tl - else - if tl = [] then ipat @ [i], [] - else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl) - | hd :: tl -> loop (ipat @ [hd]) tl - in loop [] ipats in - ((clr, ipat), binders), simpl - -let pr_hpats (((clr, ipat), binders), simpl) = - pr_clear mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl -let pr_ssrhpats _ _ _ = pr_hpats -let pr_ssrhpats_wtransp _ _ _ (_, x) = pr_hpats x - -ARGUMENT EXTEND ssrhpats TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat -PRINTED BY pr_ssrhpats - | [ ssripats(i) ] -> [ check_ssrhpats loc true i ] -END - -ARGUMENT EXTEND ssrhpats_wtransp - TYPED AS bool * (((ssrclear * ssripats) * ssripats) * ssripats) - PRINTED BY pr_ssrhpats_wtransp - | [ ssripats(i) ] -> [ false,check_ssrhpats loc true i ] - | [ ssripats(i) "@" ssripats(j) ] -> [ true,check_ssrhpats loc true (i @ j) ] -END - -ARGUMENT EXTEND ssrhpats_nobs -TYPED AS ((ssrclear * ssripats) * ssripats) * ssripats PRINTED BY pr_ssrhpats - | [ ssripats(i) ] -> [ check_ssrhpats loc false i ] -END - -ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY pr_ssripat - | [ "->" ] -> [ IPatRewrite (allocc, L2R) ] - | [ "<-" ] -> [ IPatRewrite (allocc, R2L) ] -END - -let pr_intros sep intrs = - if intrs = [] then mt() else sep () ++ str "=>" ++ pr_ipats intrs -let pr_ssrintros _ _ _ = pr_intros mt - -ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat - PRINTED BY pr_ssrintros - | [ "=>" ssripats_ne(pats) ] -> [ pats ] -(* TODO | [ "=>" ">" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ] - | [ "=>>" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ] *) -END - -ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY pr_ssrintros - | [ ssrintros_ne(intrs) ] -> [ intrs ] - | [ ] -> [ [] ] -END - -let pr_ssrintrosarg _ _ prt (tac, ipats) = - prt tacltop tac ++ pr_intros spc ipats - -ARGUMENT EXTEND ssrintrosarg TYPED AS tactic * ssrintros - PRINTED BY pr_ssrintrosarg -| [ "YouShouldNotTypeThis" ssrtacarg(arg) ssrintros_ne(ipats) ] -> [ arg, ipats ] -END - -TACTIC EXTEND ssrtclintros -| [ "YouShouldNotTypeThis" ssrintrosarg(arg) ] -> - [ let tac, intros = arg in - ssrevaltac ist tac <*> tclIPATssr intros ] -END - -(** Defined identifier *) -let pr_ssrfwdid id = pr_spc () ++ pr_id id - -let pr_ssrfwdidx _ _ _ = pr_ssrfwdid - -(* We use a primitive parser for the head identifier of forward *) -(* tactis to avoid syntactic conflicts with basic Coq tactics. *) -ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY pr_ssrfwdidx - | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -let accept_ssrfwdid strm = - match stream_nth 0 strm with - | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm - | _ -> raise Stream.Failure - - -let test_ssrfwdid = Gram.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid - -GEXTEND Gram - GLOBAL: ssrfwdid; - ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> id ]]; - END - - -(* by *) -(** Tactical arguments. *) - -(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) -(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) -(* and subgoal reordering tacticals (; first & ; last), respectively. *) - - -let pr_ortacs prt = - let rec pr_rec = function - | [None] -> spc() ++ str "|" ++ spc() - | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs - | Some tac :: tacs -> spc() ++ str "| " ++ prt tacltop tac ++ pr_rec tacs - | [] -> mt() in - function - | [None] -> spc() - | None :: tacs -> pr_rec tacs - | Some tac :: tacs -> prt tacltop tac ++ pr_rec tacs - | [] -> mt() -let pr_ssrortacs _ _ = pr_ortacs - -ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY pr_ssrortacs -| [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> [ Some tac :: tacs ] -| [ ssrtacarg(tac) "|" ] -> [ [Some tac; None] ] -| [ ssrtacarg(tac) ] -> [ [Some tac] ] -| [ "|" ssrortacs(tacs) ] -> [ None :: tacs ] -| [ "|" ] -> [ [None; None] ] -END - -let pr_hintarg prt = function - | true, tacs -> hv 0 (str "[ " ++ pr_ortacs prt tacs ++ str " ]") - | false, [Some tac] -> prt tacltop tac - | _, _ -> mt() - -let pr_ssrhintarg _ _ = pr_hintarg - - -ARGUMENT EXTEND ssrhintarg TYPED AS bool * ssrortacs PRINTED BY pr_ssrhintarg -| [ "[" "]" ] -> [ nullhint ] -| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] -| [ ssrtacarg(arg) ] -> [ mk_hint arg ] -END - -ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY pr_ssrhintarg -| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] -END - - -let pr_hint prt arg = - if arg = nohint then mt() else str "by " ++ pr_hintarg prt arg -let pr_ssrhint _ _ = pr_hint - -ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY pr_ssrhint -| [ ] -> [ nohint ] -END -(** The "in" pseudo-tactical *)(* {{{ **********************************************) - -(* We can't make "in" into a general tactical because this would create a *) -(* crippling conflict with the ltac let .. in construct. Hence, we add *) -(* explicitly an "in" suffix to all the extended tactics for which it is *) -(* relevant (including move, case, elim) and to the extended do tactical *) -(* below, which yields a general-purpose "in" of the form do [...] in ... *) - -(* This tactical needs to come before the intro tactics because the latter *) -(* must take precautions in order not to interfere with the discharged *) -(* assumptions. This is especially difficult for discharged "let"s, which *) -(* the default simpl and unfold tactics would erase blindly. *) - -open Ssrmatching_plugin.Ssrmatching -open Ssrmatching_plugin.G_ssrmatching - -let pr_wgen = function - | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id - | (clr, Some((id,k),Some p)) -> - spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++ - pr_cpattern p ++ str ")" - | (clr, None) -> spc () ++ pr_clear mt clr -let pr_ssrwgen _ _ _ = pr_wgen - -(* no globwith for char *) -ARGUMENT EXTEND ssrwgen - TYPED AS ssrclear * ((ssrhoi_hyp * string) * cpattern option) option - PRINTED BY pr_ssrwgen -| [ ssrclear_ne(clr) ] -> [ clr, None ] -| [ ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, " "), None) ] -| [ "@" ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, "@"), None) ] -| [ "(" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> - [ [], Some ((id," "),Some p) ] -| [ "(" ssrhoi_id(id) ")" ] -> [ [], Some ((id,"("), None) ] -| [ "(@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> - [ [], Some ((id,"@"),Some p) ] -| [ "(" "@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> - [ [], Some ((id,"@"),Some p) ] -END - -let pr_clseq = function - | InGoal | InHyps -> mt () - | InSeqGoal -> str "|- *" - | InHypsSeqGoal -> str " |- *" - | InHypsGoal -> str " *" - | InAll -> str "*" - | InHypsSeq -> str " |-" - | InAllHyps -> str "* |-" - -let wit_ssrclseq = add_genarg "ssrclseq" pr_clseq -let pr_clausehyps = pr_list pr_spc pr_wgen -let pr_ssrclausehyps _ _ _ = pr_clausehyps - -ARGUMENT EXTEND ssrclausehyps -TYPED AS ssrwgen list PRINTED BY pr_ssrclausehyps -| [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> [ hyp :: hyps ] -| [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> [ hyp :: hyps ] -| [ ssrwgen(hyp) ] -> [ [hyp] ] -END - -(* type ssrclauses = ssrahyps * ssrclseq *) - -let pr_clauses (hyps, clseq) = - if clseq = InGoal then mt () - else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq -let pr_ssrclauses _ _ _ = pr_clauses - -ARGUMENT EXTEND ssrclauses TYPED AS ssrwgen list * ssrclseq - PRINTED BY pr_ssrclauses - | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> [ hyps, InHypsSeqGoal ] - | [ "in" ssrclausehyps(hyps) "|-" ] -> [ hyps, InHypsSeq ] - | [ "in" ssrclausehyps(hyps) "*" ] -> [ hyps, InHypsGoal ] - | [ "in" ssrclausehyps(hyps) ] -> [ hyps, InHyps ] - | [ "in" "|-" "*" ] -> [ [], InSeqGoal ] - | [ "in" "*" ] -> [ [], InAll ] - | [ "in" "*" "|-" ] -> [ [], InAllHyps ] - | [ ] -> [ [], InGoal ] -END - - - - -(** Definition value formatting *) - -(* We use an intermediate structure to correctly render the binder list *) -(* abbreviations. We use a list of hints to extract the binders and *) -(* base term from a term, for the two first levels of representation of *) -(* of constr terms. *) - -let pr_binder prl = function - | Bvar x -> - pr_name x - | Bdecl (xs, t) -> - str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")" - | Bdef (x, None, v) -> - str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")" - | Bdef (x, Some t, v) -> - str "(" ++ pr_name x ++ str " : " ++ prl t ++ - str " := " ++ prl v ++ str ")" - | Bstruct x -> - str "{struct " ++ pr_name x ++ str "}" - | Bcast t -> - str ": " ++ prl t - -let rec format_local_binders h0 bl0 = match h0, bl0 with - | BFvar :: h, CLocalAssum ([{CAst.v=x}], _, _) :: bl -> - Bvar x :: format_local_binders h bl - | BFdecl _ :: h, CLocalAssum (lxs, _, t) :: bl -> - Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: format_local_binders h bl - | BFdef :: h, CLocalDef ({CAst.v=x}, v, oty) :: bl -> - Bdef (x, oty, v) :: format_local_binders h bl - | _ -> [] - -let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with - | BFvar :: h, { v = CLambdaN ([CLocalAssum([{CAst.v=x}], _, _)], c) } -> - let bs, c' = format_constr_expr h c in - Bvar x :: bs, c' - | BFdecl _:: h, { v = CLambdaN ([CLocalAssum(lxs, _, t)], c) } -> - let bs, c' = format_constr_expr h c in - Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: bs, c' - | BFdef :: h, { v = CLetIn({CAst.v=x}, v, oty, c) } -> - let bs, c' = format_constr_expr h c in - Bdef (x, oty, v) :: bs, c' - | [BFcast], { v = CCast (c, Glob_term.CastConv t) } -> - [Bcast t], c - | BFrec (has_str, has_cast) :: h, - { v = CFix ( _, [_, (Some locn, CStructRec), bl, t, c]) } -> - let bs = format_local_binders h bl in - let bstr = if has_str then [Bstruct (Name locn.CAst.v)] else [] in - bs @ bstr @ (if has_cast then [Bcast t] else []), c - | BFrec (_, has_cast) :: h, { v = CCoFix ( _, [_, bl, t, c]) } -> - format_local_binders h bl @ (if has_cast then [Bcast t] else []), c - | _, c -> - [], c - -(** Forward chaining argument *) - -(* There are three kinds of forward definitions: *) -(* - Hint: type only, cast to Type, may have proof hint. *) -(* - Have: type option + value, no space before type *) -(* - Pose: binders + value, space before binders. *) - -let pr_fwdkind = function - | FwdHint (s,_) -> str (s ^ " ") | _ -> str " :=" ++ spc () -let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk - -let wit_ssrfwdfmt = add_genarg "ssrfwdfmt" pr_fwdfmt - -(* type ssrfwd = ssrfwdfmt * ssrterm *) - -let mkFwdVal fk c = ((fk, []), c) -let mkssrFwdVal fk c = ((fk, []), (c,None)) -let dC t = Glob_term.CastConv t - -let same_ist { interp_env = x } { interp_env = y } = - match x,y with - | None, None -> true - | Some a, Some b -> a == b - | _ -> false - -let mkFwdCast fk ?loc ?c t = - let c = match c with - | None -> mkCHole loc - | Some c -> assert (same_ist t c); c.body in - ((fk, [BFcast]), - { t with annotation = `None; - body = (CAst.make ?loc @@ CCast (c, dC t.body)) }) - -let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t)) - -let mkFwdHint s t = - let loc = Constrexpr_ops.constr_loc t.body in - mkFwdCast (FwdHint (s,false)) ?loc t -let mkFwdHintNoTC s t = - let loc = Constrexpr_ops.constr_loc t.body in - mkFwdCast (FwdHint (s,true)) ?loc t - -let pr_gen_fwd prval prc prlc fk (bs, c) = - let prc s = str s ++ spc () ++ prval prc prlc c in - match fk, bs with - | FwdHint (s,_), [Bcast t] -> str s ++ spc () ++ prlc t - | FwdHint (s,_), _ -> prc (s ^ "(* typeof *)") - | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :=" - | _, [] -> prc " :=" - | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :=" - -let pr_fwd_guarded prval prval' = function -| (fk, h), c -> - pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c.body) - -let pr_unguarded prc prlc = prlc - -let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded -let pr_ssrfwd _ _ _ = pr_fwd - -ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ast_closure_lterm) PRINTED BY pr_ssrfwd - | [ ":=" ast_closure_lterm(c) ] -> [ mkFwdVal FwdPose c ] - | [ ":" ast_closure_lterm (t) ":=" ast_closure_lterm(c) ] -> [ mkFwdCast FwdPose ~loc t ~c ] -END - -(** Independent parsing for binders *) - -(* The pose, pose fix, and pose cofix tactics use these internally to *) -(* parse argument fragments. *) - -let pr_ssrbvar prc _ _ v = prc v - -ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY pr_ssrbvar -| [ ident(id) ] -> [ mkCVar ~loc id ] -| [ "_" ] -> [ mkCHole (Some loc) ] -END - -let bvar_lname = let open CAst in function - | { v = CRef (qid, _) } when qualid_is_ident qid -> - CAst.make ?loc:qid.CAst.loc @@ Name (qualid_basename qid) - | { loc = loc } -> CAst.make ?loc Anonymous - -let pr_ssrbinder prc _ _ (_, c) = prc c - -ARGUMENT EXTEND ssrbinder TYPED AS ssrfwdfmt * constr PRINTED BY pr_ssrbinder - | [ ssrbvar(bv) ] -> - [ let { CAst.loc=xloc } as x = bvar_lname bv in - (FwdPose, [BFvar]), - CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) ] - | [ "(" ssrbvar(bv) ")" ] -> - [ let { CAst.loc=xloc } as x = bvar_lname bv in - (FwdPose, [BFvar]), - CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) ] - | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] -> - [ let x = bvar_lname bv in - (FwdPose, [BFdecl 1]), - CAst.make ~loc @@ CLambdaN ([CLocalAssum([x], Default Explicit, t)], mkCHole (Some loc)) ] - | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] -> - [ let xs = List.map bvar_lname (bv :: bvs) in - let n = List.length xs in - (FwdPose, [BFdecl n]), - CAst.make ~loc @@ CLambdaN ([CLocalAssum (xs, Default Explicit, t)], mkCHole (Some loc)) ] - | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] -> - [ (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, Some t, mkCHole (Some loc)) ] - | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] -> - [ (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, None, mkCHole (Some loc)) ] - END - -GEXTEND Gram - GLOBAL: ssrbinder; - ssrbinder: [ - [ ["of" | "&"]; c = operconstr LEVEL "99" -> - let loc = !@loc in - (FwdPose, [BFvar]), - CAst.make ~loc @@ CLambdaN ([CLocalAssum ([CAst.make ~loc Anonymous],Default Explicit,c)],mkCHole (Some loc)) ] - ]; -END - -let rec binders_fmts = function - | ((_, h), _) :: bs -> h @ binders_fmts bs - | _ -> [] - -let push_binders c2 bs = - let loc2 = constr_loc c2 in let mkloc loc1 = Loc.merge_opt loc1 loc2 in - let open CAst in - let rec loop ty c = function - | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs when ty -> - CAst.make ?loc:(mkloc loc1) @@ CProdN (b, loop ty c bs) - | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs -> - CAst.make ?loc:(mkloc loc1) @@ CLambdaN (b, loop ty c bs) - | (_, { loc = loc1; v = CLetIn (x, v, oty, _) } ) :: bs -> - CAst.make ?loc:(mkloc loc1) @@ CLetIn (x, v, oty, loop ty c bs) - | [] -> c - | _ -> anomaly "binder not a lambda nor a let in" in - match c2 with - | { loc; v = CCast (ct, Glob_term.CastConv cty) } -> - CAst.make ?loc @@ (CCast (loop false ct bs, Glob_term.CastConv (loop true cty bs))) - | ct -> loop false ct bs - -let rec fix_binders = let open CAst in function - | (_, { v = CLambdaN ([CLocalAssum(xs, _, t)], _) } ) :: bs -> - CLocalAssum (xs, Default Explicit, t) :: fix_binders bs - | (_, { v = CLetIn (x, v, oty, _) } ) :: bs -> - CLocalDef (x, v, oty) :: fix_binders bs - | _ -> [] - -let pr_ssrstruct _ _ _ = function - | Some id -> str "{struct " ++ pr_id id ++ str "}" - | None -> mt () - -ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY pr_ssrstruct -| [ "{" "struct" ident(id) "}" ] -> [ Some id ] -| [ ] -> [ None ] -END - -(** The "pose" tactic *) - -(* The plain pose form. *) - -let bind_fwd bs ((fk, h), c) = - (fk,binders_fmts bs @ h), { c with body = push_binders c.body bs } - -ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY pr_ssrfwd - | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> [ bind_fwd bs fwd ] -END - -(* The pose fix form. *) - -let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd - -let bvar_locid = function - | { CAst.v = CRef (qid, _) } when qualid_is_ident qid -> - CAst.make ?loc:qid.CAst.loc (qualid_basename qid) - | _ -> CErrors.user_err (Pp.str "Missing identifier after \"(co)fix\"") - - -ARGUMENT EXTEND ssrfixfwd TYPED AS ident * ssrfwd PRINTED BY pr_ssrfixfwd - | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] -> - [ let { CAst.v=id } as lid = bvar_locid bv in - let (fk, h), ac = fwd in - let c = ac.body in - let has_cast, t', c' = match format_constr_expr h c with - | [Bcast t'], c' -> true, t', c' - | _ -> false, mkCHole (constr_loc c), c in - let lb = fix_binders bs in - let has_struct, i = - let rec loop = function - | {CAst.loc=l'; v=Name id'} :: _ when Option.equal Id.equal sid (Some id') -> - true, CAst.make ?loc:l' id' - | [{CAst.loc=l';v=Name id'}] when sid = None -> - false, CAst.make ?loc:l' id' - | _ :: bn -> loop bn - | [] -> CErrors.user_err (Pp.str "Bad structural argument") in - loop (names_of_local_assums lb) in - let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in - let fix = CAst.make ~loc @@ CFix (lid, [lid, (Some i, CStructRec), lb, t', c']) in - id, ((fk, h'), { ac with body = fix }) ] -END - - -(* The pose cofix form. *) - -let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd - -ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY pr_ssrcofixfwd - | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] -> - [ let { CAst.v=id } as lid = bvar_locid bv in - let (fk, h), ac = fwd in - let c = ac.body in - let has_cast, t', c' = match format_constr_expr h c with - | [Bcast t'], c' -> true, t', c' - | _ -> false, mkCHole (constr_loc c), c in - let h' = BFrec (false, has_cast) :: binders_fmts bs in - let cofix = CAst.make ~loc @@ CCoFix (lid, [lid, fix_binders bs, t', c']) in - id, ((fk, h'), { ac with body = cofix }) - ] -END - -(* This does not print the type, it should be fixed... *) -let pr_ssrsetfwd _ _ _ (((fk,_),(t,_)), docc) = - pr_gen_fwd (fun _ _ -> pr_cpattern) - (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t) - -ARGUMENT EXTEND ssrsetfwd -TYPED AS (ssrfwdfmt * (lcpattern * ast_closure_lterm option)) * ssrdocc -PRINTED BY pr_ssrsetfwd -| [ ":" ast_closure_lterm(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> - [ mkssrFwdCast FwdPose loc t c, mkocc occ ] -| [ ":" ast_closure_lterm(t) ":=" lcpattern(c) ] -> - [ mkssrFwdCast FwdPose loc t c, nodocc ] -| [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> - [ mkssrFwdVal FwdPose c, mkocc occ ] -| [ ":=" lcpattern(c) ] -> [ mkssrFwdVal FwdPose c, nodocc ] -END - - -let pr_ssrhavefwd _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint prt hint - -ARGUMENT EXTEND ssrhavefwd TYPED AS ssrfwd * ssrhint PRINTED BY pr_ssrhavefwd -| [ ":" ast_closure_lterm(t) ssrhint(hint) ] -> [ mkFwdHint ":" t, hint ] -| [ ":" ast_closure_lterm(t) ":=" ast_closure_lterm(c) ] -> [ mkFwdCast FwdHave ~loc t ~c, nohint ] -| [ ":" ast_closure_lterm(t) ":=" ] -> [ mkFwdHintNoTC ":" t, nohint ] -| [ ":=" ast_closure_lterm(c) ] -> [ mkFwdVal FwdHave c, nohint ] -END - -let intro_id_to_binder = List.map (function - | IPatId id -> - let { CAst.loc=xloc } as x = bvar_lname (mkCVar id) in - (FwdPose, [BFvar]), - CAst.make @@ CLambdaN ([CLocalAssum([x], Default Explicit, mkCHole xloc)], - mkCHole None) - | _ -> anomaly "non-id accepted as binder") - -let binder_to_intro_id = CAst.(List.map (function - | (FwdPose, [BFvar]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } - | (FwdPose, [BFdecl _]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } -> - List.map (function {v=Name id} -> IPatId id | _ -> IPatAnon One) ids - | (FwdPose, [BFdef]), { v = CLetIn ({v=Name id},_,_,_) } -> [IPatId id] - | (FwdPose, [BFdef]), { v = CLetIn ({v=Anonymous},_,_,_) } -> [IPatAnon One] - | _ -> anomaly "ssrbinder is not a binder")) - -let pr_ssrhavefwdwbinders _ _ prt (tr,((hpats, (fwd, hint)))) = - pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint - -ARGUMENT EXTEND ssrhavefwdwbinders - TYPED AS bool * (ssrhpats * (ssrfwd * ssrhint)) - PRINTED BY pr_ssrhavefwdwbinders -| [ ssrhpats_wtransp(trpats) ssrbinder_list(bs) ssrhavefwd(fwd) ] -> - [ let tr, pats = trpats in - let ((clr, pats), binders), simpl = pats in - let allbs = intro_id_to_binder binders @ bs in - let allbinders = binders @ List.flatten (binder_to_intro_id bs) in - let hint = bind_fwd allbs (fst fwd), snd fwd in - tr, ((((clr, pats), allbinders), simpl), hint) ] -END - - -let pr_ssrdoarg prc _ prt (((n, m), tac), clauses) = - pr_index n ++ pr_mmod m ++ pr_hintarg prt tac ++ pr_clauses clauses - -ARGUMENT EXTEND ssrdoarg - TYPED AS ((ssrindex * ssrmmod) * ssrhintarg) * ssrclauses - PRINTED BY pr_ssrdoarg -| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -(* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *) - -let pr_seqtacarg prt = function - | (is_first, []), _ -> str (if is_first then "first" else "last") - | tac, Some dtac -> - hv 0 (pr_hintarg prt tac ++ spc() ++ str "|| " ++ prt tacltop dtac) - | tac, _ -> pr_hintarg prt tac - -let pr_ssrseqarg _ _ prt = function - | ArgArg 0, tac -> pr_seqtacarg prt tac - | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg prt tac - -(* We must parse the index separately to resolve the conflict with *) -(* an unindexed tactic. *) -ARGUMENT EXTEND ssrseqarg TYPED AS ssrindex * (ssrhintarg * tactic option) - PRINTED BY pr_ssrseqarg -| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -let sq_brace_tacnames = - ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"] - (* "by" is a keyword *) -let accept_ssrseqvar strm = - match stream_nth 0 strm with - | Tok.IDENT id when not (List.mem id sq_brace_tacnames) -> - accept_before_syms_or_ids ["["] ["first";"last"] strm - | _ -> raise Stream.Failure - -let test_ssrseqvar = Gram.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar - -let swaptacarg (loc, b) = (b, []), Some (TacId []) - -let check_seqtacarg dir arg = match snd arg, dir with - | ((true, []), Some (TacAtom (loc, _))), L2R -> - CErrors.user_err ?loc (str "expected \"last\"") - | ((false, []), Some (TacAtom (loc, _))), R2L -> - CErrors.user_err ?loc (str "expected \"first\"") - | _, _ -> arg - -let ssrorelse = Entry.create "ssrorelse" -GEXTEND Gram - GLOBAL: ssrorelse ssrseqarg; - ssrseqidx: [ - [ test_ssrseqvar; id = Prim.ident -> ArgVar (CAst.make ~loc:!@loc id) - | n = Prim.natural -> ArgArg (check_index ~loc:!@loc n) - ] ]; - ssrswap: [[ IDENT "first" -> !@loc, true | IDENT "last" -> !@loc, false ]]; - ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> tac ]]; - ssrseqarg: [ - [ arg = ssrswap -> noindex, swaptacarg arg - | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> i, (tac, def) - | i = ssrseqidx; arg = ssrswap -> i, swaptacarg arg - | tac = tactic_expr LEVEL "3" -> noindex, (mk_hint tac, None) - ] ]; -END - -let tactic_expr = Pltac.tactic_expr - -(** 1. Utilities *) - -(** Tactic-level diagnosis *) - -(* debug *) - -(* Let's play with the new proof engine API *) -let old_tac = V82.tactic - - -(** Name generation *)(* {{{ *******************************************************) - -(* Since Coq now does repeated internal checks of its external lexical *) -(* rules, we now need to carve ssreflect reserved identifiers out of *) -(* out of the user namespace. We use identifiers of the form _id_ for *) -(* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *) -(* an extra leading _ if this might clash with an internal identifier. *) -(* We check for ssreflect identifiers in the ident grammar rule; *) -(* when the ssreflect Module is present this is normally an error, *) -(* but we provide a compatibility flag to reduce this to a warning. *) - -let ssr_reserved_ids = Summary.ref ~name:"SSR:idents" true - -let _ = - Goptions.declare_bool_option - { Goptions.optname = "ssreflect identifiers"; - Goptions.optkey = ["SsrIdents"]; - Goptions.optdepr = false; - Goptions.optread = (fun _ -> !ssr_reserved_ids); - Goptions.optwrite = (fun b -> ssr_reserved_ids := b) - } - -let is_ssr_reserved s = - let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_' - -let ssr_id_of_string loc s = - if is_ssr_reserved s && is_ssr_loaded () then begin - if !ssr_reserved_ids then - CErrors.user_err ~loc (str ("The identifier " ^ s ^ " is reserved.")) - else if is_internal_name s then - Feedback.msg_warning (str ("Conflict between " ^ s ^ " and ssreflect internal names.")) - else Feedback.msg_warning (str ( - "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" - ^ "Scripts with explicit references to anonymous variables are fragile.")) - end; Id.of_string s - -let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ()) - -let (!@) = Pcoq.to_coqloc - -GEXTEND Gram - GLOBAL: Prim.ident; - Prim.ident: [[ s = IDENT; ssr_null_entry -> ssr_id_of_string !@loc s ]]; -END - -let perm_tag = "_perm_Hyp_" -let _ = add_internal_name (is_tagged perm_tag) - -(* }}} *) - -(* We must not anonymize context names discharged by the "in" tactical. *) - -(** Tactical extensions. *)(* {{{ **************************************************) - -(* The TACTIC EXTEND facility can't be used for defining new user *) -(* tacticals, because: *) -(* - the concrete syntax must start with a fixed string *) -(* We use the following workaround: *) -(* - We use the (unparsable) "YouShouldNotTypeThis" token for tacticals that *) -(* don't start with a token, then redefine the grammar and *) -(* printer using GEXTEND and set_pr_ssrtac, respectively. *) - -type ssrargfmt = ArgSsr of string | ArgSep of string - -let ssrtac_name name = { - mltac_plugin = "ssreflect_plugin"; - mltac_tactic = "ssr" ^ name; -} - -let ssrtac_entry name n = { - mltac_name = ssrtac_name name; - mltac_index = n; -} - -let set_pr_ssrtac name prec afmt = (* FIXME *) () (* - let fmt = List.map (function - | ArgSep s -> Egramml.GramTerminal s - | ArgSsr s -> Egramml.GramTerminal s - | ArgCoq at -> Egramml.GramTerminal "COQ_ARG") afmt in - let tacname = ssrtac_name name in () *) - -let ssrtac_atom ?loc name args = TacML (Loc.tag ?loc (ssrtac_entry name 0, args)) -let ssrtac_expr ?loc name args = ssrtac_atom ?loc name args - -let tclintros_expr ?loc tac ipats = - let args = [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrintrosarg) (tac, ipats))] in - ssrtac_expr ?loc "tclintros" args - -GEXTEND Gram - GLOBAL: tactic_expr; - tactic_expr: LEVEL "1" [ RIGHTA - [ tac = tactic_expr; intros = ssrintros_ne -> tclintros_expr ~loc:!@loc tac intros - ] ]; -END - -(* }}} *) - - -(** Bracketing tactical *) - -(* The tactic pretty-printer doesn't know that some extended tactics *) -(* are actually tacticals. To prevent it from improperly removing *) -(* parentheses we override the parsing rule for bracketed tactic *) -(* expressions so that the pretty-print always reflects the input. *) -(* (Removing user-specified parentheses is dubious anyway). *) - -GEXTEND Gram - GLOBAL: tactic_expr; - ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> Loc.tag ~loc:!@loc (Tacexp tac) ]]; - tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> TacArg arg ]]; -END - -(** The internal "done" and "ssrautoprop" tactics. *) - -(* For additional flexibility, "done" and "ssrautoprop" are *) -(* defined in Ltac. *) -(* Although we provide a default definition in ssreflect, *) -(* we look up the definition dynamically at each call point, *) -(* to allow for user extensions. "ssrautoprop" defaults to *) -(* trivial. *) - -let ssrautoprop gl = - try - let tacname = - try Tacenv.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop")) - with Not_found -> Tacenv.locate_tactic (ssrqid "ssrautoprop") in - let tacexpr = Loc.tag @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in - V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl - with Not_found -> V82.of_tactic (Auto.full_trivial []) gl - -let () = ssrautoprop_tac := ssrautoprop - -let tclBY tac = Tacticals.tclTHEN tac (donetac ~-1) - -(** Tactical arguments. *) - -(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) -(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) -(* and subgoal reordering tacticals (; first & ; last), respectively. *) - -(* Force use of the tactic_expr parsing entry, to rule out tick marks. *) - -(** The "by" tactical. *) - - -open Ssrfwd - -TACTIC EXTEND ssrtclby -| [ "by" ssrhintarg(tac) ] -> [ V82.tactic (hinttac ist true tac) ] -END - -(* }}} *) -(* We can't parse "by" in ARGUMENT EXTEND because it will only be made *) -(* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *) - -GEXTEND Gram - GLOBAL: ssrhint simple_tactic; - ssrhint: [[ "by"; arg = ssrhintarg -> arg ]]; -END - -(** The "do" tactical. ********************************************************) - -(* -type ssrdoarg = ((ssrindex * ssrmmod) * ssrhint) * ssrclauses -*) -TACTIC EXTEND ssrtcldo -| [ "YouShouldNotTypeThis" "do" ssrdoarg(arg) ] -> [ V82.tactic (ssrdotac ist arg) ] -END -set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"] - -let ssrdotac_expr ?loc n m tac clauses = - let arg = ((n, m), tac), clauses in - ssrtac_expr ?loc "tcldo" [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrdoarg) arg)] - -GEXTEND Gram - GLOBAL: tactic_expr; - ssrdotac: [ - [ tac = tactic_expr LEVEL "3" -> mk_hint tac - | tacs = ssrortacarg -> tacs - ] ]; - tactic_expr: LEVEL "3" [ RIGHTA - [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> - ssrdotac_expr ~loc:!@loc noindex m tac clauses - | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses -> - ssrdotac_expr ~loc:!@loc noindex Once tac clauses - | IDENT "do"; n = int_or_var; m = ssrmmod; - tac = ssrdotac; clauses = ssrclauses -> - ssrdotac_expr ~loc:!@loc (mk_index ~loc:!@loc n) m tac clauses - ] ]; -END -(* }}} *) - - -(* We can't actually parse the direction separately because this *) -(* would introduce conflicts with the basic ltac syntax. *) -let pr_ssrseqdir _ _ _ = function - | L2R -> str ";" ++ spc () ++ str "first " - | R2L -> str ";" ++ spc () ++ str "last " - -ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY pr_ssrseqdir -| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -TACTIC EXTEND ssrtclseq -| [ "YouShouldNotTypeThis" ssrtclarg(tac) ssrseqdir(dir) ssrseqarg(arg) ] -> - [ V82.tactic (tclSEQAT ist tac dir arg) ] -END -set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"] - -let tclseq_expr ?loc tac dir arg = - let arg1 = in_gen (rawwit wit_ssrtclarg) tac in - let arg2 = in_gen (rawwit wit_ssrseqdir) dir in - let arg3 = in_gen (rawwit wit_ssrseqarg) (check_seqtacarg dir arg) in - ssrtac_expr ?loc "tclseq" (List.map (fun x -> Tacexpr.TacGeneric x) [arg1; arg2; arg3]) - -GEXTEND Gram - GLOBAL: tactic_expr; - ssr_first: [ - [ tac = ssr_first; ipats = ssrintros_ne -> tclintros_expr ~loc:!@loc tac ipats - | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> TacFirst tacl - ] ]; - ssr_first_else: [ - [ tac1 = ssr_first; tac2 = ssrorelse -> TacOrelse (tac1, tac2) - | tac = ssr_first -> tac ]]; - tactic_expr: LEVEL "4" [ LEFTA - [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else -> - TacThen (tac1, tac2) - | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg -> - tclseq_expr ~loc:!@loc tac L2R arg - | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg -> - tclseq_expr ~loc:!@loc tac R2L arg - ] ]; -END -(* }}} *) - -(** 5. Bookkeeping tactics (clear, move, case, elim) *) - -(** Generalization (discharge) item *) - -(* An item is a switch + term pair. *) - -(* type ssrgen = ssrdocc * ssrterm *) - -let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt - -let pr_ssrgen _ _ _ = pr_gen - -ARGUMENT EXTEND ssrgen TYPED AS ssrdocc * cpattern PRINTED BY pr_ssrgen -| [ ssrdocc(docc) cpattern(dt) ] -> [ - match docc with - | Some [], _ -> CErrors.user_err ~loc (str"Clear flag {} not allowed here") - | _ -> docc, dt ] -| [ cpattern(dt) ] -> [ nodocc, dt ] -END - -let has_occ ((_, occ), _) = occ <> None - -(** Generalization (discharge) sequence *) - -(* A discharge sequence is represented as a list of up to two *) -(* lists of d-items, plus an ident list set (the possibly empty *) -(* final clear switch). The main list is empty iff the command *) -(* is defective, and has length two if there is a sequence of *) -(* dependent terms (and in that case it is the first of the two *) -(* lists). Thus, the first of the two lists is never empty. *) - -(* type ssrgens = ssrgen list *) -(* type ssrdgens = ssrgens list * ssrclear *) - -let gens_sep = function [], [] -> mt | _ -> spc - -let pr_dgens pr_gen (gensl, clr) = - let prgens s gens = str s ++ pr_list spc pr_gen gens in - let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in - match gensl with - | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr - | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr - | [gens] -> prgens ": " gens ++ pr_clear spc clr - | _ -> pr_clear pr_spc clr - -let pr_ssrdgens _ _ _ = pr_dgens pr_gen - -let cons_gen gen = function - | gens :: gensl, clr -> (gen :: gens) :: gensl, clr - | _ -> anomaly "missing gen list" - -let cons_dep (gensl, clr) = - if List.length gensl = 1 then ([] :: gensl, clr) else - CErrors.user_err (Pp.str "multiple dependents switches '/'") - -ARGUMENT EXTEND ssrdgens_tl TYPED AS ssrgen list list * ssrclear - PRINTED BY pr_ssrdgens -| [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> - [ cons_gen (mkclr clr, dt) dgens ] -| [ "{" ne_ssrhyp_list(clr) "}" ] -> - [ [[]], clr ] -| [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> - [ cons_gen (mkocc occ, dt) dgens ] -| [ "/" ssrdgens_tl(dgens) ] -> - [ cons_dep dgens ] -| [ cpattern(dt) ssrdgens_tl(dgens) ] -> - [ cons_gen (nodocc, dt) dgens ] -| [ ] -> - [ [[]], [] ] -END - -ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY pr_ssrdgens -| [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> [ cons_gen gen dgens ] -END - -(** Equations *) - -(* argument *) - -let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt () -let pr_ssreqid _ _ _ = pr_eqid - -(* We must use primitive parsing here to avoid conflicts with the *) -(* basic move, case, and elim tactics. *) -ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY pr_ssreqid -| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -let accept_ssreqid strm = - match Util.stream_nth 0 strm with - | Tok.IDENT _ -> accept_before_syms [":"] strm - | Tok.KEYWORD ":" -> () - | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] -> - accept_before_syms [":"] strm - | _ -> raise Stream.Failure - -let test_ssreqid = Gram.Entry.of_parser "test_ssreqid" accept_ssreqid - -GEXTEND Gram - GLOBAL: ssreqid; - ssreqpat: [ - [ id = Prim.ident -> IPatId id - | "_" -> IPatAnon Drop - | "?" -> IPatAnon One - | occ = ssrdocc; "->" -> (match occ with - | None, occ -> IPatRewrite (occ, L2R) - | _ -> CErrors.user_err ~loc:!@loc (str"Only occurrences are allowed here")) - | occ = ssrdocc; "<-" -> (match occ with - | None, occ -> IPatRewrite (occ, R2L) - | _ -> CErrors.user_err ~loc:!@loc (str "Only occurrences are allowed here")) - | "->" -> IPatRewrite (allocc, L2R) - | "<-" -> IPatRewrite (allocc, R2L) - ]]; - ssreqid: [ - [ test_ssreqid; pat = ssreqpat -> Some pat - | test_ssreqid -> None - ]]; -END - -(** Bookkeeping (discharge-intro) argument *) - -(* Since all bookkeeping ssr commands have the same discharge-intro *) -(* argument format we use a single grammar entry point to parse them. *) -(* the entry point parses only non-empty arguments to avoid conflicts *) -(* with the basic Coq tactics. *) - -(* type ssrarg = ssrbwdview * (ssreqid * (ssrdgens * ssripats)) *) - -let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) = - let pri = pr_intros (gens_sep dgens) in - pr_view2 view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats - -ARGUMENT EXTEND ssrarg TYPED AS ssrfwdview * (ssreqid * (ssrdgens * ssrintros)) - PRINTED BY pr_ssrarg -| [ ssrfwdview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> - [ view, (eqid, (dgens, ipats)) ] -| [ ssrfwdview(view) ssrclear(clr) ssrintros(ipats) ] -> - [ view, (None, (([], clr), ipats)) ] -| [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> - [ [], (eqid, (dgens, ipats)) ] -| [ ssrclear_ne(clr) ssrintros(ipats) ] -> - [ [], (None, (([], clr), ipats)) ] -| [ ssrintros_ne(ipats) ] -> - [ [], (None, (([], []), ipats)) ] -END - -(** The "clear" tactic *) - -(* We just add a numeric version that clears the n top assumptions. *) - -TACTIC EXTEND ssrclear - | [ "clear" natural(n) ] -> [ tclIPAT (List.init n (fun _ -> IPatAnon Drop)) ] -END - -(** The "move" tactic *) - -(* TODO: review this, in particular the => _ and => [] cases *) -let rec improper_intros = function - | IPatSimpl _ :: ipats -> improper_intros ipats - | (IPatId _ | IPatAnon _ | IPatCase _ | IPatDispatch _) :: _ -> false - | _ -> true (* FIXME *) - -let check_movearg = function - | view, (eqid, _) when view <> [] && eqid <> None -> - CErrors.user_err (Pp.str "incompatible view and equation in move tactic") - | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen -> - CErrors.user_err (Pp.str "incompatible view and occurrence switch in move tactic") - | _, (_, ((dgens, _), _)) when List.length dgens > 1 -> - CErrors.user_err (Pp.str "dependents switch `/' in move tactic") - | _, (eqid, (_, ipats)) when eqid <> None && improper_intros ipats -> - CErrors.user_err (Pp.str "no proper intro pattern for equation in move tactic") - | arg -> arg - -ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY pr_ssrarg -| [ ssrarg(arg) ] -> [ check_movearg arg ] -END - -let movearg_of_parsed_movearg (v,(eq,(dg,ip))) = - (v,(eq,(ssrdgens_of_parsed_dgens dg,ip))) - -TACTIC EXTEND ssrmove -| [ "move" ssrmovearg(arg) ssrrpat(pat) ] -> - [ ssrmovetac (movearg_of_parsed_movearg arg) <*> tclIPAT [pat] ] -| [ "move" ssrmovearg(arg) ssrclauses(clauses) ] -> - [ tclCLAUSES (ssrmovetac (movearg_of_parsed_movearg arg)) clauses ] -| [ "move" ssrrpat(pat) ] -> [ tclIPAT [pat] ] -| [ "move" ] -> [ ssrsmovetac ] -END - -let check_casearg = function -| view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen -> - CErrors.user_err (Pp.str "incompatible view and occurrence switch in dependent case tactic") -| arg -> arg - -ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY pr_ssrarg -| [ ssrarg(arg) ] -> [ check_casearg arg ] -END - -TACTIC EXTEND ssrcase -| [ "case" ssrcasearg(arg) ssrclauses(clauses) ] -> - [ tclCLAUSES (ssrcasetac (movearg_of_parsed_movearg arg)) clauses ] -| [ "case" ] -> [ ssrscasetoptac ] -END - -(** The "elim" tactic *) - -TACTIC EXTEND ssrelim -| [ "elim" ssrarg(arg) ssrclauses(clauses) ] -> - [ tclCLAUSES (ssrelimtac (movearg_of_parsed_movearg arg)) clauses ] -| [ "elim" ] -> [ ssrselimtoptac ] -END - -(** 6. Backward chaining tactics: apply, exact, congr. *) - -(** The "apply" tactic *) - -let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt -let pr_ssragen _ _ _ = pr_agen -let pr_ssragens _ _ _ = pr_dgens pr_agen - -ARGUMENT EXTEND ssragen TYPED AS ssrdocc * ssrterm PRINTED BY pr_ssragen -| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> [ mkclr clr, dt ] -| [ ssrterm(dt) ] -> [ nodocc, dt ] -END - -ARGUMENT EXTEND ssragens TYPED AS ssragen list list * ssrclear -PRINTED BY pr_ssragens -| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] -> - [ cons_gen (mkclr clr, dt) agens ] -| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ [[]], clr] -| [ ssrterm(dt) ssragens(agens) ] -> - [ cons_gen (nodocc, dt) agens ] -| [ ] -> [ [[]], [] ] -END - -let mk_applyarg views agens intros = views, (agens, intros) - -let pr_ssraarg _ _ _ (view, (dgens, ipats)) = - let pri = pr_intros (gens_sep dgens) in - pr_view view ++ pr_dgens pr_agen dgens ++ pri ipats - -ARGUMENT EXTEND ssrapplyarg -TYPED AS ssrbwdview * (ssragens * ssrintros) -PRINTED BY pr_ssraarg -| [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> - [ mk_applyarg [] (cons_gen gen dgens) intros ] -| [ ssrclear_ne(clr) ssrintros(intros) ] -> - [ mk_applyarg [] ([], clr) intros ] -| [ ssrintros_ne(intros) ] -> - [ mk_applyarg [] ([], []) intros ] -| [ ssrbwdview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> - [ mk_applyarg view (cons_gen gen dgens) intros ] -| [ ssrbwdview(view) ssrclear(clr) ssrintros(intros) ] -> - [ mk_applyarg view ([], clr) intros ] - END - -TACTIC EXTEND ssrapply -| [ "apply" ssrapplyarg(arg) ] -> [ - let views, (gens_clr, intros) = arg in - inner_ssrapplytac views gens_clr ist <*> tclIPATssr intros ] -| [ "apply" ] -> [ apply_top_tac ] -END - -(** The "exact" tactic *) - -let mk_exactarg views dgens = mk_applyarg views dgens [] - -ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY pr_ssraarg -| [ ":" ssragen(gen) ssragens(dgens) ] -> - [ mk_exactarg [] (cons_gen gen dgens) ] -| [ ssrbwdview(view) ssrclear(clr) ] -> - [ mk_exactarg view ([], clr) ] -| [ ssrclear_ne(clr) ] -> - [ mk_exactarg [] ([], clr) ] -END - -let vmexacttac pf = - Goal.enter begin fun gl -> - exact_no_check (EConstr.mkCast (pf, _vmcast, Tacmach.New.pf_concl gl)) - end - -TACTIC EXTEND ssrexact -| [ "exact" ssrexactarg(arg) ] -> [ - let views, (gens_clr, _) = arg in - V82.tactic (tclBY (V82.of_tactic (inner_ssrapplytac views gens_clr ist))) ] -| [ "exact" ] -> [ - V82.tactic (Tacticals.tclORELSE (donetac ~-1) (tclBY (V82.of_tactic apply_top_tac))) ] -| [ "exact" "<:" lconstr(pf) ] -> [ vmexacttac pf ] -END - -(** The "congr" tactic *) - -(* type ssrcongrarg = open_constr * (int * constr) *) - -let pr_ssrcongrarg _ _ _ ((n, f), dgens) = - (if n <= 0 then mt () else str " " ++ int n) ++ - str " " ++ pr_term f ++ pr_dgens pr_gen dgens - -ARGUMENT EXTEND ssrcongrarg TYPED AS (int * ssrterm) * ssrdgens - PRINTED BY pr_ssrcongrarg -| [ natural(n) constr(c) ssrdgens(dgens) ] -> [ (n, mk_term xNoFlag c), dgens ] -| [ natural(n) constr(c) ] -> [ (n, mk_term xNoFlag c),([[]],[]) ] -| [ constr(c) ssrdgens(dgens) ] -> [ (0, mk_term xNoFlag c), dgens ] -| [ constr(c) ] -> [ (0, mk_term xNoFlag c), ([[]],[]) ] -END - - - -TACTIC EXTEND ssrcongr -| [ "congr" ssrcongrarg(arg) ] -> -[ let arg, dgens = arg in - V82.tactic begin - match dgens with - | [gens], clr -> Tacticals.tclTHEN (genstac (gens,clr)) (newssrcongrtac arg ist) - | _ -> errorstrm (str"Dependent family abstractions not allowed in congr") - end] -END - -(** 7. Rewriting tactics (rewrite, unlock) *) - -(** Coq rewrite compatibility flag *) - -(** Rewrite clear/occ switches *) - -let pr_rwocc = function - | None, None -> mt () - | None, occ -> pr_occ occ - | Some clr, _ -> pr_clear_ne clr - -let pr_ssrrwocc _ _ _ = pr_rwocc - -ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY pr_ssrrwocc -| [ "{" ssrhyp_list(clr) "}" ] -> [ mkclr clr ] -| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] -| [ ] -> [ noclr ] -END - -(** Rewrite rules *) - -let pr_rwkind = function - | RWred s -> pr_simpl s - | RWdef -> str "/" - | RWeq -> mt () - -let wit_ssrrwkind = add_genarg "ssrrwkind" pr_rwkind - -let pr_rule = function - | RWred s, _ -> pr_simpl s - | RWdef, r-> str "/" ++ pr_term r - | RWeq, r -> pr_term r - -let pr_ssrrule _ _ _ = pr_rule - -let noruleterm loc = mk_term xNoFlag (mkCProp loc) - -ARGUMENT EXTEND ssrrule_ne TYPED AS ssrrwkind * ssrterm PRINTED BY pr_ssrrule - | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -GEXTEND Gram - GLOBAL: ssrrule_ne; - ssrrule_ne : [ - [ test_not_ssrslashnum; x = - [ "/"; t = ssrterm -> RWdef, t - | t = ssrterm -> RWeq, t - | s = ssrsimpl_ne -> RWred s, noruleterm (Some !@loc) - ] -> x - | s = ssrsimpl_ne -> RWred s, noruleterm (Some !@loc) - ]]; -END - -ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY pr_ssrrule - | [ ssrrule_ne(r) ] -> [ r ] - | [ ] -> [ RWred Nop, noruleterm (Some loc) ] -END - -(** Rewrite arguments *) - -let pr_option f = function None -> mt() | Some x -> f x -let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]") -let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep -let pr_rwarg ((d, m), ((docc, rx), r)) = - pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r - -let pr_ssrrwarg _ _ _ = pr_rwarg - -ARGUMENT EXTEND ssrpattern_squarep -TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep - | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] - | [ ] -> [ None ] -END - -ARGUMENT EXTEND ssrpattern_ne_squarep -TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep - | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] -END - - -ARGUMENT EXTEND ssrrwarg - TYPED AS (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule) - PRINTED BY pr_ssrrwarg - | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg (R2L, m) (docc, rx) r ] - | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *) - [ mk_rwarg (R2L, nomult) norwocc (RWdef, t) ] - | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg (L2R, m) (docc, rx) r ] - | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg norwmult (mkclr clr, rx) r ] - | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] -> - [ mk_rwarg norwmult (mkclr clr, None) r ] - | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg norwmult (mkocc occ, rx) r ] - | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg norwmult (nodocc, rx) r ] - | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> - [ mk_rwarg norwmult (noclr, rx) r ] - | [ ssrrule_ne(r) ] -> - [ mk_rwarg norwmult norwocc r ] -END - -TACTIC EXTEND ssrinstofruleL2R -| [ "ssrinstancesofruleL2R" ssrterm(arg) ] -> [ V82.tactic (ssrinstancesofrule ist L2R arg) ] -END -TACTIC EXTEND ssrinstofruleR2L -| [ "ssrinstancesofruleR2L" ssrterm(arg) ] -> [ V82.tactic (ssrinstancesofrule ist R2L arg) ] -END - -(** Rewrite argument sequence *) - -(* type ssrrwargs = ssrrwarg list *) - -let pr_ssrrwargs _ _ _ rwargs = pr_list spc pr_rwarg rwargs - -ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list PRINTED BY pr_ssrrwargs - | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] -END - -let ssr_rw_syntax = Summary.ref ~name:"SSR:rewrite" true - -let _ = - Goptions.declare_bool_option - { Goptions.optname = "ssreflect rewrite"; - Goptions.optkey = ["SsrRewrite"]; - Goptions.optread = (fun _ -> !ssr_rw_syntax); - Goptions.optdepr = false; - Goptions.optwrite = (fun b -> ssr_rw_syntax := b) } - -let test_ssr_rw_syntax = - let test strm = - if not !ssr_rw_syntax then raise Stream.Failure else - if is_ssr_loaded () then () else - match Util.stream_nth 0 strm with - | Tok.KEYWORD key when List.mem key.[0] ['{'; '['; '/'] -> () - | _ -> raise Stream.Failure in - Gram.Entry.of_parser "test_ssr_rw_syntax" test - -GEXTEND Gram - GLOBAL: ssrrwargs; - ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> a ]]; -END - -(** The "rewrite" tactic *) - -TACTIC EXTEND ssrrewrite - | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] -> - [ tclCLAUSES (old_tac (ssrrewritetac ist args)) clauses ] -END - -(** The "unlock" tactic *) - -let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t -let pr_ssrunlockarg _ _ _ = pr_unlockarg - -ARGUMENT EXTEND ssrunlockarg TYPED AS ssrocc * ssrterm - PRINTED BY pr_ssrunlockarg - | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> [ occ, t ] - | [ ssrterm(t) ] -> [ None, t ] -END - -let pr_ssrunlockargs _ _ _ args = pr_list spc pr_unlockarg args - -ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list - PRINTED BY pr_ssrunlockargs - | [ ssrunlockarg_list(args) ] -> [ args ] -END - -TACTIC EXTEND ssrunlock - | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] -> - [ tclCLAUSES (old_tac (unlocktac ist args)) clauses ] -END - -(** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *) - - -TACTIC EXTEND ssrpose -| [ "pose" ssrfixfwd(ffwd) ] -> [ V82.tactic (ssrposetac ffwd) ] -| [ "pose" ssrcofixfwd(ffwd) ] -> [ V82.tactic (ssrposetac ffwd) ] -| [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> [ V82.tactic (ssrposetac (id, fwd)) ] -END - -(** The "set" tactic *) - -(* type ssrsetfwd = ssrfwd * ssrdocc *) - -TACTIC EXTEND ssrset -| [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] -> - [ tclCLAUSES (old_tac (ssrsettac id fwd)) clauses ] -END - -(** The "have" tactic *) - -(* type ssrhavefwd = ssrfwd * ssrhint *) - - -(* Pltac. *) - -(* The standard TACTIC EXTEND does not work for abstract *) -GEXTEND Gram - GLOBAL: tactic_expr; - tactic_expr: LEVEL "3" - [ RIGHTA [ IDENT "abstract"; gens = ssrdgens -> - ssrtac_expr ~loc:!@loc "abstract" - [Tacexpr.TacGeneric (Genarg.in_gen (Genarg.rawwit wit_ssrdgens) gens)] ]]; -END -TACTIC EXTEND ssrabstract -| [ "abstract" ssrdgens(gens) ] -> [ - if List.length (fst gens) <> 1 then - errorstrm (str"dependents switches '/' not allowed here"); - Ssripats.ssrabstract (ssrdgens_of_parsed_dgens gens) ] -END - -TACTIC EXTEND ssrhave -| [ "have" ssrhavefwdwbinders(fwd) ] -> - [ V82.tactic (havetac ist fwd false false) ] -END - -TACTIC EXTEND ssrhavesuff -| [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> - [ V82.tactic (havetac ist (false,(pats,fwd)) true false) ] -END - -TACTIC EXTEND ssrhavesuffices -| [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> - [ V82.tactic (havetac ist (false,(pats,fwd)) true false) ] -END - -TACTIC EXTEND ssrsuffhave -| [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> - [ V82.tactic (havetac ist (false,(pats,fwd)) true true) ] -END - -TACTIC EXTEND ssrsufficeshave -| [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> - [ V82.tactic (havetac ist (false,(pats,fwd)) true true) ] -END - -(** The "suffice" tactic *) - -let pr_ssrsufffwdwbinders _ _ prt (hpats, (fwd, hint)) = - pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint - -ARGUMENT EXTEND ssrsufffwd - TYPED AS ssrhpats * (ssrfwd * ssrhint) PRINTED BY pr_ssrsufffwdwbinders -| [ ssrhpats(pats) ssrbinder_list(bs) ":" ast_closure_lterm(t) ssrhint(hint) ] -> - [ let ((clr, pats), binders), simpl = pats in - let allbs = intro_id_to_binder binders @ bs in - let allbinders = binders @ List.flatten (binder_to_intro_id bs) in - let fwd = mkFwdHint ":" t in - (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) ] -END - - -TACTIC EXTEND ssrsuff -| [ "suff" ssrsufffwd(fwd) ] -> [ V82.tactic (sufftac ist fwd) ] -END - -TACTIC EXTEND ssrsuffices -| [ "suffices" ssrsufffwd(fwd) ] -> [ V82.tactic (sufftac ist fwd) ] -END - -(** The "wlog" (Without Loss Of Generality) tactic *) - -(* type ssrwlogfwd = ssrwgen list * ssrfwd *) - -let pr_ssrwlogfwd _ _ _ (gens, t) = - str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t - -ARGUMENT EXTEND ssrwlogfwd TYPED AS ssrwgen list * ssrfwd - PRINTED BY pr_ssrwlogfwd -| [ ":" ssrwgen_list(gens) "/" ast_closure_lterm(t) ] -> [ gens, mkFwdHint "/" t] -END - - -TACTIC EXTEND ssrwlog -| [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ V82.tactic (wlogtac ist pats fwd hint false `NoGen) ] -END - -TACTIC EXTEND ssrwlogs -| [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] -END - -TACTIC EXTEND ssrwlogss -| [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> - [ V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] -END - -TACTIC EXTEND ssrwithoutloss -| [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ V82.tactic (wlogtac ist pats fwd hint false `NoGen) ] -END - -TACTIC EXTEND ssrwithoutlosss -| [ "without" "loss" "suff" - ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] -END - -TACTIC EXTEND ssrwithoutlossss -| [ "without" "loss" "suffices" - ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> - [ V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] -END - -(* Generally have *) -let pr_idcomma _ _ _ = function - | None -> mt() - | Some None -> str"_, " - | Some (Some id) -> pr_id id ++ str", " - -ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY pr_idcomma - | [ ] -> [ None ] -END - -let accept_idcomma strm = - match stream_nth 0 strm with - | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm - | _ -> raise Stream.Failure - -let test_idcomma = Gram.Entry.of_parser "test_idcomma" accept_idcomma - -GEXTEND Gram - GLOBAL: ssr_idcomma; - ssr_idcomma: [ [ test_idcomma; - ip = [ id = IDENT -> Some (Id.of_string id) | "_" -> None ]; "," -> - Some ip - ] ]; -END - -let augment_preclr clr1 (((clr0, x),y),z) = (((clr1 @ clr0, x),y),z) - -TACTIC EXTEND ssrgenhave -| [ "gen" "have" ssrclear(clr) - ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ let pats = augment_preclr clr pats in - V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ] -END - -TACTIC EXTEND ssrgenhave2 -| [ "generally" "have" ssrclear(clr) - ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> - [ let pats = augment_preclr clr pats in - V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ] -END - -(* We wipe out all the keywords generated by the grammar rules we defined. *) -(* The user is supposed to Require Import ssreflect or Require ssreflect *) -(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) -(* consequence the extended ssreflect grammar. *) -let () = CLexer.set_keyword_state frozen_lexer ;; - - -(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/plugins/ssr/ssrparser.mlg b/plugins/ssr/ssrparser.mlg new file mode 100644 index 0000000000..8699b62c39 --- /dev/null +++ b/plugins/ssr/ssrparser.mlg @@ -0,0 +1,2628 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* { CErrors.anomaly (Pp.str "Grammar placeholder match") } +END +GRAMMAR EXTEND Gram + GLOBAL: ssrtacarg; + ssrtacarg: [[ tac = tactic_expr LEVEL "5" -> { tac } ]]; +END + +{ + +(* Lexically closed tactic for tacticals. *) +let pr_ssrtclarg _ _ prt tac = prt tacltop tac + +} + +ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg + PRINTED BY { pr_ssrtclarg } +| [ ssrtacarg(tac) ] -> { tac } +END + +{ + +open Genarg + +(** Adding a new uninterpreted generic argument type *) +let add_genarg tag pr = + let wit = Genarg.make0 tag in + let tag = Geninterp.Val.create tag in + let glob ist x = (ist, x) in + let subst _ x = x in + let interp ist x = Ftactic.return (Geninterp.Val.Dyn (tag, x)) in + let gen_pr _ _ _ = pr in + let () = Genintern.register_intern0 wit glob in + let () = Genintern.register_subst0 wit subst in + let () = Geninterp.register_interp0 wit interp in + let () = Geninterp.register_val0 wit (Some (Geninterp.Val.Base tag)) in + Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; + wit + +(** Primitive parsing to avoid syntax conflicts with basic tactics. *) + +let accept_before_syms syms strm = + match Util.stream_nth 1 strm with + | Tok.KEYWORD sym when List.mem sym syms -> () + | _ -> raise Stream.Failure + +let accept_before_syms_or_any_id syms strm = + match Util.stream_nth 1 strm with + | Tok.KEYWORD sym when List.mem sym syms -> () + | Tok.IDENT _ -> () + | _ -> raise Stream.Failure + +let accept_before_syms_or_ids syms ids strm = + match Util.stream_nth 1 strm with + | Tok.KEYWORD sym when List.mem sym syms -> () + | Tok.IDENT id when List.mem id ids -> () + | _ -> raise Stream.Failure + +open Ssrast +let pr_id = Ppconstr.pr_id +let pr_name = function Name id -> pr_id id | Anonymous -> str "_" +let pr_spc () = str " " +let pr_list = prlist_with_sep + +(**************************** ssrhyp **************************************) + +let pr_ssrhyp _ _ _ = pr_hyp + +let wit_ssrhyprep = add_genarg "ssrhyprep" pr_hyp + +let intern_hyp ist (SsrHyp (loc, id) as hyp) = + let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_var) CAst.(make ?loc id)) in + if not_section_id id then hyp else + hyp_err ?loc "Can't clear section hypothesis " id + +open Pcoq.Prim + +} + +ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY { pr_ssrhyp } + INTERPRETED BY { interp_hyp } + GLOBALIZED BY { intern_hyp } + | [ ident(id) ] -> { SsrHyp (Loc.tag ~loc id) } +END + +{ + +let pr_hoi = hoik pr_hyp +let pr_ssrhoi _ _ _ = pr_hoi + +let wit_ssrhoirep = add_genarg "ssrhoirep" pr_hoi + +let intern_ssrhoi ist = function + | Hyp h -> Hyp (intern_hyp ist h) + | Id (SsrHyp (_, id)) as hyp -> + let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_ident) id) in + hyp + +let interp_ssrhoi ist gl = function + | Hyp h -> let s, h' = interp_hyp ist gl h in s, Hyp h' + | Id (SsrHyp (loc, id)) -> + let s, id' = interp_wit wit_ident ist gl id in + s, Id (SsrHyp (loc, id')) + +} + +ARGUMENT EXTEND ssrhoi_hyp TYPED AS ssrhoirep PRINTED BY { pr_ssrhoi } + INTERPRETED BY { interp_ssrhoi } + GLOBALIZED BY { intern_ssrhoi } + | [ ident(id) ] -> { Hyp (SsrHyp(Loc.tag ~loc id)) } +END +ARGUMENT EXTEND ssrhoi_id TYPED AS ssrhoirep PRINTED BY { pr_ssrhoi } + INTERPRETED BY { interp_ssrhoi } + GLOBALIZED BY { intern_ssrhoi } + | [ ident(id) ] -> { Id (SsrHyp(Loc.tag ~loc id)) } +END + +{ + +let pr_ssrhyps _ _ _ = pr_hyps + +} + +ARGUMENT EXTEND ssrhyps TYPED AS ssrhyp list PRINTED BY { pr_ssrhyps } + INTERPRETED BY { interp_hyps } + | [ ssrhyp_list(hyps) ] -> { check_hyps_uniq [] hyps; hyps } +END + +(** Rewriting direction *) + +{ + +let pr_rwdir = function L2R -> mt() | R2L -> str "-" + +let wit_ssrdir = add_genarg "ssrdir" pr_dir + +(** Simpl switch *) + +let pr_ssrsimpl _ _ _ = pr_simpl + +let wit_ssrsimplrep = add_genarg "ssrsimplrep" pr_simpl + +let test_ssrslashnum b1 b2 strm = + match Util.stream_nth 0 strm with + | Tok.KEYWORD "/" -> + (match Util.stream_nth 1 strm with + | Tok.INT _ when b1 -> + (match Util.stream_nth 2 strm with + | Tok.KEYWORD "=" | Tok.KEYWORD "/=" when not b2 -> () + | Tok.KEYWORD "/" -> + if not b2 then () else begin + match Util.stream_nth 3 strm with + | Tok.INT _ -> () + | _ -> raise Stream.Failure + end + | _ -> raise Stream.Failure) + | Tok.KEYWORD "/" when not b1 -> + (match Util.stream_nth 2 strm with + | Tok.KEYWORD "=" when not b2 -> () + | Tok.INT _ when b2 -> + (match Util.stream_nth 3 strm with + | Tok.KEYWORD "=" -> () + | _ -> raise Stream.Failure) + | _ when not b2 -> () + | _ -> raise Stream.Failure) + | Tok.KEYWORD "=" when not b1 && not b2 -> () + | _ -> raise Stream.Failure) + | Tok.KEYWORD "//" when not b1 -> + (match Util.stream_nth 1 strm with + | Tok.KEYWORD "=" when not b2 -> () + | Tok.INT _ when b2 -> + (match Util.stream_nth 2 strm with + | Tok.KEYWORD "=" -> () + | _ -> raise Stream.Failure) + | _ when not b2 -> () + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure + +let test_ssrslashnum10 = test_ssrslashnum true false +let test_ssrslashnum11 = test_ssrslashnum true true +let test_ssrslashnum01 = test_ssrslashnum false true +let test_ssrslashnum00 = test_ssrslashnum false false + +let negate_parser f x = + let rc = try Some (f x) with Stream.Failure -> None in + match rc with + | None -> () + | Some _ -> raise Stream.Failure + +let test_not_ssrslashnum = + Pcoq.Gram.Entry.of_parser + "test_not_ssrslashnum" (negate_parser test_ssrslashnum10) +let test_ssrslashnum00 = + Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum00 +let test_ssrslashnum10 = + Pcoq.Gram.Entry.of_parser "test_ssrslashnum10" test_ssrslashnum10 +let test_ssrslashnum11 = + Pcoq.Gram.Entry.of_parser "test_ssrslashnum11" test_ssrslashnum11 +let test_ssrslashnum01 = + Pcoq.Gram.Entry.of_parser "test_ssrslashnum01" test_ssrslashnum01 + +} + +ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY { pr_ssrsimpl } +| [ "//=" ] -> { SimplCut (~-1,~-1) } +| [ "/=" ] -> { Simpl ~-1 } +END + +(* Pcoq.Prim. *) +GRAMMAR EXTEND Gram + GLOBAL: ssrsimpl_ne; + ssrsimpl_ne: [ + [ test_ssrslashnum11; "/"; n = natural; "/"; m = natural; "=" -> { SimplCut(n,m) } + | test_ssrslashnum10; "/"; n = natural; "/" -> { Cut n } + | test_ssrslashnum10; "/"; n = natural; "=" -> { Simpl n } + | test_ssrslashnum10; "/"; n = natural; "/=" -> { SimplCut (n,~-1) } + | test_ssrslashnum10; "/"; n = natural; "/"; "=" -> { SimplCut (n,~-1) } + | test_ssrslashnum01; "//"; m = natural; "=" -> { SimplCut (~-1,m) } + | test_ssrslashnum00; "//" -> { Cut ~-1 } + ]]; + +END + +ARGUMENT EXTEND ssrsimpl TYPED AS ssrsimplrep PRINTED BY { pr_ssrsimpl } +| [ ssrsimpl_ne(sim) ] -> { sim } +| [ ] -> { Nop } +END + +{ + +let pr_ssrclear _ _ _ = pr_clear mt + +} + +ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyps PRINTED BY { pr_ssrclear } +| [ "{" ne_ssrhyp_list(clr) "}" ] -> { check_hyps_uniq [] clr; clr } +END + +ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY { pr_ssrclear } +| [ ssrclear_ne(clr) ] -> { clr } +| [ ] -> { [] } +END + +(** Indexes *) + +(* Since SSR indexes are always positive numbers, we use the 0 value *) +(* to encode an omitted index. We reuse the in or_var type, but we *) +(* supply our own interpretation function, which checks for non *) +(* positive values, and allows the use of constr numerals, so that *) +(* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *) + +{ + +let pr_index = function + | ArgVar {CAst.v=id} -> pr_id id + | ArgArg n when n > 0 -> int n + | _ -> mt () +let pr_ssrindex _ _ _ = pr_index + +let noindex = ArgArg 0 + +let check_index ?loc i = + if i > 0 then i else CErrors.user_err ?loc (str"Index not positive") +let mk_index ?loc = function + | ArgArg i -> ArgArg (check_index ?loc i) + | iv -> iv + +let interp_index ist gl idx = + Tacmach.project gl, + match idx with + | ArgArg _ -> idx + | ArgVar id -> + let i = + try + let v = Id.Map.find id.CAst.v ist.Tacinterp.lfun in + begin match Tacinterp.Value.to_int v with + | Some i -> i + | None -> + begin match Tacinterp.Value.to_constr v with + | Some c -> + let rc = Detyping.detype Detyping.Now false Id.Set.empty (pf_env gl) (project gl) c in + begin match Notation.uninterp_prim_token rc with + | _, Constrexpr.Numeral (s,b) -> + let n = int_of_string s in if b then n else -n + | _ -> raise Not_found + end + | None -> raise Not_found + end end + with _ -> CErrors.user_err ?loc:id.CAst.loc (str"Index not a number") in + ArgArg (check_index ?loc:id.CAst.loc i) + +open Pltac + +} + +ARGUMENT EXTEND ssrindex PRINTED BY { pr_ssrindex } + INTERPRETED BY { interp_index } +| [ int_or_var(i) ] -> { mk_index ~loc i } +END + + +(** Occurrence switch *) + +(* The standard syntax of complemented occurrence lists involves a single *) +(* initial "-", e.g., {-1 3 5}. An initial *) +(* "+" may be used to indicate positive occurrences (the default). The *) +(* "+" is optional, except if the list of occurrences starts with a *) +(* variable or is empty (to avoid confusion with a clear switch). The *) +(* empty positive switch "{+}" selects no occurrences, while the empty *) +(* negative switch "{-}" selects all occurrences explicitly; this is the *) +(* default, but "{-}" prevents the implicit clear, and can be used to *) +(* force dependent elimination -- see ndefectelimtac below. *) + +{ + +let pr_ssrocc _ _ _ = pr_occ + +open Pcoq.Prim + +} + +ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY { pr_ssrocc } +| [ natural(n) natural_list(occ) ] -> { + Some (false, List.map (check_index ~loc) (n::occ)) } +| [ "-" natural_list(occ) ] -> { Some (true, occ) } +| [ "+" natural_list(occ) ] -> { Some (false, occ) } +END + + +(* modality *) + +{ + +let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt () + +let wit_ssrmmod = add_genarg "ssrmmod" pr_mmod +let ssrmmod = Pcoq.create_generic_entry Pcoq.utactic "ssrmmod" (Genarg.rawwit wit_ssrmmod);; + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssrmmod; + ssrmmod: [[ "!" -> { Must } | LEFTQMARK -> { May } | "?" -> { May } ]]; +END + +(** Rewrite multiplier: !n ?n *) + +{ + +let pr_mult (n, m) = + if n > 0 && m <> Once then int n ++ pr_mmod m else pr_mmod m + +let pr_ssrmult _ _ _ = pr_mult + +} + +ARGUMENT EXTEND ssrmult_ne TYPED AS (int * ssrmmod) PRINTED BY { pr_ssrmult } + | [ natural(n) ssrmmod(m) ] -> { check_index ~loc n, m } + | [ ssrmmod(m) ] -> { notimes, m } +END + +ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY { pr_ssrmult } + | [ ssrmult_ne(m) ] -> { m } + | [ ] -> { nomult } +END + +{ + +(** Discharge occ switch (combined occurrence / clear switch *) + +let pr_docc = function + | None, occ -> pr_occ occ + | Some clr, _ -> pr_clear mt clr + +let pr_ssrdocc _ _ _ = pr_docc + +} + +ARGUMENT EXTEND ssrdocc TYPED AS (ssrclear option * ssrocc) PRINTED BY { pr_ssrdocc } +| [ "{" ssrocc(occ) "}" ] -> { mkocc occ } +| [ "{" ssrhyp_list(clr) "}" ] -> { mkclr clr } +END + +{ + +(* Old kinds of terms *) + +let input_ssrtermkind strm = match Util.stream_nth 0 strm with + | Tok.KEYWORD "(" -> xInParens + | Tok.KEYWORD "@" -> xWithAt + | _ -> xNoFlag + +let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind + +(* New kinds of terms *) + +let input_term_annotation strm = + match Stream.npeek 2 strm with + | Tok.KEYWORD "(" :: Tok.KEYWORD "(" :: _ -> `DoubleParens + | Tok.KEYWORD "(" :: _ -> `Parens + | Tok.KEYWORD "@" :: _ -> `At + | _ -> `None +let term_annotation = + Gram.Entry.of_parser "term_annotation" input_term_annotation + +(* terms *) + +(** Terms parsing. ********************************************************) + +(* Because we allow wildcards in term references, we need to stage the *) +(* interpretation of terms so that it occurs at the right time during *) +(* the execution of the tactic (e.g., so that we don't report an error *) +(* for a term that isn't actually used in the execution). *) +(* The term representation tracks whether the concrete initial term *) +(* started with an opening paren, which might avoid a conflict between *) +(* the ssrreflect term syntax and Gallina notation. *) + +(* Old terms *) +let pr_ssrterm _ _ _ = pr_term +let glob_ssrterm gs = function + | k, (_, Some c) -> k, Tacintern.intern_constr gs c + | ct -> ct +let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c +let interp_ssrterm _ gl t = Tacmach.project gl, t + +open Pcoq.Constr + +} + +ARGUMENT EXTEND ssrterm + PRINTED BY { pr_ssrterm } + INTERPRETED BY { interp_ssrterm } + GLOBALIZED BY { glob_ssrterm } SUBSTITUTED BY { subst_ssrterm } + RAW_PRINTED BY { pr_ssrterm } + GLOB_PRINTED BY { pr_ssrterm } +| [ "YouShouldNotTypeThis" constr(c) ] -> { mk_lterm c } +END + +GRAMMAR EXTEND Gram + GLOBAL: ssrterm; + ssrterm: [[ k = ssrtermkind; c = Pcoq.Constr.constr -> { mk_term k c } ]]; +END + +(* New terms *) + +{ + +let pp_ast_closure_term _ _ _ = pr_ast_closure_term + +} + +ARGUMENT EXTEND ast_closure_term + PRINTED BY { pp_ast_closure_term } + INTERPRETED BY { interp_ast_closure_term } + GLOBALIZED BY { glob_ast_closure_term } + SUBSTITUTED BY { subst_ast_closure_term } + RAW_PRINTED BY { pp_ast_closure_term } + GLOB_PRINTED BY { pp_ast_closure_term } + | [ term_annotation(a) constr(c) ] -> { mk_ast_closure_term a c } +END +ARGUMENT EXTEND ast_closure_lterm + PRINTED BY { pp_ast_closure_term } + INTERPRETED BY { interp_ast_closure_term } + GLOBALIZED BY { glob_ast_closure_term } + SUBSTITUTED BY { subst_ast_closure_term } + RAW_PRINTED BY { pp_ast_closure_term } + GLOB_PRINTED BY { pp_ast_closure_term } + | [ term_annotation(a) lconstr(c) ] -> { mk_ast_closure_term a c } +END + +(* Old Views *) + +{ + +let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c) + +let pr_ssrbwdview _ _ _ = pr_view + +} + +ARGUMENT EXTEND ssrbwdview TYPED AS ssrterm list + PRINTED BY { pr_ssrbwdview } +| [ "YouShouldNotTypeThis" ] -> { [] } +END + +(* Pcoq *) +GRAMMAR EXTEND Gram + GLOBAL: ssrbwdview; + ssrbwdview: [ + [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> { [mk_term xNoFlag c] } + | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrbwdview -> { + (mk_term xNoFlag c) :: w } ]]; +END + +(* New Views *) + +{ + +let pr_ssrfwdview _ _ _ = pr_view2 + +} + +ARGUMENT EXTEND ssrfwdview TYPED AS ast_closure_term list + PRINTED BY { pr_ssrfwdview } +| [ "YouShouldNotTypeThis" ] -> { [] } +END + +(* Pcoq *) +GRAMMAR EXTEND Gram + GLOBAL: ssrfwdview; + ssrfwdview: [ + [ test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr -> + { [mk_ast_closure_term `None c] } + | test_not_ssrslashnum; "/"; c = Pcoq.Constr.constr; w = ssrfwdview -> + { (mk_ast_closure_term `None c) :: w } ]]; +END + +(* ipats *) + +{ + +let remove_loc x = x.CAst.v + +let ipat_of_intro_pattern p = Tactypes.( + let rec ipat_of_intro_pattern = function + | IntroNaming (IntroIdentifier id) -> IPatId id + | IntroAction IntroWildcard -> IPatAnon Drop + | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> + IPatCase + (List.map (List.map ipat_of_intro_pattern) + (List.map (List.map remove_loc) iorpat)) + | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> + IPatCase + [List.map ipat_of_intro_pattern (List.map remove_loc iandpat)] + | IntroNaming IntroAnonymous -> IPatAnon One + | IntroAction (IntroRewrite b) -> IPatRewrite (allocc, if b then L2R else R2L) + | IntroNaming (IntroFresh id) -> IPatAnon One + | IntroAction (IntroApplyOn _) -> (* to do *) CErrors.user_err (Pp.str "TO DO") + | IntroAction (IntroInjection ips) -> + IPatInj [List.map ipat_of_intro_pattern (List.map remove_loc ips)] + | IntroForthcoming _ -> + (* Unable to determine which kind of ipat interp_introid could + * return [HH] *) + assert false + in + ipat_of_intro_pattern p +) + +let rec map_ipat map_id map_ssrhyp map_ast_closure_term = function + | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop) as x -> x + | IPatId id -> IPatId (map_id id) + | IPatAbstractVars l -> IPatAbstractVars (List.map map_id l) + | IPatClear clr -> IPatClear (List.map map_ssrhyp clr) + | IPatCase iorpat -> IPatCase (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) + | IPatDispatch iorpat -> IPatDispatch (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) + | IPatInj iorpat -> IPatInj (List.map (List.map (map_ipat map_id map_ssrhyp map_ast_closure_term)) iorpat) + | IPatView (clr,v) -> IPatView (clr,List.map map_ast_closure_term v) + | IPatTac _ -> assert false (*internal usage only *) + +let wit_ssripatrep = add_genarg "ssripatrep" pr_ipat + +let pr_ssripat _ _ _ = pr_ipat +let pr_ssripats _ _ _ = pr_ipats +let pr_ssriorpat _ _ _ = pr_iorpat + +let intern_ipat ist = + map_ipat + (fun id -> id) + (intern_hyp ist) + (glob_ast_closure_term ist) + +let intern_ipats ist = List.map (intern_ipat ist) + +let interp_intro_pattern = interp_wit wit_intro_pattern + +let interp_introid ist gl id = + try IntroNaming (IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (Loc.tag id)))))) + with _ -> (snd (interp_intro_pattern ist gl (CAst.make @@ IntroNaming (IntroIdentifier id)))).CAst.v + +let get_intro_id = function + | IntroNaming (IntroIdentifier id) -> id + | _ -> assert false + +let rec add_intro_pattern_hyps ipat hyps = + let {CAst.loc=loc;v=ipat} = ipat in + match ipat with + | IntroNaming (IntroIdentifier id) -> + if not_section_id id then SsrHyp (loc, id) :: hyps else + hyp_err ?loc "Can't delete section hypothesis " id + | IntroAction IntroWildcard -> hyps + | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> + List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps + | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> + List.fold_right add_intro_pattern_hyps iandpat hyps + | IntroNaming IntroAnonymous -> [] + | IntroNaming (IntroFresh _) -> [] + | IntroAction (IntroRewrite _) -> hyps + | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps + | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps + | IntroForthcoming _ -> + (* As in ipat_of_intro_pattern, was unable to determine which kind + of ipat interp_introid could return [HH] *) assert false + +(* We interp the ipat using the standard ltac machinery for ids, since + * we have no clue what a name could be bound to (maybe another ipat) *) +let interp_ipat ist gl = + let ltacvar id = Id.Map.mem id ist.Tacinterp.lfun in + let rec interp = function + | IPatId id when ltacvar id -> + ipat_of_intro_pattern (interp_introid ist gl id) + | IPatId _ as x -> x + | IPatClear clr -> + let add_hyps (SsrHyp (loc, id) as hyp) hyps = + if not (ltacvar id) then hyp :: hyps else + add_intro_pattern_hyps CAst.(make ?loc (interp_introid ist gl id)) hyps in + let clr' = List.fold_right add_hyps clr [] in + check_hyps_uniq [] clr'; IPatClear clr' + | IPatCase(iorpat) -> + IPatCase(List.map (List.map interp) iorpat) + | IPatDispatch(iorpat) -> + IPatDispatch(List.map (List.map interp) iorpat) + | IPatInj iorpat -> IPatInj (List.map (List.map interp) iorpat) + | IPatAbstractVars l -> + IPatAbstractVars (List.map get_intro_id (List.map (interp_introid ist gl) l)) + | IPatView (clr,l) -> IPatView (clr,List.map (fun x -> snd(interp_ast_closure_term ist + gl x)) l) + | (IPatSimpl _ | IPatAnon _ | IPatRewrite _ | IPatNoop) as x -> x + | IPatTac _ -> assert false (*internal usage only *) + in + interp + +let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l + +let pushIPatRewrite = function + | pats :: orpat -> (IPatRewrite (allocc, L2R) :: pats) :: orpat + | [] -> [] + +let pushIPatNoop = function + | pats :: orpat -> (IPatNoop :: pats) :: orpat + | [] -> [] + +} + +ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY { pr_ssripats } + INTERPRETED BY { interp_ipats } + GLOBALIZED BY { intern_ipats } + | [ "_" ] -> { [IPatAnon Drop] } + | [ "*" ] -> { [IPatAnon All] } + (* + | [ "^" "*" ] -> { [IPatFastMode] } + | [ "^" "_" ] -> { [IPatSeed `Wild] } + | [ "^_" ] -> { [IPatSeed `Wild] } + | [ "^" "?" ] -> { [IPatSeed `Anon] } + | [ "^?" ] -> { [IPatSeed `Anon] } + | [ "^" ident(id) ] -> { [IPatSeed (`Id(id,`Pre))] } + | [ "^" "~" ident(id) ] -> { [IPatSeed (`Id(id,`Post))] } + | [ "^~" ident(id) ] -> { [IPatSeed (`Id(id,`Post))] } + *) + | [ ident(id) ] -> { [IPatId id] } + | [ "?" ] -> { [IPatAnon One] } +(* TODO | [ "+" ] -> [ [IPatAnon Temporary] ] *) + | [ ssrsimpl_ne(sim) ] -> { [IPatSimpl sim] } + | [ ssrdocc(occ) "->" ] -> { match occ with + | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") + | None, occ -> [IPatRewrite (occ, L2R)] + | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, L2R)] } + | [ ssrdocc(occ) "<-" ] -> { match occ with + | Some [], _ -> CErrors.user_err ~loc (str"occ_switch expected") + | None, occ -> [IPatRewrite (occ, R2L)] + | Some clr, _ -> [IPatClear clr; IPatRewrite (allocc, R2L)] } + | [ ssrdocc(occ) ssrfwdview(v) ] -> { match occ with + | Some [], _ -> [IPatView (true,v)] + | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl;IPatView (false,v)] + | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here") } + | [ ssrdocc(occ) ] -> { match occ with + | Some cl, _ -> check_hyps_uniq [] cl; [IPatClear cl] + | _ -> CErrors.user_err ~loc (str"Only identifiers are allowed here") } + | [ "->" ] -> { [IPatRewrite (allocc, L2R)] } + | [ "<-" ] -> { [IPatRewrite (allocc, R2L)] } + | [ "-" ] -> { [IPatNoop] } + | [ "-/" "=" ] -> { [IPatNoop;IPatSimpl(Simpl ~-1)] } + | [ "-/=" ] -> { [IPatNoop;IPatSimpl(Simpl ~-1)] } + | [ "-/" "/" ] -> { [IPatNoop;IPatSimpl(Cut ~-1)] } + | [ "-//" ] -> { [IPatNoop;IPatSimpl(Cut ~-1)] } + | [ "-/" integer(n) "/" ] -> { [IPatNoop;IPatSimpl(Cut n)] } + | [ "-/" "/=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } + | [ "-//" "=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } + | [ "-//=" ] -> { [IPatNoop;IPatSimpl(SimplCut (~-1,~-1))] } + | [ "-/" integer(n) "/=" ] -> { [IPatNoop;IPatSimpl(SimplCut (n,~-1))] } + | [ "-/" integer(n) "/" integer (m) "=" ] -> + { [IPatNoop;IPatSimpl(SimplCut(n,m))] } + | [ ssrfwdview(v) ] -> { [IPatView (false,v)] } + | [ "[" ":" ident_list(idl) "]" ] -> { [IPatAbstractVars idl] } + | [ "[:" ident_list(idl) "]" ] -> { [IPatAbstractVars idl] } +END + +ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY { pr_ssripats } + | [ ssripat(i) ssripats(tl) ] -> { i @ tl } + | [ ] -> { [] } +END + +ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY { pr_ssriorpat } +| [ ssripats(pats) "|" ssriorpat(orpat) ] -> { pats :: orpat } +| [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> { pats :: pushIPatRewrite orpat } +| [ ssripats(pats) "|-" ssriorpat(orpat) ] -> { pats :: pushIPatNoop orpat } +| [ ssripats(pats) "|->" ssriorpat(orpat) ] -> { pats :: pushIPatRewrite orpat } +| [ ssripats(pats) "||" ssriorpat(orpat) ] -> { pats :: [] :: orpat } +| [ ssripats(pats) "|||" ssriorpat(orpat) ] -> { pats :: [] :: [] :: orpat } +| [ ssripats(pats) "||||" ssriorpat(orpat) ] -> { [pats; []; []; []] @ orpat } +| [ ssripats(pats) ] -> { [pats] } +END + +{ + +let reject_ssrhid strm = + match Util.stream_nth 0 strm with + | Tok.KEYWORD "[" -> + (match Util.stream_nth 1 strm with + | Tok.KEYWORD ":" -> raise Stream.Failure + | _ -> ()) + | _ -> () + +let test_nohidden = Pcoq.Gram.Entry.of_parser "test_ssrhid" reject_ssrhid + +} + +ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY { pr_ssripat } + | [ "YouShouldNotTypeThis" ssriorpat(x) ] -> { IPatCase(x) } +END + +(* Pcoq *) +GRAMMAR EXTEND Gram + GLOBAL: ssrcpat; + ssrcpat: [ + [ test_nohidden; "["; iorpat = ssriorpat; "]" -> { +(* check_no_inner_seed !@loc false iorpat; + IPatCase (understand_case_type iorpat) *) + IPatCase iorpat } +(* + | test_nohidden; "("; iorpat = ssriorpat; ")" -> +(* check_no_inner_seed !@loc false iorpat; + IPatCase (understand_case_type iorpat) *) + IPatDispatch iorpat +*) + | test_nohidden; "[="; iorpat = ssriorpat; "]" -> { +(* check_no_inner_seed !@loc false iorpat; *) + IPatInj iorpat } ]]; +END + +GRAMMAR EXTEND Gram + GLOBAL: ssripat; + ssripat: [[ pat = ssrcpat -> { [pat] } ]]; +END + +ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY { pr_ssripats } + | [ ssripat(i) ssripats(tl) ] -> { i @ tl } + END + +(* subsets of patterns *) + +{ + +(* TODO: review what this function does, it looks suspicious *) +let check_ssrhpats loc w_binders ipats = + let err_loc s = CErrors.user_err ~loc ~hdr:"ssreflect" s in + let clr, ipats = + let rec aux clr = function + | IPatClear cl :: tl -> aux (clr @ cl) tl +(* | IPatSimpl (cl, sim) :: tl -> clr @ cl, IPatSimpl ([], sim) :: tl *) + | tl -> clr, tl + in aux [] ipats in + let simpl, ipats = + match List.rev ipats with + | IPatSimpl _ as s :: tl -> [s], List.rev tl + | _ -> [], ipats in + if simpl <> [] && not w_binders then + err_loc (str "No s-item allowed here: " ++ pr_ipats simpl); + let ipat, binders = + let rec loop ipat = function + | [] -> ipat, [] + | ( IPatId _| IPatAnon _| IPatCase _ | IPatDispatch _ | IPatRewrite _ as i) :: tl -> + if w_binders then + if simpl <> [] && tl <> [] then + err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) + else if not (List.for_all (function IPatId _ -> true | _ -> false) tl) + then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) + else ipat @ [i], tl + else + if tl = [] then ipat @ [i], [] + else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl) + | hd :: tl -> loop (ipat @ [hd]) tl + in loop [] ipats in + ((clr, ipat), binders), simpl + +let pr_hpats (((clr, ipat), binders), simpl) = + pr_clear mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl +let pr_ssrhpats _ _ _ = pr_hpats +let pr_ssrhpats_wtransp _ _ _ (_, x) = pr_hpats x + +} + +ARGUMENT EXTEND ssrhpats TYPED AS (((ssrclear * ssripat) * ssripat) * ssripat) +PRINTED BY { pr_ssrhpats } + | [ ssripats(i) ] -> { check_ssrhpats loc true i } +END + +ARGUMENT EXTEND ssrhpats_wtransp + TYPED AS (bool * (((ssrclear * ssripats) * ssripats) * ssripats)) + PRINTED BY { pr_ssrhpats_wtransp } + | [ ssripats(i) ] -> { false,check_ssrhpats loc true i } + | [ ssripats(i) "@" ssripats(j) ] -> { true,check_ssrhpats loc true (i @ j) } +END + +ARGUMENT EXTEND ssrhpats_nobs +TYPED AS (((ssrclear * ssripats) * ssripats) * ssripats) PRINTED BY { pr_ssrhpats } + | [ ssripats(i) ] -> { check_ssrhpats loc false i } +END + +ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY { pr_ssripat } + | [ "->" ] -> { IPatRewrite (allocc, L2R) } + | [ "<-" ] -> { IPatRewrite (allocc, R2L) } +END + +{ + +let pr_intros sep intrs = + if intrs = [] then mt() else sep () ++ str "=>" ++ pr_ipats intrs +let pr_ssrintros _ _ _ = pr_intros mt + +} + +ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat + PRINTED BY { pr_ssrintros } + | [ "=>" ssripats_ne(pats) ] -> { pats } +(* TODO | [ "=>" ">" ssripats_ne(pats) ] -> { IPatFastMode :: pats } + | [ "=>>" ssripats_ne(pats) ] -> [ IPatFastMode :: pats ] *) +END + +ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY { pr_ssrintros } + | [ ssrintros_ne(intrs) ] -> { intrs } + | [ ] -> { [] } +END + +{ + +let pr_ssrintrosarg _ _ prt (tac, ipats) = + prt tacltop tac ++ pr_intros spc ipats + +} + +ARGUMENT EXTEND ssrintrosarg TYPED AS (tactic * ssrintros) + PRINTED BY { pr_ssrintrosarg } +| [ "YouShouldNotTypeThis" ssrtacarg(arg) ssrintros_ne(ipats) ] -> { arg, ipats } +END + +TACTIC EXTEND ssrtclintros +| [ "YouShouldNotTypeThis" ssrintrosarg(arg) ] -> + { let tac, intros = arg in + ssrevaltac ist tac <*> tclIPATssr intros } +END + +{ + +(** Defined identifier *) +let pr_ssrfwdid id = pr_spc () ++ pr_id id + +let pr_ssrfwdidx _ _ _ = pr_ssrfwdid + +} + +(* We use a primitive parser for the head identifier of forward *) +(* tactis to avoid syntactic conflicts with basic Coq tactics. *) +ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY { pr_ssrfwdidx } + | [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +{ + +let accept_ssrfwdid strm = + match stream_nth 0 strm with + | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm + | _ -> raise Stream.Failure + +let test_ssrfwdid = Gram.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssrfwdid; + ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> { id } ]]; + END + + +(* by *) +(** Tactical arguments. *) + +(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) +(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) +(* and subgoal reordering tacticals (; first & ; last), respectively. *) + +{ + +let pr_ortacs prt = + let rec pr_rec = function + | [None] -> spc() ++ str "|" ++ spc() + | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs + | Some tac :: tacs -> spc() ++ str "| " ++ prt tacltop tac ++ pr_rec tacs + | [] -> mt() in + function + | [None] -> spc() + | None :: tacs -> pr_rec tacs + | Some tac :: tacs -> prt tacltop tac ++ pr_rec tacs + | [] -> mt() +let pr_ssrortacs _ _ = pr_ortacs + +} + +ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY { pr_ssrortacs } +| [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> { Some tac :: tacs } +| [ ssrtacarg(tac) "|" ] -> { [Some tac; None] } +| [ ssrtacarg(tac) ] -> { [Some tac] } +| [ "|" ssrortacs(tacs) ] -> { None :: tacs } +| [ "|" ] -> { [None; None] } +END + +{ + +let pr_hintarg prt = function + | true, tacs -> hv 0 (str "[ " ++ pr_ortacs prt tacs ++ str " ]") + | false, [Some tac] -> prt tacltop tac + | _, _ -> mt() + +let pr_ssrhintarg _ _ = pr_hintarg + +} + +ARGUMENT EXTEND ssrhintarg TYPED AS (bool * ssrortacs) PRINTED BY { pr_ssrhintarg } +| [ "[" "]" ] -> { nullhint } +| [ "[" ssrortacs(tacs) "]" ] -> { mk_orhint tacs } +| [ ssrtacarg(arg) ] -> { mk_hint arg } +END + +ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY { pr_ssrhintarg } +| [ "[" ssrortacs(tacs) "]" ] -> { mk_orhint tacs } +END + +{ + +let pr_hint prt arg = + if arg = nohint then mt() else str "by " ++ pr_hintarg prt arg +let pr_ssrhint _ _ = pr_hint + +} + +ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY { pr_ssrhint } +| [ ] -> { nohint } +END +(** The "in" pseudo-tactical *) + +(* We can't make "in" into a general tactical because this would create a *) +(* crippling conflict with the ltac let .. in construct. Hence, we add *) +(* explicitly an "in" suffix to all the extended tactics for which it is *) +(* relevant (including move, case, elim) and to the extended do tactical *) +(* below, which yields a general-purpose "in" of the form do [...] in ... *) + +(* This tactical needs to come before the intro tactics because the latter *) +(* must take precautions in order not to interfere with the discharged *) +(* assumptions. This is especially difficult for discharged "let"s, which *) +(* the default simpl and unfold tactics would erase blindly. *) + +{ + +open Ssrmatching_plugin.Ssrmatching +open Ssrmatching_plugin.G_ssrmatching + +let pr_wgen = function + | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id + | (clr, Some((id,k),Some p)) -> + spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++ + pr_cpattern p ++ str ")" + | (clr, None) -> spc () ++ pr_clear mt clr +let pr_ssrwgen _ _ _ = pr_wgen + +} + +(* no globwith for char *) +ARGUMENT EXTEND ssrwgen + TYPED AS (ssrclear * ((ssrhoi_hyp * string) * cpattern option) option) + PRINTED BY { pr_ssrwgen } +| [ ssrclear_ne(clr) ] -> { clr, None } +| [ ssrhoi_hyp(hyp) ] -> { [], Some((hyp, " "), None) } +| [ "@" ssrhoi_hyp(hyp) ] -> { [], Some((hyp, "@"), None) } +| [ "(" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + { [], Some ((id," "),Some p) } +| [ "(" ssrhoi_id(id) ")" ] -> { [], Some ((id,"("), None) } +| [ "(@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + { [], Some ((id,"@"),Some p) } +| [ "(" "@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + { [], Some ((id,"@"),Some p) } +END + +{ + +let pr_clseq = function + | InGoal | InHyps -> mt () + | InSeqGoal -> str "|- *" + | InHypsSeqGoal -> str " |- *" + | InHypsGoal -> str " *" + | InAll -> str "*" + | InHypsSeq -> str " |-" + | InAllHyps -> str "* |-" + +let wit_ssrclseq = add_genarg "ssrclseq" pr_clseq +let pr_clausehyps = pr_list pr_spc pr_wgen +let pr_ssrclausehyps _ _ _ = pr_clausehyps + +} + +ARGUMENT EXTEND ssrclausehyps +TYPED AS ssrwgen list PRINTED BY { pr_ssrclausehyps } +| [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> { hyp :: hyps } +| [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> { hyp :: hyps } +| [ ssrwgen(hyp) ] -> { [hyp] } +END + +{ + +(* type ssrclauses = ssrahyps * ssrclseq *) + +let pr_clauses (hyps, clseq) = + if clseq = InGoal then mt () + else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq +let pr_ssrclauses _ _ _ = pr_clauses + +} + +ARGUMENT EXTEND ssrclauses TYPED AS (ssrwgen list * ssrclseq) + PRINTED BY { pr_ssrclauses } + | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> { hyps, InHypsSeqGoal } + | [ "in" ssrclausehyps(hyps) "|-" ] -> { hyps, InHypsSeq } + | [ "in" ssrclausehyps(hyps) "*" ] -> { hyps, InHypsGoal } + | [ "in" ssrclausehyps(hyps) ] -> { hyps, InHyps } + | [ "in" "|-" "*" ] -> { [], InSeqGoal } + | [ "in" "*" ] -> { [], InAll } + | [ "in" "*" "|-" ] -> { [], InAllHyps } + | [ ] -> { [], InGoal } +END + + +{ + +(** Definition value formatting *) + +(* We use an intermediate structure to correctly render the binder list *) +(* abbreviations. We use a list of hints to extract the binders and *) +(* base term from a term, for the two first levels of representation of *) +(* of constr terms. *) + +let pr_binder prl = function + | Bvar x -> + pr_name x + | Bdecl (xs, t) -> + str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")" + | Bdef (x, None, v) -> + str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")" + | Bdef (x, Some t, v) -> + str "(" ++ pr_name x ++ str " : " ++ prl t ++ + str " := " ++ prl v ++ str ")" + | Bstruct x -> + str "{struct " ++ pr_name x ++ str "}" + | Bcast t -> + str ": " ++ prl t + +let rec format_local_binders h0 bl0 = match h0, bl0 with + | BFvar :: h, CLocalAssum ([{CAst.v=x}], _, _) :: bl -> + Bvar x :: format_local_binders h bl + | BFdecl _ :: h, CLocalAssum (lxs, _, t) :: bl -> + Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: format_local_binders h bl + | BFdef :: h, CLocalDef ({CAst.v=x}, v, oty) :: bl -> + Bdef (x, oty, v) :: format_local_binders h bl + | _ -> [] + +let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with + | BFvar :: h, { v = CLambdaN ([CLocalAssum([{CAst.v=x}], _, _)], c) } -> + let bs, c' = format_constr_expr h c in + Bvar x :: bs, c' + | BFdecl _:: h, { v = CLambdaN ([CLocalAssum(lxs, _, t)], c) } -> + let bs, c' = format_constr_expr h c in + Bdecl (List.map (fun x -> x.CAst.v) lxs, t) :: bs, c' + | BFdef :: h, { v = CLetIn({CAst.v=x}, v, oty, c) } -> + let bs, c' = format_constr_expr h c in + Bdef (x, oty, v) :: bs, c' + | [BFcast], { v = CCast (c, Glob_term.CastConv t) } -> + [Bcast t], c + | BFrec (has_str, has_cast) :: h, + { v = CFix ( _, [_, (Some locn, CStructRec), bl, t, c]) } -> + let bs = format_local_binders h bl in + let bstr = if has_str then [Bstruct (Name locn.CAst.v)] else [] in + bs @ bstr @ (if has_cast then [Bcast t] else []), c + | BFrec (_, has_cast) :: h, { v = CCoFix ( _, [_, bl, t, c]) } -> + format_local_binders h bl @ (if has_cast then [Bcast t] else []), c + | _, c -> + [], c + +(** Forward chaining argument *) + +(* There are three kinds of forward definitions: *) +(* - Hint: type only, cast to Type, may have proof hint. *) +(* - Have: type option + value, no space before type *) +(* - Pose: binders + value, space before binders. *) + +let pr_fwdkind = function + | FwdHint (s,_) -> str (s ^ " ") | _ -> str " :=" ++ spc () +let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk + +let wit_ssrfwdfmt = add_genarg "ssrfwdfmt" pr_fwdfmt + +(* type ssrfwd = ssrfwdfmt * ssrterm *) + +let mkFwdVal fk c = ((fk, []), c) +let mkssrFwdVal fk c = ((fk, []), (c,None)) +let dC t = Glob_term.CastConv t + +let same_ist { interp_env = x } { interp_env = y } = + match x,y with + | None, None -> true + | Some a, Some b -> a == b + | _ -> false + +let mkFwdCast fk ?loc ?c t = + let c = match c with + | None -> mkCHole loc + | Some c -> assert (same_ist t c); c.body in + ((fk, [BFcast]), + { t with annotation = `None; + body = (CAst.make ?loc @@ CCast (c, dC t.body)) }) + +let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t)) + +let mkFwdHint s t = + let loc = Constrexpr_ops.constr_loc t.body in + mkFwdCast (FwdHint (s,false)) ?loc t +let mkFwdHintNoTC s t = + let loc = Constrexpr_ops.constr_loc t.body in + mkFwdCast (FwdHint (s,true)) ?loc t + +let pr_gen_fwd prval prc prlc fk (bs, c) = + let prc s = str s ++ spc () ++ prval prc prlc c in + match fk, bs with + | FwdHint (s,_), [Bcast t] -> str s ++ spc () ++ prlc t + | FwdHint (s,_), _ -> prc (s ^ "(* typeof *)") + | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :=" + | _, [] -> prc " :=" + | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :=" + +let pr_fwd_guarded prval prval' = function +| (fk, h), c -> + pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c.body) + +let pr_unguarded prc prlc = prlc + +let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded +let pr_ssrfwd _ _ _ = pr_fwd + +} + +ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ast_closure_lterm) PRINTED BY { pr_ssrfwd } + | [ ":=" ast_closure_lterm(c) ] -> { mkFwdVal FwdPose c } + | [ ":" ast_closure_lterm (t) ":=" ast_closure_lterm(c) ] -> { mkFwdCast FwdPose ~loc t ~c } +END + +(** Independent parsing for binders *) + +(* The pose, pose fix, and pose cofix tactics use these internally to *) +(* parse argument fragments. *) + +{ + +let pr_ssrbvar prc _ _ v = prc v + +} + +ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY { pr_ssrbvar } +| [ ident(id) ] -> { mkCVar ~loc id } +| [ "_" ] -> { mkCHole (Some loc) } +END + +{ + +let bvar_lname = let open CAst in function + | { v = CRef (qid, _) } when qualid_is_ident qid -> + CAst.make ?loc:qid.CAst.loc @@ Name (qualid_basename qid) + | { loc = loc } -> CAst.make ?loc Anonymous + +let pr_ssrbinder prc _ _ (_, c) = prc c + +} + +ARGUMENT EXTEND ssrbinder TYPED AS (ssrfwdfmt * constr) PRINTED BY { pr_ssrbinder } + | [ ssrbvar(bv) ] -> + { let { CAst.loc=xloc } as x = bvar_lname bv in + (FwdPose, [BFvar]), + CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) } + | [ "(" ssrbvar(bv) ")" ] -> + { let { CAst.loc=xloc } as x = bvar_lname bv in + (FwdPose, [BFvar]), + CAst.make ~loc @@ CLambdaN ([CLocalAssum([x],Default Explicit,mkCHole xloc)],mkCHole (Some loc)) } + | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] -> + { let x = bvar_lname bv in + (FwdPose, [BFdecl 1]), + CAst.make ~loc @@ CLambdaN ([CLocalAssum([x], Default Explicit, t)], mkCHole (Some loc)) } + | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] -> + { let xs = List.map bvar_lname (bv :: bvs) in + let n = List.length xs in + (FwdPose, [BFdecl n]), + CAst.make ~loc @@ CLambdaN ([CLocalAssum (xs, Default Explicit, t)], mkCHole (Some loc)) } + | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] -> + { (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, Some t, mkCHole (Some loc)) } + | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] -> + { (FwdPose,[BFdef]), CAst.make ~loc @@ CLetIn (bvar_lname id, v, None, mkCHole (Some loc)) } + END + +GRAMMAR EXTEND Gram + GLOBAL: ssrbinder; + ssrbinder: [ + [ ["of" -> { () } | "&" -> { () } ]; c = operconstr LEVEL "99" -> { + (FwdPose, [BFvar]), + CAst.make ~loc @@ CLambdaN ([CLocalAssum ([CAst.make ~loc Anonymous],Default Explicit,c)],mkCHole (Some loc)) } ] + ]; +END + +{ + +let rec binders_fmts = function + | ((_, h), _) :: bs -> h @ binders_fmts bs + | _ -> [] + +let push_binders c2 bs = + let loc2 = constr_loc c2 in let mkloc loc1 = Loc.merge_opt loc1 loc2 in + let open CAst in + let rec loop ty c = function + | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs when ty -> + CAst.make ?loc:(mkloc loc1) @@ CProdN (b, loop ty c bs) + | (_, { loc = loc1; v = CLambdaN (b, _) } ) :: bs -> + CAst.make ?loc:(mkloc loc1) @@ CLambdaN (b, loop ty c bs) + | (_, { loc = loc1; v = CLetIn (x, v, oty, _) } ) :: bs -> + CAst.make ?loc:(mkloc loc1) @@ CLetIn (x, v, oty, loop ty c bs) + | [] -> c + | _ -> anomaly "binder not a lambda nor a let in" in + match c2 with + | { loc; v = CCast (ct, Glob_term.CastConv cty) } -> + CAst.make ?loc @@ (CCast (loop false ct bs, Glob_term.CastConv (loop true cty bs))) + | ct -> loop false ct bs + +let rec fix_binders = let open CAst in function + | (_, { v = CLambdaN ([CLocalAssum(xs, _, t)], _) } ) :: bs -> + CLocalAssum (xs, Default Explicit, t) :: fix_binders bs + | (_, { v = CLetIn (x, v, oty, _) } ) :: bs -> + CLocalDef (x, v, oty) :: fix_binders bs + | _ -> [] + +let pr_ssrstruct _ _ _ = function + | Some id -> str "{struct " ++ pr_id id ++ str "}" + | None -> mt () + +} + +ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY { pr_ssrstruct } +| [ "{" "struct" ident(id) "}" ] -> { Some id } +| [ ] -> { None } +END + +(** The "pose" tactic *) + +(* The plain pose form. *) + +{ + +let bind_fwd bs ((fk, h), c) = + (fk,binders_fmts bs @ h), { c with body = push_binders c.body bs } + +} + +ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY { pr_ssrfwd } + | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> { bind_fwd bs fwd } +END + +(* The pose fix form. *) + +{ + +let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd + +let bvar_locid = function + | { CAst.v = CRef (qid, _) } when qualid_is_ident qid -> + CAst.make ?loc:qid.CAst.loc (qualid_basename qid) + | _ -> CErrors.user_err (Pp.str "Missing identifier after \"(co)fix\"") + +} + +ARGUMENT EXTEND ssrfixfwd TYPED AS (ident * ssrfwd) PRINTED BY { pr_ssrfixfwd } + | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] -> + { let { CAst.v=id } as lid = bvar_locid bv in + let (fk, h), ac = fwd in + let c = ac.body in + let has_cast, t', c' = match format_constr_expr h c with + | [Bcast t'], c' -> true, t', c' + | _ -> false, mkCHole (constr_loc c), c in + let lb = fix_binders bs in + let has_struct, i = + let rec loop = function + | {CAst.loc=l'; v=Name id'} :: _ when Option.equal Id.equal sid (Some id') -> + true, CAst.make ?loc:l' id' + | [{CAst.loc=l';v=Name id'}] when sid = None -> + false, CAst.make ?loc:l' id' + | _ :: bn -> loop bn + | [] -> CErrors.user_err (Pp.str "Bad structural argument") in + loop (names_of_local_assums lb) in + let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in + let fix = CAst.make ~loc @@ CFix (lid, [lid, (Some i, CStructRec), lb, t', c']) in + id, ((fk, h'), { ac with body = fix }) } +END + + +(* The pose cofix form. *) + +{ + +let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd + +} + +ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY { pr_ssrcofixfwd } + | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] -> + { let { CAst.v=id } as lid = bvar_locid bv in + let (fk, h), ac = fwd in + let c = ac.body in + let has_cast, t', c' = match format_constr_expr h c with + | [Bcast t'], c' -> true, t', c' + | _ -> false, mkCHole (constr_loc c), c in + let h' = BFrec (false, has_cast) :: binders_fmts bs in + let cofix = CAst.make ~loc @@ CCoFix (lid, [lid, fix_binders bs, t', c']) in + id, ((fk, h'), { ac with body = cofix }) + } +END + +{ + +(* This does not print the type, it should be fixed... *) +let pr_ssrsetfwd _ _ _ (((fk,_),(t,_)), docc) = + pr_gen_fwd (fun _ _ -> pr_cpattern) + (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t) + +} + +ARGUMENT EXTEND ssrsetfwd +TYPED AS ((ssrfwdfmt * (lcpattern * ast_closure_lterm option)) * ssrdocc) +PRINTED BY { pr_ssrsetfwd } +| [ ":" ast_closure_lterm(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> + { mkssrFwdCast FwdPose loc t c, mkocc occ } +| [ ":" ast_closure_lterm(t) ":=" lcpattern(c) ] -> + { mkssrFwdCast FwdPose loc t c, nodocc } +| [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> + { mkssrFwdVal FwdPose c, mkocc occ } +| [ ":=" lcpattern(c) ] -> { mkssrFwdVal FwdPose c, nodocc } +END + +{ + +let pr_ssrhavefwd _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint prt hint + +} + +ARGUMENT EXTEND ssrhavefwd TYPED AS (ssrfwd * ssrhint) PRINTED BY { pr_ssrhavefwd } +| [ ":" ast_closure_lterm(t) ssrhint(hint) ] -> { mkFwdHint ":" t, hint } +| [ ":" ast_closure_lterm(t) ":=" ast_closure_lterm(c) ] -> { mkFwdCast FwdHave ~loc t ~c, nohint } +| [ ":" ast_closure_lterm(t) ":=" ] -> { mkFwdHintNoTC ":" t, nohint } +| [ ":=" ast_closure_lterm(c) ] -> { mkFwdVal FwdHave c, nohint } +END + +{ + +let intro_id_to_binder = List.map (function + | IPatId id -> + let { CAst.loc=xloc } as x = bvar_lname (mkCVar id) in + (FwdPose, [BFvar]), + CAst.make @@ CLambdaN ([CLocalAssum([x], Default Explicit, mkCHole xloc)], + mkCHole None) + | _ -> anomaly "non-id accepted as binder") + +let binder_to_intro_id = CAst.(List.map (function + | (FwdPose, [BFvar]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } + | (FwdPose, [BFdecl _]), { v = CLambdaN ([CLocalAssum(ids,_,_)],_) } -> + List.map (function {v=Name id} -> IPatId id | _ -> IPatAnon One) ids + | (FwdPose, [BFdef]), { v = CLetIn ({v=Name id},_,_,_) } -> [IPatId id] + | (FwdPose, [BFdef]), { v = CLetIn ({v=Anonymous},_,_,_) } -> [IPatAnon One] + | _ -> anomaly "ssrbinder is not a binder")) + +let pr_ssrhavefwdwbinders _ _ prt (tr,((hpats, (fwd, hint)))) = + pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint + +} + +ARGUMENT EXTEND ssrhavefwdwbinders + TYPED AS (bool * (ssrhpats * (ssrfwd * ssrhint))) + PRINTED BY { pr_ssrhavefwdwbinders } +| [ ssrhpats_wtransp(trpats) ssrbinder_list(bs) ssrhavefwd(fwd) ] -> + { let tr, pats = trpats in + let ((clr, pats), binders), simpl = pats in + let allbs = intro_id_to_binder binders @ bs in + let allbinders = binders @ List.flatten (binder_to_intro_id bs) in + let hint = bind_fwd allbs (fst fwd), snd fwd in + tr, ((((clr, pats), allbinders), simpl), hint) } +END + +{ + +let pr_ssrdoarg prc _ prt (((n, m), tac), clauses) = + pr_index n ++ pr_mmod m ++ pr_hintarg prt tac ++ pr_clauses clauses + +} + +ARGUMENT EXTEND ssrdoarg + TYPED AS (((ssrindex * ssrmmod) * ssrhintarg) * ssrclauses) + PRINTED BY { pr_ssrdoarg } +| [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +{ + +(* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *) + +let pr_seqtacarg prt = function + | (is_first, []), _ -> str (if is_first then "first" else "last") + | tac, Some dtac -> + hv 0 (pr_hintarg prt tac ++ spc() ++ str "|| " ++ prt tacltop dtac) + | tac, _ -> pr_hintarg prt tac + +let pr_ssrseqarg _ _ prt = function + | ArgArg 0, tac -> pr_seqtacarg prt tac + | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg prt tac + +} + +(* We must parse the index separately to resolve the conflict with *) +(* an unindexed tactic. *) +ARGUMENT EXTEND ssrseqarg TYPED AS (ssrindex * (ssrhintarg * tactic option)) + PRINTED BY { pr_ssrseqarg } +| [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +{ + +let sq_brace_tacnames = + ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"] + (* "by" is a keyword *) +let accept_ssrseqvar strm = + match stream_nth 0 strm with + | Tok.IDENT id when not (List.mem id sq_brace_tacnames) -> + accept_before_syms_or_ids ["["] ["first";"last"] strm + | _ -> raise Stream.Failure + +let test_ssrseqvar = Gram.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar + +let swaptacarg (loc, b) = (b, []), Some (TacId []) + +let check_seqtacarg dir arg = match snd arg, dir with + | ((true, []), Some (TacAtom (loc, _))), L2R -> + CErrors.user_err ?loc (str "expected \"last\"") + | ((false, []), Some (TacAtom (loc, _))), R2L -> + CErrors.user_err ?loc (str "expected \"first\"") + | _, _ -> arg + +let ssrorelse = Entry.create "ssrorelse" + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssrorelse ssrseqarg; + ssrseqidx: [ + [ test_ssrseqvar; id = Prim.ident -> { ArgVar (CAst.make ~loc id) } + | n = Prim.natural -> { ArgArg (check_index ~loc n) } + ] ]; + ssrswap: [[ IDENT "first" -> { loc, true } | IDENT "last" -> { loc, false } ]]; + ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> { tac } ]]; + ssrseqarg: [ + [ arg = ssrswap -> { noindex, swaptacarg arg } + | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> { i, (tac, def) } + | i = ssrseqidx; arg = ssrswap -> { i, swaptacarg arg } + | tac = tactic_expr LEVEL "3" -> { noindex, (mk_hint tac, None) } + ] ]; +END + +{ + +let tactic_expr = Pltac.tactic_expr + +} + +(** 1. Utilities *) + +(** Tactic-level diagnosis *) + +(* debug *) + +{ + +(* Let's play with the new proof engine API *) +let old_tac = V82.tactic + +} + +(** Name generation *) + +(* Since Coq now does repeated internal checks of its external lexical *) +(* rules, we now need to carve ssreflect reserved identifiers out of *) +(* out of the user namespace. We use identifiers of the form _id_ for *) +(* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *) +(* an extra leading _ if this might clash with an internal identifier. *) +(* We check for ssreflect identifiers in the ident grammar rule; *) +(* when the ssreflect Module is present this is normally an error, *) +(* but we provide a compatibility flag to reduce this to a warning. *) + +{ + +let ssr_reserved_ids = Summary.ref ~name:"SSR:idents" true + +let _ = + Goptions.declare_bool_option + { Goptions.optname = "ssreflect identifiers"; + Goptions.optkey = ["SsrIdents"]; + Goptions.optdepr = false; + Goptions.optread = (fun _ -> !ssr_reserved_ids); + Goptions.optwrite = (fun b -> ssr_reserved_ids := b) + } + +let is_ssr_reserved s = + let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_' + +let ssr_id_of_string loc s = + if is_ssr_reserved s && is_ssr_loaded () then begin + if !ssr_reserved_ids then + CErrors.user_err ~loc (str ("The identifier " ^ s ^ " is reserved.")) + else if is_internal_name s then + Feedback.msg_warning (str ("Conflict between " ^ s ^ " and ssreflect internal names.")) + else Feedback.msg_warning (str ( + "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" + ^ "Scripts with explicit references to anonymous variables are fragile.")) + end; Id.of_string s + +let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ()) + +} + +GRAMMAR EXTEND Gram + GLOBAL: Prim.ident; + Prim.ident: [[ s = IDENT; ssr_null_entry -> { ssr_id_of_string loc s } ]]; +END + +{ + +let perm_tag = "_perm_Hyp_" +let _ = add_internal_name (is_tagged perm_tag) + +} + +(* We must not anonymize context names discharged by the "in" tactical. *) + +(** Tactical extensions. *) + +(* The TACTIC EXTEND facility can't be used for defining new user *) +(* tacticals, because: *) +(* - the concrete syntax must start with a fixed string *) +(* We use the following workaround: *) +(* - We use the (unparsable) "YouShouldNotTypeThis" token for tacticals that *) +(* don't start with a token, then redefine the grammar and *) +(* printer using GEXTEND and set_pr_ssrtac, respectively. *) + +{ + +type ssrargfmt = ArgSsr of string | ArgSep of string + +let ssrtac_name name = { + mltac_plugin = "ssreflect_plugin"; + mltac_tactic = "ssr" ^ name; +} + +let ssrtac_entry name n = { + mltac_name = ssrtac_name name; + mltac_index = n; +} + +let set_pr_ssrtac name prec afmt = (* FIXME *) () (* + let fmt = List.map (function + | ArgSep s -> Egramml.GramTerminal s + | ArgSsr s -> Egramml.GramTerminal s + | ArgCoq at -> Egramml.GramTerminal "COQ_ARG") afmt in + let tacname = ssrtac_name name in () *) + +let ssrtac_atom ?loc name args = TacML (Loc.tag ?loc (ssrtac_entry name 0, args)) +let ssrtac_expr ?loc name args = ssrtac_atom ?loc name args + +let tclintros_expr ?loc tac ipats = + let args = [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrintrosarg) (tac, ipats))] in + ssrtac_expr ?loc "tclintros" args + +} + +GRAMMAR EXTEND Gram + GLOBAL: tactic_expr; + tactic_expr: LEVEL "1" [ RIGHTA + [ tac = tactic_expr; intros = ssrintros_ne -> { tclintros_expr ~loc tac intros } + ] ]; +END + + +(** Bracketing tactical *) + +(* The tactic pretty-printer doesn't know that some extended tactics *) +(* are actually tacticals. To prevent it from improperly removing *) +(* parentheses we override the parsing rule for bracketed tactic *) +(* expressions so that the pretty-print always reflects the input. *) +(* (Removing user-specified parentheses is dubious anyway). *) + +GRAMMAR EXTEND Gram + GLOBAL: tactic_expr; + ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> { Loc.tag ~loc (Tacexp tac) } ]]; + tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> { TacArg arg } ]]; +END + +(** The internal "done" and "ssrautoprop" tactics. *) + +(* For additional flexibility, "done" and "ssrautoprop" are *) +(* defined in Ltac. *) +(* Although we provide a default definition in ssreflect, *) +(* we look up the definition dynamically at each call point, *) +(* to allow for user extensions. "ssrautoprop" defaults to *) +(* trivial. *) + +{ + +let ssrautoprop gl = + try + let tacname = + try Tacenv.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop")) + with Not_found -> Tacenv.locate_tactic (ssrqid "ssrautoprop") in + let tacexpr = Loc.tag @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in + V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl + with Not_found -> V82.of_tactic (Auto.full_trivial []) gl + +let () = ssrautoprop_tac := ssrautoprop + +let tclBY tac = Tacticals.tclTHEN tac (donetac ~-1) + +(** Tactical arguments. *) + +(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) +(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) +(* and subgoal reordering tacticals (; first & ; last), respectively. *) + +(* Force use of the tactic_expr parsing entry, to rule out tick marks. *) + +(** The "by" tactical. *) + + +open Ssrfwd + +} + +TACTIC EXTEND ssrtclby +| [ "by" ssrhintarg(tac) ] -> { V82.tactic (hinttac ist true tac) } +END + +(* We can't parse "by" in ARGUMENT EXTEND because it will only be made *) +(* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *) + +GRAMMAR EXTEND Gram + GLOBAL: ssrhint simple_tactic; + ssrhint: [[ "by"; arg = ssrhintarg -> { arg } ]]; +END + +(** The "do" tactical. ********************************************************) + +(* +type ssrdoarg = ((ssrindex * ssrmmod) * ssrhint) * ssrclauses +*) +TACTIC EXTEND ssrtcldo +| [ "YouShouldNotTypeThis" "do" ssrdoarg(arg) ] -> { V82.tactic (ssrdotac ist arg) } +END + +{ + +let _ = set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"] + +let ssrdotac_expr ?loc n m tac clauses = + let arg = ((n, m), tac), clauses in + ssrtac_expr ?loc "tcldo" [Tacexpr.TacGeneric (in_gen (rawwit wit_ssrdoarg) arg)] + +} + +GRAMMAR EXTEND Gram + GLOBAL: tactic_expr; + ssrdotac: [ + [ tac = tactic_expr LEVEL "3" -> { mk_hint tac } + | tacs = ssrortacarg -> { tacs } + ] ]; + tactic_expr: LEVEL "3" [ RIGHTA + [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> + { ssrdotac_expr ~loc noindex m tac clauses } + | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses -> + { ssrdotac_expr ~loc noindex Once tac clauses } + | IDENT "do"; n = int_or_var; m = ssrmmod; + tac = ssrdotac; clauses = ssrclauses -> + { ssrdotac_expr ~loc (mk_index ~loc n) m tac clauses } + ] ]; +END + +{ + +(* We can't actually parse the direction separately because this *) +(* would introduce conflicts with the basic ltac syntax. *) +let pr_ssrseqdir _ _ _ = function + | L2R -> str ";" ++ spc () ++ str "first " + | R2L -> str ";" ++ spc () ++ str "last " + +} + +ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY { pr_ssrseqdir } +| [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +TACTIC EXTEND ssrtclseq +| [ "YouShouldNotTypeThis" ssrtclarg(tac) ssrseqdir(dir) ssrseqarg(arg) ] -> + { V82.tactic (tclSEQAT ist tac dir arg) } +END + +{ + +let _ = set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"] + +let tclseq_expr ?loc tac dir arg = + let arg1 = in_gen (rawwit wit_ssrtclarg) tac in + let arg2 = in_gen (rawwit wit_ssrseqdir) dir in + let arg3 = in_gen (rawwit wit_ssrseqarg) (check_seqtacarg dir arg) in + ssrtac_expr ?loc "tclseq" (List.map (fun x -> Tacexpr.TacGeneric x) [arg1; arg2; arg3]) + +} + +GRAMMAR EXTEND Gram + GLOBAL: tactic_expr; + ssr_first: [ + [ tac = ssr_first; ipats = ssrintros_ne -> { tclintros_expr ~loc tac ipats } + | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> { TacFirst tacl } + ] ]; + ssr_first_else: [ + [ tac1 = ssr_first; tac2 = ssrorelse -> { TacOrelse (tac1, tac2) } + | tac = ssr_first -> { tac } ]]; + tactic_expr: LEVEL "4" [ LEFTA + [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else -> + { TacThen (tac1, tac2) } + | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg -> + { tclseq_expr ~loc tac L2R arg } + | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg -> + { tclseq_expr ~loc tac R2L arg } + ] ]; +END + +(** 5. Bookkeeping tactics (clear, move, case, elim) *) + +(** Generalization (discharge) item *) + +(* An item is a switch + term pair. *) + +(* type ssrgen = ssrdocc * ssrterm *) + +{ + +let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt + +let pr_ssrgen _ _ _ = pr_gen + +} + +ARGUMENT EXTEND ssrgen TYPED AS (ssrdocc * cpattern) PRINTED BY { pr_ssrgen } +| [ ssrdocc(docc) cpattern(dt) ] -> { + match docc with + | Some [], _ -> CErrors.user_err ~loc (str"Clear flag {} not allowed here") + | _ -> docc, dt } +| [ cpattern(dt) ] -> { nodocc, dt } +END + +{ + +let has_occ ((_, occ), _) = occ <> None + +(** Generalization (discharge) sequence *) + +(* A discharge sequence is represented as a list of up to two *) +(* lists of d-items, plus an ident list set (the possibly empty *) +(* final clear switch). The main list is empty iff the command *) +(* is defective, and has length two if there is a sequence of *) +(* dependent terms (and in that case it is the first of the two *) +(* lists). Thus, the first of the two lists is never empty. *) + +(* type ssrgens = ssrgen list *) +(* type ssrdgens = ssrgens list * ssrclear *) + +let gens_sep = function [], [] -> mt | _ -> spc + +let pr_dgens pr_gen (gensl, clr) = + let prgens s gens = str s ++ pr_list spc pr_gen gens in + let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in + match gensl with + | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr + | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr + | [gens] -> prgens ": " gens ++ pr_clear spc clr + | _ -> pr_clear pr_spc clr + +let pr_ssrdgens _ _ _ = pr_dgens pr_gen + +let cons_gen gen = function + | gens :: gensl, clr -> (gen :: gens) :: gensl, clr + | _ -> anomaly "missing gen list" + +let cons_dep (gensl, clr) = + if List.length gensl = 1 then ([] :: gensl, clr) else + CErrors.user_err (Pp.str "multiple dependents switches '/'") + +} + +ARGUMENT EXTEND ssrdgens_tl TYPED AS (ssrgen list list * ssrclear) + PRINTED BY { pr_ssrdgens } +| [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> + { cons_gen (mkclr clr, dt) dgens } +| [ "{" ne_ssrhyp_list(clr) "}" ] -> + { [[]], clr } +| [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> + { cons_gen (mkocc occ, dt) dgens } +| [ "/" ssrdgens_tl(dgens) ] -> + { cons_dep dgens } +| [ cpattern(dt) ssrdgens_tl(dgens) ] -> + { cons_gen (nodocc, dt) dgens } +| [ ] -> + { [[]], [] } +END + +ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY { pr_ssrdgens } +| [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> { cons_gen gen dgens } +END + +(** Equations *) + +(* argument *) + +{ + +let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt () +let pr_ssreqid _ _ _ = pr_eqid + +} + +(* We must use primitive parsing here to avoid conflicts with the *) +(* basic move, case, and elim tactics. *) +ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY { pr_ssreqid } +| [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +{ + +let accept_ssreqid strm = + match Util.stream_nth 0 strm with + | Tok.IDENT _ -> accept_before_syms [":"] strm + | Tok.KEYWORD ":" -> () + | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] -> + accept_before_syms [":"] strm + | _ -> raise Stream.Failure + +let test_ssreqid = Gram.Entry.of_parser "test_ssreqid" accept_ssreqid + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssreqid; + ssreqpat: [ + [ id = Prim.ident -> { IPatId id } + | "_" -> { IPatAnon Drop } + | "?" -> { IPatAnon One } + | occ = ssrdocc; "->" -> { match occ with + | None, occ -> IPatRewrite (occ, L2R) + | _ -> CErrors.user_err ~loc (str"Only occurrences are allowed here") } + | occ = ssrdocc; "<-" -> { match occ with + | None, occ -> IPatRewrite (occ, R2L) + | _ -> CErrors.user_err ~loc (str "Only occurrences are allowed here") } + | "->" -> { IPatRewrite (allocc, L2R) } + | "<-" -> { IPatRewrite (allocc, R2L) } + ]]; + ssreqid: [ + [ test_ssreqid; pat = ssreqpat -> { Some pat } + | test_ssreqid -> { None } + ]]; +END + +(** Bookkeeping (discharge-intro) argument *) + +(* Since all bookkeeping ssr commands have the same discharge-intro *) +(* argument format we use a single grammar entry point to parse them. *) +(* the entry point parses only non-empty arguments to avoid conflicts *) +(* with the basic Coq tactics. *) + +(* type ssrarg = ssrbwdview * (ssreqid * (ssrdgens * ssripats)) *) + +{ + +let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) = + let pri = pr_intros (gens_sep dgens) in + pr_view2 view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats + +} + +ARGUMENT EXTEND ssrarg TYPED AS (ssrfwdview * (ssreqid * (ssrdgens * ssrintros))) + PRINTED BY { pr_ssrarg } +| [ ssrfwdview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> + { view, (eqid, (dgens, ipats)) } +| [ ssrfwdview(view) ssrclear(clr) ssrintros(ipats) ] -> + { view, (None, (([], clr), ipats)) } +| [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> + { [], (eqid, (dgens, ipats)) } +| [ ssrclear_ne(clr) ssrintros(ipats) ] -> + { [], (None, (([], clr), ipats)) } +| [ ssrintros_ne(ipats) ] -> + { [], (None, (([], []), ipats)) } +END + +(** The "clear" tactic *) + +(* We just add a numeric version that clears the n top assumptions. *) + +TACTIC EXTEND ssrclear + | [ "clear" natural(n) ] -> { tclIPAT (List.init n (fun _ -> IPatAnon Drop)) } +END + +(** The "move" tactic *) + +{ + +(* TODO: review this, in particular the => _ and => [] cases *) +let rec improper_intros = function + | IPatSimpl _ :: ipats -> improper_intros ipats + | (IPatId _ | IPatAnon _ | IPatCase _ | IPatDispatch _) :: _ -> false + | _ -> true (* FIXME *) + +let check_movearg = function + | view, (eqid, _) when view <> [] && eqid <> None -> + CErrors.user_err (Pp.str "incompatible view and equation in move tactic") + | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen -> + CErrors.user_err (Pp.str "incompatible view and occurrence switch in move tactic") + | _, (_, ((dgens, _), _)) when List.length dgens > 1 -> + CErrors.user_err (Pp.str "dependents switch `/' in move tactic") + | _, (eqid, (_, ipats)) when eqid <> None && improper_intros ipats -> + CErrors.user_err (Pp.str "no proper intro pattern for equation in move tactic") + | arg -> arg + +} + +ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY { pr_ssrarg } +| [ ssrarg(arg) ] -> { check_movearg arg } +END + +{ + +let movearg_of_parsed_movearg (v,(eq,(dg,ip))) = + (v,(eq,(ssrdgens_of_parsed_dgens dg,ip))) + +} + +TACTIC EXTEND ssrmove +| [ "move" ssrmovearg(arg) ssrrpat(pat) ] -> + { ssrmovetac (movearg_of_parsed_movearg arg) <*> tclIPAT [pat] } +| [ "move" ssrmovearg(arg) ssrclauses(clauses) ] -> + { tclCLAUSES (ssrmovetac (movearg_of_parsed_movearg arg)) clauses } +| [ "move" ssrrpat(pat) ] -> { tclIPAT [pat] } +| [ "move" ] -> { ssrsmovetac } +END + +{ + +let check_casearg = function +| view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen -> + CErrors.user_err (Pp.str "incompatible view and occurrence switch in dependent case tactic") +| arg -> arg + +} + +ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY { pr_ssrarg } +| [ ssrarg(arg) ] -> { check_casearg arg } +END + +TACTIC EXTEND ssrcase +| [ "case" ssrcasearg(arg) ssrclauses(clauses) ] -> + { tclCLAUSES (ssrcasetac (movearg_of_parsed_movearg arg)) clauses } +| [ "case" ] -> { ssrscasetoptac } +END + +(** The "elim" tactic *) + +TACTIC EXTEND ssrelim +| [ "elim" ssrarg(arg) ssrclauses(clauses) ] -> + { tclCLAUSES (ssrelimtac (movearg_of_parsed_movearg arg)) clauses } +| [ "elim" ] -> { ssrselimtoptac } +END + +(** 6. Backward chaining tactics: apply, exact, congr. *) + +(** The "apply" tactic *) + +{ + +let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt +let pr_ssragen _ _ _ = pr_agen +let pr_ssragens _ _ _ = pr_dgens pr_agen + +} + +ARGUMENT EXTEND ssragen TYPED AS (ssrdocc * ssrterm) PRINTED BY { pr_ssragen } +| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> { mkclr clr, dt } +| [ ssrterm(dt) ] -> { nodocc, dt } +END + +ARGUMENT EXTEND ssragens TYPED AS (ssragen list list * ssrclear) +PRINTED BY { pr_ssragens } +| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] -> + { cons_gen (mkclr clr, dt) agens } +| [ "{" ne_ssrhyp_list(clr) "}" ] -> { [[]], clr} +| [ ssrterm(dt) ssragens(agens) ] -> + { cons_gen (nodocc, dt) agens } +| [ ] -> { [[]], [] } +END + +{ + +let mk_applyarg views agens intros = views, (agens, intros) + +let pr_ssraarg _ _ _ (view, (dgens, ipats)) = + let pri = pr_intros (gens_sep dgens) in + pr_view view ++ pr_dgens pr_agen dgens ++ pri ipats + +} + +ARGUMENT EXTEND ssrapplyarg +TYPED AS (ssrbwdview * (ssragens * ssrintros)) +PRINTED BY { pr_ssraarg } +| [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> + { mk_applyarg [] (cons_gen gen dgens) intros } +| [ ssrclear_ne(clr) ssrintros(intros) ] -> + { mk_applyarg [] ([], clr) intros } +| [ ssrintros_ne(intros) ] -> + { mk_applyarg [] ([], []) intros } +| [ ssrbwdview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> + { mk_applyarg view (cons_gen gen dgens) intros } +| [ ssrbwdview(view) ssrclear(clr) ssrintros(intros) ] -> + { mk_applyarg view ([], clr) intros } + END + +TACTIC EXTEND ssrapply +| [ "apply" ssrapplyarg(arg) ] -> { + let views, (gens_clr, intros) = arg in + inner_ssrapplytac views gens_clr ist <*> tclIPATssr intros } +| [ "apply" ] -> { apply_top_tac } +END + +(** The "exact" tactic *) + +{ + +let mk_exactarg views dgens = mk_applyarg views dgens [] + +} + +ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY { pr_ssraarg } +| [ ":" ssragen(gen) ssragens(dgens) ] -> + { mk_exactarg [] (cons_gen gen dgens) } +| [ ssrbwdview(view) ssrclear(clr) ] -> + { mk_exactarg view ([], clr) } +| [ ssrclear_ne(clr) ] -> + { mk_exactarg [] ([], clr) } +END + +{ + +let vmexacttac pf = + Goal.enter begin fun gl -> + exact_no_check (EConstr.mkCast (pf, _vmcast, Tacmach.New.pf_concl gl)) + end + +} + +TACTIC EXTEND ssrexact +| [ "exact" ssrexactarg(arg) ] -> { + let views, (gens_clr, _) = arg in + V82.tactic (tclBY (V82.of_tactic (inner_ssrapplytac views gens_clr ist))) } +| [ "exact" ] -> { + V82.tactic (Tacticals.tclORELSE (donetac ~-1) (tclBY (V82.of_tactic apply_top_tac))) } +| [ "exact" "<:" lconstr(pf) ] -> { vmexacttac pf } +END + +(** The "congr" tactic *) + +(* type ssrcongrarg = open_constr * (int * constr) *) + +{ + +let pr_ssrcongrarg _ _ _ ((n, f), dgens) = + (if n <= 0 then mt () else str " " ++ int n) ++ + str " " ++ pr_term f ++ pr_dgens pr_gen dgens + +} + +ARGUMENT EXTEND ssrcongrarg TYPED AS ((int * ssrterm) * ssrdgens) + PRINTED BY { pr_ssrcongrarg } +| [ natural(n) constr(c) ssrdgens(dgens) ] -> { (n, mk_term xNoFlag c), dgens } +| [ natural(n) constr(c) ] -> { (n, mk_term xNoFlag c),([[]],[]) } +| [ constr(c) ssrdgens(dgens) ] -> { (0, mk_term xNoFlag c), dgens } +| [ constr(c) ] -> { (0, mk_term xNoFlag c), ([[]],[]) } +END + + + +TACTIC EXTEND ssrcongr +| [ "congr" ssrcongrarg(arg) ] -> +{ let arg, dgens = arg in + V82.tactic begin + match dgens with + | [gens], clr -> Tacticals.tclTHEN (genstac (gens,clr)) (newssrcongrtac arg ist) + | _ -> errorstrm (str"Dependent family abstractions not allowed in congr") + end } +END + +(** 7. Rewriting tactics (rewrite, unlock) *) + +(** Coq rewrite compatibility flag *) + +(** Rewrite clear/occ switches *) + +{ + +let pr_rwocc = function + | None, None -> mt () + | None, occ -> pr_occ occ + | Some clr, _ -> pr_clear_ne clr + +let pr_ssrrwocc _ _ _ = pr_rwocc + +} + +ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY { pr_ssrrwocc } +| [ "{" ssrhyp_list(clr) "}" ] -> { mkclr clr } +| [ "{" ssrocc(occ) "}" ] -> { mkocc occ } +| [ ] -> { noclr } +END + +(** Rewrite rules *) + +{ + +let pr_rwkind = function + | RWred s -> pr_simpl s + | RWdef -> str "/" + | RWeq -> mt () + +let wit_ssrrwkind = add_genarg "ssrrwkind" pr_rwkind + +let pr_rule = function + | RWred s, _ -> pr_simpl s + | RWdef, r-> str "/" ++ pr_term r + | RWeq, r -> pr_term r + +let pr_ssrrule _ _ _ = pr_rule + +let noruleterm loc = mk_term xNoFlag (mkCProp loc) + +} + +ARGUMENT EXTEND ssrrule_ne TYPED AS (ssrrwkind * ssrterm) PRINTED BY { pr_ssrrule } + | [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +GRAMMAR EXTEND Gram + GLOBAL: ssrrule_ne; + ssrrule_ne : [ + [ test_not_ssrslashnum; x = + [ "/"; t = ssrterm -> { RWdef, t } + | t = ssrterm -> { RWeq, t } + | s = ssrsimpl_ne -> { RWred s, noruleterm (Some loc) } + ] -> { x } + | s = ssrsimpl_ne -> { RWred s, noruleterm (Some loc) } + ]]; +END + +ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY { pr_ssrrule } + | [ ssrrule_ne(r) ] -> { r } + | [ ] -> { RWred Nop, noruleterm (Some loc) } +END + +(** Rewrite arguments *) + +{ + +let pr_option f = function None -> mt() | Some x -> f x +let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]") +let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep +let pr_rwarg ((d, m), ((docc, rx), r)) = + pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r + +let pr_ssrrwarg _ _ _ = pr_rwarg + +} + +ARGUMENT EXTEND ssrpattern_squarep +TYPED AS rpattern option PRINTED BY { pr_ssrpattern_squarep } + | [ "[" rpattern(rdx) "]" ] -> { Some rdx } + | [ ] -> { None } +END + +ARGUMENT EXTEND ssrpattern_ne_squarep +TYPED AS rpattern option PRINTED BY { pr_ssrpattern_squarep } + | [ "[" rpattern(rdx) "]" ] -> { Some rdx } +END + + +ARGUMENT EXTEND ssrrwarg + TYPED AS ((ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule)) + PRINTED BY { pr_ssrrwarg } + | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg (R2L, m) (docc, rx) r } + | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *) + { mk_rwarg (R2L, nomult) norwocc (RWdef, t) } + | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg (L2R, m) (docc, rx) r } + | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg norwmult (mkclr clr, rx) r } + | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] -> + { mk_rwarg norwmult (mkclr clr, None) r } + | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg norwmult (mkocc occ, rx) r } + | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg norwmult (nodocc, rx) r } + | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> + { mk_rwarg norwmult (noclr, rx) r } + | [ ssrrule_ne(r) ] -> + { mk_rwarg norwmult norwocc r } +END + +TACTIC EXTEND ssrinstofruleL2R +| [ "ssrinstancesofruleL2R" ssrterm(arg) ] -> { V82.tactic (ssrinstancesofrule ist L2R arg) } +END +TACTIC EXTEND ssrinstofruleR2L +| [ "ssrinstancesofruleR2L" ssrterm(arg) ] -> { V82.tactic (ssrinstancesofrule ist R2L arg) } +END + +(** Rewrite argument sequence *) + +(* type ssrrwargs = ssrrwarg list *) + +{ + +let pr_ssrrwargs _ _ _ rwargs = pr_list spc pr_rwarg rwargs + +} + +ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list PRINTED BY { pr_ssrrwargs } + | [ "YouShouldNotTypeThis" ] -> { anomaly "Grammar placeholder match" } +END + +{ + +let ssr_rw_syntax = Summary.ref ~name:"SSR:rewrite" true + +let _ = + Goptions.declare_bool_option + { Goptions.optname = "ssreflect rewrite"; + Goptions.optkey = ["SsrRewrite"]; + Goptions.optread = (fun _ -> !ssr_rw_syntax); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> ssr_rw_syntax := b) } + +let lbrace = Char.chr 123 +(** Workaround to a limitation of coqpp *) + +let test_ssr_rw_syntax = + let test strm = + if not !ssr_rw_syntax then raise Stream.Failure else + if is_ssr_loaded () then () else + match Util.stream_nth 0 strm with + | Tok.KEYWORD key when List.mem key.[0] [lbrace; '['; '/'] -> () + | _ -> raise Stream.Failure in + Gram.Entry.of_parser "test_ssr_rw_syntax" test + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssrrwargs; + ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> { a } ]]; +END + +(** The "rewrite" tactic *) + +TACTIC EXTEND ssrrewrite + | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] -> + { tclCLAUSES (old_tac (ssrrewritetac ist args)) clauses } +END + +(** The "unlock" tactic *) + +{ + +let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t +let pr_ssrunlockarg _ _ _ = pr_unlockarg + +} + +ARGUMENT EXTEND ssrunlockarg TYPED AS (ssrocc * ssrterm) + PRINTED BY { pr_ssrunlockarg } + | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> { occ, t } + | [ ssrterm(t) ] -> { None, t } +END + +{ + +let pr_ssrunlockargs _ _ _ args = pr_list spc pr_unlockarg args + +} + +ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list + PRINTED BY { pr_ssrunlockargs } + | [ ssrunlockarg_list(args) ] -> { args } +END + +TACTIC EXTEND ssrunlock + | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] -> + { tclCLAUSES (old_tac (unlocktac ist args)) clauses } +END + +(** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *) + + +TACTIC EXTEND ssrpose +| [ "pose" ssrfixfwd(ffwd) ] -> { V82.tactic (ssrposetac ffwd) } +| [ "pose" ssrcofixfwd(ffwd) ] -> { V82.tactic (ssrposetac ffwd) } +| [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> { V82.tactic (ssrposetac (id, fwd)) } +END + +(** The "set" tactic *) + +(* type ssrsetfwd = ssrfwd * ssrdocc *) + +TACTIC EXTEND ssrset +| [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] -> + { tclCLAUSES (old_tac (ssrsettac id fwd)) clauses } +END + +(** The "have" tactic *) + +(* type ssrhavefwd = ssrfwd * ssrhint *) + + +(* Pltac. *) + +(* The standard TACTIC EXTEND does not work for abstract *) +GRAMMAR EXTEND Gram + GLOBAL: tactic_expr; + tactic_expr: LEVEL "3" + [ RIGHTA [ IDENT "abstract"; gens = ssrdgens -> + { ssrtac_expr ~loc "abstract" + [Tacexpr.TacGeneric (Genarg.in_gen (Genarg.rawwit wit_ssrdgens) gens)] } ]]; +END +TACTIC EXTEND ssrabstract +| [ "abstract" ssrdgens(gens) ] -> { + if List.length (fst gens) <> 1 then + errorstrm (str"dependents switches '/' not allowed here"); + Ssripats.ssrabstract (ssrdgens_of_parsed_dgens gens) } +END + +TACTIC EXTEND ssrhave +| [ "have" ssrhavefwdwbinders(fwd) ] -> + { V82.tactic (havetac ist fwd false false) } +END + +TACTIC EXTEND ssrhavesuff +| [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + { V82.tactic (havetac ist (false,(pats,fwd)) true false) } +END + +TACTIC EXTEND ssrhavesuffices +| [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + { V82.tactic (havetac ist (false,(pats,fwd)) true false) } +END + +TACTIC EXTEND ssrsuffhave +| [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + { V82.tactic (havetac ist (false,(pats,fwd)) true true) } +END + +TACTIC EXTEND ssrsufficeshave +| [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + { V82.tactic (havetac ist (false,(pats,fwd)) true true) } +END + +(** The "suffice" tactic *) + +{ + +let pr_ssrsufffwdwbinders _ _ prt (hpats, (fwd, hint)) = + pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint + +} + +ARGUMENT EXTEND ssrsufffwd + TYPED AS (ssrhpats * (ssrfwd * ssrhint)) PRINTED BY { pr_ssrsufffwdwbinders } +| [ ssrhpats(pats) ssrbinder_list(bs) ":" ast_closure_lterm(t) ssrhint(hint) ] -> + { let ((clr, pats), binders), simpl = pats in + let allbs = intro_id_to_binder binders @ bs in + let allbinders = binders @ List.flatten (binder_to_intro_id bs) in + let fwd = mkFwdHint ":" t in + (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) } +END + + +TACTIC EXTEND ssrsuff +| [ "suff" ssrsufffwd(fwd) ] -> { V82.tactic (sufftac ist fwd) } +END + +TACTIC EXTEND ssrsuffices +| [ "suffices" ssrsufffwd(fwd) ] -> { V82.tactic (sufftac ist fwd) } +END + +(** The "wlog" (Without Loss Of Generality) tactic *) + +(* type ssrwlogfwd = ssrwgen list * ssrfwd *) + +{ + +let pr_ssrwlogfwd _ _ _ (gens, t) = + str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t + +} + +ARGUMENT EXTEND ssrwlogfwd TYPED AS (ssrwgen list * ssrfwd) + PRINTED BY { pr_ssrwlogfwd } +| [ ":" ssrwgen_list(gens) "/" ast_closure_lterm(t) ] -> { gens, mkFwdHint "/" t} +END + + +TACTIC EXTEND ssrwlog +| [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { V82.tactic (wlogtac ist pats fwd hint false `NoGen) } +END + +TACTIC EXTEND ssrwlogs +| [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } +END + +TACTIC EXTEND ssrwlogss +| [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> + { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } +END + +TACTIC EXTEND ssrwithoutloss +| [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { V82.tactic (wlogtac ist pats fwd hint false `NoGen) } +END + +TACTIC EXTEND ssrwithoutlosss +| [ "without" "loss" "suff" + ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } +END + +TACTIC EXTEND ssrwithoutlossss +| [ "without" "loss" "suffices" + ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> + { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } +END + +{ + +(* Generally have *) +let pr_idcomma _ _ _ = function + | None -> mt() + | Some None -> str"_, " + | Some (Some id) -> pr_id id ++ str", " + +} + +ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY { pr_idcomma } + | [ ] -> { None } +END + +{ + +let accept_idcomma strm = + match stream_nth 0 strm with + | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm + | _ -> raise Stream.Failure + +let test_idcomma = Gram.Entry.of_parser "test_idcomma" accept_idcomma + +} + +GRAMMAR EXTEND Gram + GLOBAL: ssr_idcomma; + ssr_idcomma: [ [ test_idcomma; + ip = [ id = IDENT -> { Some (Id.of_string id) } | "_" -> { None } ]; "," -> + { Some ip } + ] ]; +END + +{ + +let augment_preclr clr1 (((clr0, x),y),z) = (((clr1 @ clr0, x),y),z) + +} + +TACTIC EXTEND ssrgenhave +| [ "gen" "have" ssrclear(clr) + ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { let pats = augment_preclr clr pats in + V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) } +END + +TACTIC EXTEND ssrgenhave2 +| [ "generally" "have" ssrclear(clr) + ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + { let pats = augment_preclr clr pats in + V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) } +END + +{ + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +let () = CLexer.set_keyword_state frozen_lexer ;; + +} + +(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4 deleted file mode 100644 index 02a5d08507..0000000000 --- a/plugins/ssr/ssrvernac.ml4 +++ /dev/null @@ -1,625 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* is then ... else ... *) -(* if is [in ..] return ... then ... else ... *) -(* let: := in ... *) -(* let: [in ...] := return ... in ... *) -(* The scope of a top-level 'as' in the pattern extends over the *) -(* 'return' type (dependent if/let). *) -(* Note that the optional "in ..." appears next to the *) -(* rather than the in then "let:" syntax. The alternative *) -(* would lead to ambiguities in, e.g., *) -(* let: p1 := (*v---INNER LET:---v *) *) -(* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *) -(* in b (*^--ALTERNATIVE INNER LET--------^ *) *) - -(* Caveat : There is no pretty-printing support, since this would *) -(* require a modification to the Coq kernel (adding a new match *) -(* display style -- why aren't these strings?); also, the v8.1 *) -(* pretty-printer only allows extension hooks for printing *) -(* integer or string literals. *) -(* Also note that in the v8 grammar "is" needs to be a keyword; *) -(* as this can't be done from an ML extension file, the new *) -(* syntax will only work when ssreflect.v is imported. *) - -let no_ct = None, None and no_rt = None in -let aliasvar = function - | [[{ CAst.v = CPatAlias (_, na); loc }]] -> Some na - | _ -> None in -let mk_cnotype mp = aliasvar mp, None in -let mk_ctype mp t = aliasvar mp, Some t in -let mk_rtype t = Some t in -let mk_dthen ?loc (mp, ct, rt) c = (CAst.make ?loc (mp, c)), ct, rt in -let mk_let ?loc rt ct mp c1 = - CAst.make ?loc @@ CCases (LetPatternStyle, rt, ct, [CAst.make ?loc (mp, c1)]) in -let mk_pat c (na, t) = (c, na, t) in -GEXTEND Gram - GLOBAL: binder_constr; - ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> mk_rtype t ]]; - ssr_mpat: [[ p = pattern -> [[p]] ]]; - ssr_dpat: [ - [ mp = ssr_mpat; "in"; t = pattern; rt = ssr_rtype -> mp, mk_ctype mp t, rt - | mp = ssr_mpat; rt = ssr_rtype -> mp, mk_cnotype mp, rt - | mp = ssr_mpat -> mp, no_ct, no_rt - ] ]; - ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> mk_dthen ~loc:!@loc dp c ]]; - ssr_elsepat: [[ "else" -> [[CAst.make ~loc:!@loc @@ CPatAtom None]] ]]; - ssr_else: [[ mp = ssr_elsepat; c = lconstr -> CAst.make ~loc:!@loc (mp, c) ]]; - binder_constr: [ - [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> - let b1, ct, rt = db1 in CAst.make ~loc:!@loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) - | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> - let b1, ct, rt = db1 in - let b1, b2 = let open CAst in - let {loc=l1; v=(p1, r1)}, {loc=l2; v=(p2, r2)} = b1, b2 in - (make ?loc:l1 (p1, r2), make ?loc:l2 (p2, r1)) - in - CAst.make ~loc:!@loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) - | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr -> - mk_let ~loc:!@loc no_rt [mk_pat c no_ct] mp c1 - | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; - rt = ssr_rtype; "in"; c1 = lconstr -> - mk_let ~loc:!@loc rt [mk_pat c (mk_cnotype mp)] mp c1 - | "let"; ":"; mp = ssr_mpat; "in"; t = pattern; ":="; c = lconstr; - rt = ssr_rtype; "in"; c1 = lconstr -> - mk_let ~loc:!@loc rt [mk_pat c (mk_ctype mp t)] mp c1 - ] ]; -END - -GEXTEND Gram - GLOBAL: closed_binder; - closed_binder: [ - [ ["of" | "&"]; c = operconstr LEVEL "99" -> - [CLocalAssum ([CAst.make ~loc:!@loc Anonymous], Default Explicit, c)] - ] ]; -END -(* }}} *) - -(** Vernacular commands: Prenex Implicits and Search *)(* {{{ **********************) - -(* This should really be implemented as an extension to the implicit *) -(* arguments feature, but unfortuately that API is sealed. The current *) -(* workaround uses a combination of notations that works reasonably, *) -(* with the following caveats: *) -(* - The pretty-printing always elides prenex implicits, even when *) -(* they are obviously needed. *) -(* - Prenex Implicits are NEVER exported from a module, because this *) -(* would lead to faulty pretty-printing and scoping errors. *) -(* - The command "Import Prenex Implicits" can be used to reassert *) -(* Prenex Implicits for all the visible constants that had been *) -(* declared as Prenex Implicits. *) - -let declare_one_prenex_implicit locality f = - let fref = - try Smartlocate.global_with_alias f - with _ -> errorstrm (pr_qualid f ++ str " is not declared") in - let rec loop = function - | a :: args' when Impargs.is_status_implicit a -> - (ExplByName (Impargs.name_of_implicit a), (true, true, true)) :: loop args' - | args' when List.exists Impargs.is_status_implicit args' -> - errorstrm (str "Expected prenex implicits for " ++ pr_qualid f) - | _ -> [] in - let impls = - match Impargs.implicits_of_global fref with - | [cond,impls] -> impls - | [] -> errorstrm (str "Expected some implicits for " ++ pr_qualid f) - | _ -> errorstrm (str "Multiple implicits not supported") in - match loop impls with - | [] -> - errorstrm (str "Expected some implicits for " ++ pr_qualid f) - | impls -> - Impargs.declare_manual_implicits locality fref ~enriching:false [impls] - -VERNAC COMMAND FUNCTIONAL EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF - | [ "Prenex" "Implicits" ne_global_list(fl) ] - -> [ fun ~atts ~st -> - let open Vernacinterp in - let locality = Locality.make_section_locality atts.locality in - List.iter (declare_one_prenex_implicit locality) fl; - st - ] -END - -(* Vernac grammar visibility patch *) - -GEXTEND Gram - GLOBAL: gallina_ext; - gallina_ext: - [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" -> - Vernacexpr.VernacUnsetOption (false, ["Printing"; "Implicit"; "Defensive"]) - ] ] - ; -END - -(** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *) - -(* Main prefilter *) - -type raw_glob_search_about_item = - | RGlobSearchSubPattern of constr_expr - | RGlobSearchString of Loc.t * string * string option - -let pr_search_item = function - | RGlobSearchString (_,s,_) -> str s - | RGlobSearchSubPattern p -> pr_constr_expr p - -let wit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item - -let pr_ssr_search_item _ _ _ = pr_search_item - -(* Workaround the notation API that can only print notations *) - -let is_ident s = try CLexer.check_ident s; true with _ -> false - -let is_ident_part s = is_ident ("H" ^ s) - -let interp_search_notation ?loc tag okey = - let err msg = CErrors.user_err ?loc ~hdr:"interp_search_notation" msg in - let mk_pntn s for_key = - let n = String.length s in - let s' = Bytes.make (n + 2) ' ' in - let rec loop i i' = - if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else - let j = try String.index_from s (i + 1) ' ' with _ -> n in - let m = j - i in - if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then - (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1)) - else if for_key && is_ident (String.sub s i m) then - (Bytes.set s' i' '_'; loop (j + 1) (i' + 2)) - else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in - loop 0 1 in - let trim_ntn (pntn, m) = (InConstrEntrySomeLevel,Bytes.sub_string pntn 1 (max 0 m)) in - let pr_ntn ntn = str "(" ++ Notation.pr_notation ntn ++ str ")" in - let pr_and_list pr = function - | [x] -> pr x - | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x - | [] -> mt () in - let pr_sc sc = str (if sc = "" then "independently" else sc) in - let pr_scs = function - | [""] -> pr_sc "" - | scs -> str "in " ++ pr_and_list pr_sc scs in - let generator, pr_tag_sc = - let ign _ = mt () in match okey with - | Some key -> - let sc = Notation.find_delimiters_scope ?loc key in - let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in - Notation.pr_scope ign sc, pr_sc - | None -> Notation.pr_scopes ign, ign in - let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in - let ptag, ttag = - let ptag, m = mk_pntn tag false in - if m <= 0 then err (str "empty notation fragment"); - ptag, trim_ntn (ptag, m) in - let last = ref "" and last_sc = ref "" in - let scs = ref [] and ntns = ref [] in - let push_sc sc = match !scs with - | "" :: scs' -> scs := "" :: sc :: scs' - | scs' -> scs := sc :: scs' in - let get s _ _ = match !last with - | "Scope " -> last_sc := s; last := "" - | "Lonely notation" -> last_sc := ""; last := "" - | "\"" -> - let pntn, m = mk_pntn s true in - if String.string_contains ~where:(Bytes.to_string pntn) ~what:(Bytes.to_string ptag) then begin - let ntn = trim_ntn (pntn, m) in - match !ntns with - | [] -> ntns := [ntn]; scs := [!last_sc] - | ntn' :: _ when ntn' = ntn -> push_sc !last_sc - | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc] - | _ :: ntns' when List.mem ntn ntns' -> () - | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns' - end; - last := "" - | _ -> last := s in - pp_with (Format.make_formatter get (fun _ -> ())) generator; - let ntn = match !ntns with - | [] -> - err (hov 0 (qtag "in" ++ str "does not occur in any notation")) - | ntn :: ntns' when ntn = ttag -> - if ntns' <> [] then begin - let pr_ntns' = pr_and_list pr_ntn ntns' in - Feedback.msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns')) - end; ntn - | [ntn] -> - Feedback.msg_info (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn - | ntns' -> - let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in - err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in - let (nvars, body), ((_, pat), osc) = match !scs with - | [sc] -> Notation.interp_notation ?loc ntn (None, [sc]) - | scs' -> - try Notation.interp_notation ?loc ntn (None, []) with _ -> - let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in - err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in - let sc = Option.default "" osc in - let _ = - let m_sc = - if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in - let ntn_pat = trim_ntn (mk_pntn pat false) in - let rbody = glob_constr_of_notation_constr ?loc body in - let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in - let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in - Feedback.msg_info (hov 0 m) in - if List.length !scs > 1 then - let scs' = List.remove (=) sc !scs in - let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in - Feedback.msg_warning (hov 4 w) - else if String.string_contains ~where:(snd ntn) ~what:" .. " then - err (pr_ntn ntn ++ str " is an n-ary notation"); - let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in - let rec sub () = function - | NVar x when List.mem_assoc x nvars -> DAst.make ?loc @@ GPatVar (FirstOrderPatVar x) - | c -> - glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), None, x) sub () c in - let _, npat = Patternops.pattern_of_glob_constr (sub () body) in - Search.GlobSearchSubPattern npat - -ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem - PRINTED BY pr_ssr_search_item - | [ string(s) ] -> [ RGlobSearchString (loc,s,None) ] - | [ string(s) "%" preident(key) ] -> [ RGlobSearchString (loc,s,Some key) ] - | [ constr_pattern(p) ] -> [ RGlobSearchSubPattern p ] -END - -let pr_ssr_search_arg _ _ _ = - let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item p in - pr_list spc pr_item - -ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list - PRINTED BY pr_ssr_search_arg - | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> [ (false, p) :: a ] - | [ ssr_search_item(p) ssr_search_arg(a) ] -> [ (true, p) :: a ] - | [ ] -> [ [] ] -END - -(* Main type conclusion pattern filter *) - -let rec splay_search_pattern na = function - | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp - | Pattern.PLetIn (_, _, _, bp) -> splay_search_pattern na bp - | Pattern.PRef hr -> hr, na - | _ -> CErrors.user_err (Pp.str "no head constant in head search pattern") - -let push_rels_assum l e = - let l = List.map (fun (n,t) -> n, EConstr.Unsafe.to_constr t) l in - push_rels_assum l e - -let coerce_search_pattern_to_sort hpat = - let env = Global.env () in - let sigma = Evd.(from_env env) in - let mkPApp fp n_imps args = - let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in - Pattern.PApp (fp, args') in - let hr, na = splay_search_pattern 0 hpat in - let dc, ht = - let hr, _ = Global.type_of_global_in_context (Global.env ()) hr (** FIXME *) in - Reductionops.splay_prod env sigma (EConstr.of_constr hr) in - let np = List.length dc in - if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else - let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in - let warn () = - Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++ - pr_constr_pattern_env env sigma hpat') in - if EConstr.isSort sigma ht then begin warn (); true, hpat' end else - let filter_head, coe_path = - try - let _, cp = - Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in - warn (); - true, cp - with _ -> false, [] in - let coerce hp coe_index = - let coe_ref = coe_index.Classops.coe_value in - try - let n_imps = Option.get (Classops.hide_coercion coe_ref) in - mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] - with Not_found | Option.IsNone -> - errorstrm (str "need explicit coercion " ++ pr_global coe_ref ++ spc () - ++ str "to interpret head search pattern as type") in - filter_head, List.fold_left coerce hpat' coe_path - -let interp_head_pat hpat = - let filter_head, p = coerce_search_pattern_to_sort hpat in - let rec loop c = match CoqConstr.kind c with - | Cast (c', _, _) -> loop c' - | Prod (_, _, c') -> loop c' - | LetIn (_, _, _, c') -> loop c' - | _ -> - let env = Global.env () in - let sigma = Evd.from_env env in - Constr_matching.is_matching env sigma p (EConstr.of_constr c) in - filter_head, loop - -let all_true _ = true - -let rec interp_search_about args accu = match args with -| [] -> accu -| (flag, arg) :: rem -> - fun gr env typ -> - let ans = Search.search_about_filter arg gr env typ in - (if flag then ans else not ans) && interp_search_about rem accu gr env typ - -let interp_search_arg arg = - let arg = List.map (fun (x,arg) -> x, match arg with - | RGlobSearchString (loc,s,key) -> - if is_ident_part s then Search.GlobSearchString s else - interp_search_notation ~loc s key - | RGlobSearchSubPattern p -> - try - let env = Global.env () in - let _, p = Constrintern.intern_constr_pattern env (Evd.from_env env) p in - Search.GlobSearchSubPattern p - with e -> let e = CErrors.push e in iraise (ExplainErr.process_vernac_interp_error e)) arg in - let hpat, a1 = match arg with - | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' - | (true, Search.GlobSearchSubPattern p) :: a' -> - let filter_head, p = interp_head_pat p in - if filter_head then p, a' else all_true, arg - | _ -> all_true, arg in - let is_string = - function (_, Search.GlobSearchString _) -> true | _ -> false in - let a2, a3 = List.partition is_string a1 in - interp_search_about (a2 @ a3) (fun gr env typ -> hpat typ) - -(* Module path postfilter *) - -let pr_modloc (b, m) = if b then str "-" ++ pr_qualid m else pr_qualid m - -let wit_ssrmodloc = add_genarg "ssrmodloc" pr_modloc - -let pr_ssr_modlocs _ _ _ ml = - if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml - -ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY pr_ssr_modlocs - | [ ] -> [ [] ] -END - -GEXTEND Gram - GLOBAL: ssr_modlocs; - modloc: [[ "-"; m = global -> true, m | m = global -> false, m]]; - ssr_modlocs: [[ "in"; ml = LIST1 modloc -> ml ]]; -END - -let interp_modloc mr = - let interp_mod (_, qid) = - try Nametab.full_name_module qid with Not_found -> - CErrors.user_err ?loc:qid.CAst.loc (str "No Module " ++ pr_qualid qid) in - let mr_out, mr_in = List.partition fst mr in - let interp_bmod b = function - | [] -> fun _ _ _ -> true - | rmods -> Search.module_filter (List.map interp_mod rmods, b) in - let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in - fun gr env typ -> is_in gr env typ && is_out gr env typ - -(* The unified, extended vernacular "Search" command *) - -let ssrdisplaysearch gr env t = - let pr_res = pr_global gr ++ spc () ++ str " " ++ pr_lconstr_env env Evd.empty t in - Feedback.msg_info (hov 2 pr_res ++ fnl ()) - -VERNAC COMMAND EXTEND SsrSearchPattern CLASSIFIED AS QUERY -| [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] -> - [ let hpat = interp_search_arg a in - let in_mod = interp_modloc mr in - let post_filter gr env typ = in_mod gr env typ && hpat gr env typ in - let display gr env typ = - if post_filter gr env typ then ssrdisplaysearch gr env typ - in - Search.generic_search None display ] -END - -(* }}} *) - -(** View hint database and View application. *)(* {{{ ******************************) - -(* There are three databases of lemmas used to mediate the application *) -(* of reflection lemmas: one for forward chaining, one for backward *) -(* chaining, and one for secondary backward chaining. *) - -(* View hints *) - -let pr_raw_ssrhintref prc _ _ = let open CAst in function - | { v = CAppExpl ((None, r,x), args) } when isCHoles args -> - prc (CAst.make @@ CRef (r,x)) ++ str "|" ++ int (List.length args) - | { v = CApp ((_, { v = CRef _ }), _) } as c -> prc c - | { v = CApp ((_, c), args) } when isCxHoles args -> - prc c ++ str "|" ++ int (List.length args) - | c -> prc c - -let pr_rawhintref c = - let _, env = Pfedit.get_current_context () in - match DAst.get c with - | GApp (f, args) when isRHoles args -> - pr_glob_constr_env env f ++ str "|" ++ int (List.length args) - | _ -> pr_glob_constr_env env c - -let pr_glob_ssrhintref _ _ _ (c, _) = pr_rawhintref c - -let pr_ssrhintref prc _ _ = prc - -let mkhintref ?loc c n = match c.CAst.v with - | CRef (r,x) -> CAst.make ?loc @@ CAppExpl ((None, r, x), mkCHoles ?loc n) - | _ -> mkAppC (c, mkCHoles ?loc n) - -ARGUMENT EXTEND ssrhintref - TYPED AS constr - PRINTED BY pr_ssrhintref - RAW_PRINTED BY pr_raw_ssrhintref - GLOB_PRINTED BY pr_glob_ssrhintref - | [ constr(c) ] -> [ c ] - | [ constr(c) "|" natural(n) ] -> [ mkhintref ~loc c n ] -END - -(* View purpose *) - -let pr_viewpos = function - | Some Ssrview.AdaptorDb.Forward -> str " for move/" - | Some Ssrview.AdaptorDb.Backward -> str " for apply/" - | Some Ssrview.AdaptorDb.Equivalence -> str " for apply//" - | None -> mt () - -let pr_ssrviewpos _ _ _ = pr_viewpos - -ARGUMENT EXTEND ssrviewpos PRINTED BY pr_ssrviewpos - | [ "for" "move" "/" ] -> [ Some Ssrview.AdaptorDb.Forward ] - | [ "for" "apply" "/" ] -> [ Some Ssrview.AdaptorDb.Backward ] - | [ "for" "apply" "/" "/" ] -> [ Some Ssrview.AdaptorDb.Equivalence ] - | [ "for" "apply" "//" ] -> [ Some Ssrview.AdaptorDb.Equivalence ] - | [ ] -> [ None ] -END - -let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc () - -ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY pr_ssrviewposspc - | [ ssrviewpos(i) ] -> [ i ] -END - -let print_view_hints kind l = - let pp_viewname = str "Hint View" ++ pr_viewpos (Some kind) ++ str " " in - let pp_hints = pr_list spc pr_rawhintref l in - Feedback.msg_info (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ()) - -VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY -| [ "Print" "Hint" "View" ssrviewpos(i) ] -> - [ match i with - | Some k -> print_view_hints k (Ssrview.AdaptorDb.get k) - | None -> - List.iter (fun k -> print_view_hints k (Ssrview.AdaptorDb.get k)) - [ Ssrview.AdaptorDb.Forward; - Ssrview.AdaptorDb.Backward; - Ssrview.AdaptorDb.Equivalence ] - ] -END - -let glob_view_hints lvh = - List.map (Constrintern.intern_constr (Global.env ()) (Evd.from_env (Global.env ()))) lvh - -VERNAC COMMAND EXTEND HintView CLASSIFIED AS SIDEFF - | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] -> - [ let hints = glob_view_hints lvh in - match n with - | None -> - Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Forward hints; - Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Backward hints - | Some k -> - Ssrview.AdaptorDb.declare k hints ] -END - -(* }}} *) - -(** Canonical Structure alias *) - -GEXTEND Gram - GLOBAL: gallina_ext; - - gallina_ext: - (* Canonical structure *) - [[ IDENT "Canonical"; qid = Constr.global -> - Vernacexpr.VernacCanonical (CAst.make @@ AN qid) - | IDENT "Canonical"; ntn = Prim.by_notation -> - Vernacexpr.VernacCanonical (CAst.make @@ ByNotation ntn) - | IDENT "Canonical"; qid = Constr.global; - d = G_vernac.def_body -> - let s = coerce_reference_to_id qid in - Vernacexpr.VernacDefinition - ((Decl_kinds.NoDischarge,Decl_kinds.CanonicalStructure), - ((CAst.make (Name s)),None), d) - ]]; -END - -(** Keyword compatibility fixes. *) - -(* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *) -(* identifiers used as keywords. This is incompatible with ssreflect.v *) -(* which makes "by" and "of" true keywords, because of technicalities *) -(* in the internal lexer-parser API of Coq. We patch this here by *) -(* adding new parsing rules that recognize the new keywords. *) -(* To make matters worse, the Coq grammar for tactics fails to *) -(* export the non-terminals we need to patch. Fortunately, the CamlP5 *) -(* API provides a backdoor access (with loads of Obj.magic trickery). *) - -(* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *) -(* longer and thus comment out. Such comments are marked with v8.3 *) - -open Pltac - -GEXTEND Gram - GLOBAL: hypident; - hypident: [ - [ "("; IDENT "type"; "of"; id = Prim.identref; ")" -> id, Locus.InHypTypeOnly - | "("; IDENT "value"; "of"; id = Prim.identref; ")" -> id, Locus.InHypValueOnly - ] ]; -END - -GEXTEND Gram - GLOBAL: hloc; -hloc: [ - [ "in"; "("; "Type"; "of"; id = ident; ")" -> - Tacexpr.HypLocation (CAst.make id, Locus.InHypTypeOnly) - | "in"; "("; IDENT "Value"; "of"; id = ident; ")" -> - Tacexpr.HypLocation (CAst.make id, Locus.InHypValueOnly) - ] ]; -END - -GEXTEND Gram - GLOBAL: constr_eval; - constr_eval: [ - [ IDENT "type"; "of"; c = Constr.constr -> Genredexpr.ConstrTypeOf c ] - ]; -END - -(* We wipe out all the keywords generated by the grammar rules we defined. *) -(* The user is supposed to Require Import ssreflect or Require ssreflect *) -(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) -(* consequence the extended ssreflect grammar. *) -let () = CLexer.set_keyword_state frozen_lexer ;; - -(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/plugins/ssr/ssrvernac.mlg b/plugins/ssr/ssrvernac.mlg new file mode 100644 index 0000000000..876751911b --- /dev/null +++ b/plugins/ssr/ssrvernac.mlg @@ -0,0 +1,675 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* is then ... else ... *) +(* if is [in ..] return ... then ... else ... *) +(* let: := in ... *) +(* let: [in ...] := return ... in ... *) +(* The scope of a top-level 'as' in the pattern extends over the *) +(* 'return' type (dependent if/let). *) +(* Note that the optional "in ..." appears next to the *) +(* rather than the in then "let:" syntax. The alternative *) +(* would lead to ambiguities in, e.g., *) +(* let: p1 := (*v---INNER LET:---v *) *) +(* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *) +(* in b (*^--ALTERNATIVE INNER LET--------^ *) *) + +(* Caveat : There is no pretty-printing support, since this would *) +(* require a modification to the Coq kernel (adding a new match *) +(* display style -- why aren't these strings?); also, the v8.1 *) +(* pretty-printer only allows extension hooks for printing *) +(* integer or string literals. *) +(* Also note that in the v8 grammar "is" needs to be a keyword; *) +(* as this can't be done from an ML extension file, the new *) +(* syntax will only work when ssreflect.v is imported. *) + +let no_ct = None, None and no_rt = None +let aliasvar = function + | [[{ CAst.v = CPatAlias (_, na); loc }]] -> Some na + | _ -> None +let mk_cnotype mp = aliasvar mp, None +let mk_ctype mp t = aliasvar mp, Some t +let mk_rtype t = Some t +let mk_dthen ?loc (mp, ct, rt) c = (CAst.make ?loc (mp, c)), ct, rt +let mk_let ?loc rt ct mp c1 = + CAst.make ?loc @@ CCases (LetPatternStyle, rt, ct, [CAst.make ?loc (mp, c1)]) +let mk_pat c (na, t) = (c, na, t) + +} + +GRAMMAR EXTEND Gram + GLOBAL: binder_constr; + ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> { mk_rtype t } ]]; + ssr_mpat: [[ p = pattern -> { [[p]] } ]]; + ssr_dpat: [ + [ mp = ssr_mpat; "in"; t = pattern; rt = ssr_rtype -> { mp, mk_ctype mp t, rt } + | mp = ssr_mpat; rt = ssr_rtype -> { mp, mk_cnotype mp, rt } + | mp = ssr_mpat -> { mp, no_ct, no_rt } + ] ]; + ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> { mk_dthen ~loc dp c } ]]; + ssr_elsepat: [[ "else" -> { [[CAst.make ~loc @@ CPatAtom None]] } ]]; + ssr_else: [[ mp = ssr_elsepat; c = lconstr -> { CAst.make ~loc (mp, c) } ]]; + binder_constr: [ + [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> + { let b1, ct, rt = db1 in CAst.make ~loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) } + | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> + { let b1, ct, rt = db1 in + let b1, b2 = let open CAst in + let {loc=l1; v=(p1, r1)}, {loc=l2; v=(p2, r2)} = b1, b2 in + (make ?loc:l1 (p1, r2), make ?loc:l2 (p2, r1)) + in + CAst.make ~loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) } + | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr -> + { mk_let ~loc no_rt [mk_pat c no_ct] mp c1 } + | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; + rt = ssr_rtype; "in"; c1 = lconstr -> + { mk_let ~loc rt [mk_pat c (mk_cnotype mp)] mp c1 } + | "let"; ":"; mp = ssr_mpat; "in"; t = pattern; ":="; c = lconstr; + rt = ssr_rtype; "in"; c1 = lconstr -> + { mk_let ~loc rt [mk_pat c (mk_ctype mp t)] mp c1 } + ] ]; +END + +GRAMMAR EXTEND Gram + GLOBAL: closed_binder; + closed_binder: [ + [ ["of" -> { () } | "&" -> { () } ]; c = operconstr LEVEL "99" -> + { [CLocalAssum ([CAst.make ~loc Anonymous], Default Explicit, c)] } + ] ]; +END + +(** Vernacular commands: Prenex Implicits and Search *)(***********************) + +(* This should really be implemented as an extension to the implicit *) +(* arguments feature, but unfortuately that API is sealed. The current *) +(* workaround uses a combination of notations that works reasonably, *) +(* with the following caveats: *) +(* - The pretty-printing always elides prenex implicits, even when *) +(* they are obviously needed. *) +(* - Prenex Implicits are NEVER exported from a module, because this *) +(* would lead to faulty pretty-printing and scoping errors. *) +(* - The command "Import Prenex Implicits" can be used to reassert *) +(* Prenex Implicits for all the visible constants that had been *) +(* declared as Prenex Implicits. *) + +{ + +let declare_one_prenex_implicit locality f = + let fref = + try Smartlocate.global_with_alias f + with _ -> errorstrm (pr_qualid f ++ str " is not declared") in + let rec loop = function + | a :: args' when Impargs.is_status_implicit a -> + (ExplByName (Impargs.name_of_implicit a), (true, true, true)) :: loop args' + | args' when List.exists Impargs.is_status_implicit args' -> + errorstrm (str "Expected prenex implicits for " ++ pr_qualid f) + | _ -> [] in + let impls = + match Impargs.implicits_of_global fref with + | [cond,impls] -> impls + | [] -> errorstrm (str "Expected some implicits for " ++ pr_qualid f) + | _ -> errorstrm (str "Multiple implicits not supported") in + match loop impls with + | [] -> + errorstrm (str "Expected some implicits for " ++ pr_qualid f) + | impls -> + Impargs.declare_manual_implicits locality fref ~enriching:false [impls] + +} + +VERNAC COMMAND EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF + | [ "Prenex" "Implicits" ne_global_list(fl) ] + -> { + let open Vernacinterp in + let locality = Locality.make_section_locality atts.locality in + List.iter (declare_one_prenex_implicit locality) fl; + } +END + +(* Vernac grammar visibility patch *) + +GRAMMAR EXTEND Gram + GLOBAL: gallina_ext; + gallina_ext: + [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" -> + { Vernacexpr.VernacUnsetOption (false, ["Printing"; "Implicit"; "Defensive"]) } + ] ] + ; +END + +(** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *) + +(* Main prefilter *) + +{ + +type raw_glob_search_about_item = + | RGlobSearchSubPattern of constr_expr + | RGlobSearchString of Loc.t * string * string option + +let pr_search_item = function + | RGlobSearchString (_,s,_) -> str s + | RGlobSearchSubPattern p -> pr_constr_expr p + +let wit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item + +let pr_ssr_search_item _ _ _ = pr_search_item + +(* Workaround the notation API that can only print notations *) + +let is_ident s = try CLexer.check_ident s; true with _ -> false + +let is_ident_part s = is_ident ("H" ^ s) + +let interp_search_notation ?loc tag okey = + let err msg = CErrors.user_err ?loc ~hdr:"interp_search_notation" msg in + let mk_pntn s for_key = + let n = String.length s in + let s' = Bytes.make (n + 2) ' ' in + let rec loop i i' = + if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else + let j = try String.index_from s (i + 1) ' ' with _ -> n in + let m = j - i in + if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then + (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1)) + else if for_key && is_ident (String.sub s i m) then + (Bytes.set s' i' '_'; loop (j + 1) (i' + 2)) + else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in + loop 0 1 in + let trim_ntn (pntn, m) = (InConstrEntrySomeLevel,Bytes.sub_string pntn 1 (max 0 m)) in + let pr_ntn ntn = str "(" ++ Notation.pr_notation ntn ++ str ")" in + let pr_and_list pr = function + | [x] -> pr x + | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x + | [] -> mt () in + let pr_sc sc = str (if sc = "" then "independently" else sc) in + let pr_scs = function + | [""] -> pr_sc "" + | scs -> str "in " ++ pr_and_list pr_sc scs in + let generator, pr_tag_sc = + let ign _ = mt () in match okey with + | Some key -> + let sc = Notation.find_delimiters_scope ?loc key in + let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in + Notation.pr_scope ign sc, pr_sc + | None -> Notation.pr_scopes ign, ign in + let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in + let ptag, ttag = + let ptag, m = mk_pntn tag false in + if m <= 0 then err (str "empty notation fragment"); + ptag, trim_ntn (ptag, m) in + let last = ref "" and last_sc = ref "" in + let scs = ref [] and ntns = ref [] in + let push_sc sc = match !scs with + | "" :: scs' -> scs := "" :: sc :: scs' + | scs' -> scs := sc :: scs' in + let get s _ _ = match !last with + | "Scope " -> last_sc := s; last := "" + | "Lonely notation" -> last_sc := ""; last := "" + | "\"" -> + let pntn, m = mk_pntn s true in + if String.string_contains ~where:(Bytes.to_string pntn) ~what:(Bytes.to_string ptag) then begin + let ntn = trim_ntn (pntn, m) in + match !ntns with + | [] -> ntns := [ntn]; scs := [!last_sc] + | ntn' :: _ when ntn' = ntn -> push_sc !last_sc + | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc] + | _ :: ntns' when List.mem ntn ntns' -> () + | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns' + end; + last := "" + | _ -> last := s in + pp_with (Format.make_formatter get (fun _ -> ())) generator; + let ntn = match !ntns with + | [] -> + err (hov 0 (qtag "in" ++ str "does not occur in any notation")) + | ntn :: ntns' when ntn = ttag -> + if ntns' <> [] then begin + let pr_ntns' = pr_and_list pr_ntn ntns' in + Feedback.msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns')) + end; ntn + | [ntn] -> + Feedback.msg_info (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn + | ntns' -> + let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in + err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in + let (nvars, body), ((_, pat), osc) = match !scs with + | [sc] -> Notation.interp_notation ?loc ntn (None, [sc]) + | scs' -> + try Notation.interp_notation ?loc ntn (None, []) with _ -> + let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in + err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in + let sc = Option.default "" osc in + let _ = + let m_sc = + if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in + let ntn_pat = trim_ntn (mk_pntn pat false) in + let rbody = glob_constr_of_notation_constr ?loc body in + let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in + let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in + Feedback.msg_info (hov 0 m) in + if List.length !scs > 1 then + let scs' = List.remove (=) sc !scs in + let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in + Feedback.msg_warning (hov 4 w) + else if String.string_contains ~where:(snd ntn) ~what:" .. " then + err (pr_ntn ntn ++ str " is an n-ary notation"); + let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in + let rec sub () = function + | NVar x when List.mem_assoc x nvars -> DAst.make ?loc @@ GPatVar (FirstOrderPatVar x) + | c -> + glob_constr_of_notation_constr_with_binders ?loc (fun _ x -> (), None, x) sub () c in + let _, npat = Patternops.pattern_of_glob_constr (sub () body) in + Search.GlobSearchSubPattern npat + +} + +ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem + PRINTED BY { pr_ssr_search_item } + | [ string(s) ] -> { RGlobSearchString (loc,s,None) } + | [ string(s) "%" preident(key) ] -> { RGlobSearchString (loc,s,Some key) } + | [ constr_pattern(p) ] -> { RGlobSearchSubPattern p } +END + +{ + +let pr_ssr_search_arg _ _ _ = + let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item p in + pr_list spc pr_item + +} + +ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list + PRINTED BY { pr_ssr_search_arg } + | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> { (false, p) :: a } + | [ ssr_search_item(p) ssr_search_arg(a) ] -> { (true, p) :: a } + | [ ] -> { [] } +END + +{ + +(* Main type conclusion pattern filter *) + +let rec splay_search_pattern na = function + | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp + | Pattern.PLetIn (_, _, _, bp) -> splay_search_pattern na bp + | Pattern.PRef hr -> hr, na + | _ -> CErrors.user_err (Pp.str "no head constant in head search pattern") + +let push_rels_assum l e = + let l = List.map (fun (n,t) -> n, EConstr.Unsafe.to_constr t) l in + push_rels_assum l e + +let coerce_search_pattern_to_sort hpat = + let env = Global.env () in + let sigma = Evd.(from_env env) in + let mkPApp fp n_imps args = + let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in + Pattern.PApp (fp, args') in + let hr, na = splay_search_pattern 0 hpat in + let dc, ht = + let hr, _ = Global.type_of_global_in_context (Global.env ()) hr (** FIXME *) in + Reductionops.splay_prod env sigma (EConstr.of_constr hr) in + let np = List.length dc in + if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else + let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in + let warn () = + Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++ + pr_constr_pattern_env env sigma hpat') in + if EConstr.isSort sigma ht then begin warn (); true, hpat' end else + let filter_head, coe_path = + try + let _, cp = + Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in + warn (); + true, cp + with _ -> false, [] in + let coerce hp coe_index = + let coe_ref = coe_index.Classops.coe_value in + try + let n_imps = Option.get (Classops.hide_coercion coe_ref) in + mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] + with Not_found | Option.IsNone -> + errorstrm (str "need explicit coercion " ++ pr_global coe_ref ++ spc () + ++ str "to interpret head search pattern as type") in + filter_head, List.fold_left coerce hpat' coe_path + +let interp_head_pat hpat = + let filter_head, p = coerce_search_pattern_to_sort hpat in + let rec loop c = match CoqConstr.kind c with + | Cast (c', _, _) -> loop c' + | Prod (_, _, c') -> loop c' + | LetIn (_, _, _, c') -> loop c' + | _ -> + let env = Global.env () in + let sigma = Evd.from_env env in + Constr_matching.is_matching env sigma p (EConstr.of_constr c) in + filter_head, loop + +let all_true _ = true + +let rec interp_search_about args accu = match args with +| [] -> accu +| (flag, arg) :: rem -> + fun gr env typ -> + let ans = Search.search_about_filter arg gr env typ in + (if flag then ans else not ans) && interp_search_about rem accu gr env typ + +let interp_search_arg arg = + let arg = List.map (fun (x,arg) -> x, match arg with + | RGlobSearchString (loc,s,key) -> + if is_ident_part s then Search.GlobSearchString s else + interp_search_notation ~loc s key + | RGlobSearchSubPattern p -> + try + let env = Global.env () in + let _, p = Constrintern.intern_constr_pattern env (Evd.from_env env) p in + Search.GlobSearchSubPattern p + with e -> let e = CErrors.push e in iraise (ExplainErr.process_vernac_interp_error e)) arg in + let hpat, a1 = match arg with + | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' + | (true, Search.GlobSearchSubPattern p) :: a' -> + let filter_head, p = interp_head_pat p in + if filter_head then p, a' else all_true, arg + | _ -> all_true, arg in + let is_string = + function (_, Search.GlobSearchString _) -> true | _ -> false in + let a2, a3 = List.partition is_string a1 in + interp_search_about (a2 @ a3) (fun gr env typ -> hpat typ) + +(* Module path postfilter *) + +let pr_modloc (b, m) = if b then str "-" ++ pr_qualid m else pr_qualid m + +let wit_ssrmodloc = add_genarg "ssrmodloc" pr_modloc + +let pr_ssr_modlocs _ _ _ ml = + if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml + +} + +ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY { pr_ssr_modlocs } + | [ ] -> { [] } +END + +GRAMMAR EXTEND Gram + GLOBAL: ssr_modlocs; + modloc: [[ "-"; m = global -> { true, m } | m = global -> { false, m } ]]; + ssr_modlocs: [[ "in"; ml = LIST1 modloc -> { ml } ]]; +END + +{ + +let interp_modloc mr = + let interp_mod (_, qid) = + try Nametab.full_name_module qid with Not_found -> + CErrors.user_err ?loc:qid.CAst.loc (str "No Module " ++ pr_qualid qid) in + let mr_out, mr_in = List.partition fst mr in + let interp_bmod b = function + | [] -> fun _ _ _ -> true + | rmods -> Search.module_filter (List.map interp_mod rmods, b) in + let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in + fun gr env typ -> is_in gr env typ && is_out gr env typ + +(* The unified, extended vernacular "Search" command *) + +let ssrdisplaysearch gr env t = + let pr_res = pr_global gr ++ spc () ++ str " " ++ pr_lconstr_env env Evd.empty t in + Feedback.msg_info (hov 2 pr_res ++ fnl ()) + +} + +VERNAC COMMAND EXTEND SsrSearchPattern CLASSIFIED AS QUERY +| [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] -> + { let hpat = interp_search_arg a in + let in_mod = interp_modloc mr in + let post_filter gr env typ = in_mod gr env typ && hpat gr env typ in + let display gr env typ = + if post_filter gr env typ then ssrdisplaysearch gr env typ + in + Search.generic_search None display } +END + +(** View hint database and View application. *)(* ******************************) + +(* There are three databases of lemmas used to mediate the application *) +(* of reflection lemmas: one for forward chaining, one for backward *) +(* chaining, and one for secondary backward chaining. *) + +(* View hints *) + +{ + +let pr_raw_ssrhintref prc _ _ = let open CAst in function + | { v = CAppExpl ((None, r,x), args) } when isCHoles args -> + prc (CAst.make @@ CRef (r,x)) ++ str "|" ++ int (List.length args) + | { v = CApp ((_, { v = CRef _ }), _) } as c -> prc c + | { v = CApp ((_, c), args) } when isCxHoles args -> + prc c ++ str "|" ++ int (List.length args) + | c -> prc c + +let pr_rawhintref c = + let _, env = Pfedit.get_current_context () in + match DAst.get c with + | GApp (f, args) when isRHoles args -> + pr_glob_constr_env env f ++ str "|" ++ int (List.length args) + | _ -> pr_glob_constr_env env c + +let pr_glob_ssrhintref _ _ _ (c, _) = pr_rawhintref c + +let pr_ssrhintref prc _ _ = prc + +let mkhintref ?loc c n = match c.CAst.v with + | CRef (r,x) -> CAst.make ?loc @@ CAppExpl ((None, r, x), mkCHoles ?loc n) + | _ -> mkAppC (c, mkCHoles ?loc n) + +} + +ARGUMENT EXTEND ssrhintref + TYPED AS constr + PRINTED BY { pr_ssrhintref } + RAW_PRINTED BY { pr_raw_ssrhintref } + GLOB_PRINTED BY { pr_glob_ssrhintref } + | [ constr(c) ] -> { c } + | [ constr(c) "|" natural(n) ] -> { mkhintref ~loc c n } +END + +{ + +(* View purpose *) + +let pr_viewpos = function + | Some Ssrview.AdaptorDb.Forward -> str " for move/" + | Some Ssrview.AdaptorDb.Backward -> str " for apply/" + | Some Ssrview.AdaptorDb.Equivalence -> str " for apply//" + | None -> mt () + +let pr_ssrviewpos _ _ _ = pr_viewpos + +} + +ARGUMENT EXTEND ssrviewpos PRINTED BY { pr_ssrviewpos } + | [ "for" "move" "/" ] -> { Some Ssrview.AdaptorDb.Forward } + | [ "for" "apply" "/" ] -> { Some Ssrview.AdaptorDb.Backward } + | [ "for" "apply" "/" "/" ] -> { Some Ssrview.AdaptorDb.Equivalence } + | [ "for" "apply" "//" ] -> { Some Ssrview.AdaptorDb.Equivalence } + | [ ] -> { None } +END + +{ + +let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc () + +} + +ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY { pr_ssrviewposspc } + | [ ssrviewpos(i) ] -> { i } +END + +{ + +let print_view_hints kind l = + let pp_viewname = str "Hint View" ++ pr_viewpos (Some kind) ++ str " " in + let pp_hints = pr_list spc pr_rawhintref l in + Feedback.msg_info (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ()) + +} + +VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY +| [ "Print" "Hint" "View" ssrviewpos(i) ] -> + { match i with + | Some k -> print_view_hints k (Ssrview.AdaptorDb.get k) + | None -> + List.iter (fun k -> print_view_hints k (Ssrview.AdaptorDb.get k)) + [ Ssrview.AdaptorDb.Forward; + Ssrview.AdaptorDb.Backward; + Ssrview.AdaptorDb.Equivalence ] + } +END + +{ + +let glob_view_hints lvh = + List.map (Constrintern.intern_constr (Global.env ()) (Evd.from_env (Global.env ()))) lvh + +} + +VERNAC COMMAND EXTEND HintView CLASSIFIED AS SIDEFF + | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] -> + { let hints = glob_view_hints lvh in + match n with + | None -> + Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Forward hints; + Ssrview.AdaptorDb.declare Ssrview.AdaptorDb.Backward hints + | Some k -> + Ssrview.AdaptorDb.declare k hints } +END + +(** Canonical Structure alias *) + +GRAMMAR EXTEND Gram + GLOBAL: gallina_ext; + + gallina_ext: + (* Canonical structure *) + [[ IDENT "Canonical"; qid = Constr.global -> + { Vernacexpr.VernacCanonical (CAst.make @@ AN qid) } + | IDENT "Canonical"; ntn = Prim.by_notation -> + { Vernacexpr.VernacCanonical (CAst.make @@ ByNotation ntn) } + | IDENT "Canonical"; qid = Constr.global; + d = G_vernac.def_body -> + { let s = coerce_reference_to_id qid in + Vernacexpr.VernacDefinition + ((Decl_kinds.NoDischarge,Decl_kinds.CanonicalStructure), + ((CAst.make (Name s)),None), d) } + ]]; +END + +(** Keyword compatibility fixes. *) + +(* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *) +(* identifiers used as keywords. This is incompatible with ssreflect.v *) +(* which makes "by" and "of" true keywords, because of technicalities *) +(* in the internal lexer-parser API of Coq. We patch this here by *) +(* adding new parsing rules that recognize the new keywords. *) +(* To make matters worse, the Coq grammar for tactics fails to *) +(* export the non-terminals we need to patch. Fortunately, the CamlP5 *) +(* API provides a backdoor access (with loads of Obj.magic trickery). *) + +(* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *) +(* longer and thus comment out. Such comments are marked with v8.3 *) + +{ + +open Pltac + +} + +GRAMMAR EXTEND Gram + GLOBAL: hypident; + hypident: [ + [ "("; IDENT "type"; "of"; id = Prim.identref; ")" -> { id, Locus.InHypTypeOnly } + | "("; IDENT "value"; "of"; id = Prim.identref; ")" -> { id, Locus.InHypValueOnly } + ] ]; +END + +GRAMMAR EXTEND Gram + GLOBAL: hloc; +hloc: [ + [ "in"; "("; "Type"; "of"; id = ident; ")" -> + { Tacexpr.HypLocation (CAst.make id, Locus.InHypTypeOnly) } + | "in"; "("; IDENT "Value"; "of"; id = ident; ")" -> + { Tacexpr.HypLocation (CAst.make id, Locus.InHypValueOnly) } + ] ]; +END + +GRAMMAR EXTEND Gram + GLOBAL: constr_eval; + constr_eval: [ + [ IDENT "type"; "of"; c = Constr.constr -> { Genredexpr.ConstrTypeOf c }] + ]; +END + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +{ + +let () = CLexer.set_keyword_state frozen_lexer ;; + +} + +(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/plugins/ssrmatching/g_ssrmatching.ml4 b/plugins/ssrmatching/g_ssrmatching.ml4 deleted file mode 100644 index 9e1f992f38..0000000000 --- a/plugins/ssrmatching/g_ssrmatching.ml4 +++ /dev/null @@ -1,100 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* [ mk_rpattern (T (mk_lterm c None)) ] - | [ "in" lconstr(c) ] -> [ mk_rpattern (In_T (mk_lterm c None)) ] - | [ lconstr(x) "in" lconstr(c) ] -> - [ mk_rpattern (X_In_T (mk_lterm x None, mk_lterm c None)) ] - | [ "in" lconstr(x) "in" lconstr(c) ] -> - [ mk_rpattern (In_X_In_T (mk_lterm x None, mk_lterm c None)) ] - | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> - [ mk_rpattern (E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ] - | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> - [ mk_rpattern (E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ] -END - -let pr_ssrterm _ _ _ = pr_ssrterm - -ARGUMENT EXTEND cpattern - PRINTED BY pr_ssrterm - INTERPRETED BY interp_ssrterm - GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm - RAW_PRINTED BY pr_ssrterm - GLOB_PRINTED BY pr_ssrterm -| [ "Qed" constr(c) ] -> [ mk_lterm c None ] -END - -let input_ssrtermkind strm = match Util.stream_nth 0 strm with - | Tok.KEYWORD "(" -> '(' - | Tok.KEYWORD "@" -> '@' - | _ -> ' ' -let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind - -GEXTEND Gram - GLOBAL: cpattern; - cpattern: [[ k = ssrtermkind; c = constr -> - let pattern = mk_term k c None in - if loc_of_cpattern pattern <> Some !@loc && k = '(' - then mk_term 'x' c None - else pattern ]]; -END - -ARGUMENT EXTEND lcpattern - TYPED AS cpattern - PRINTED BY pr_ssrterm - INTERPRETED BY interp_ssrterm - GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm - RAW_PRINTED BY pr_ssrterm - GLOB_PRINTED BY pr_ssrterm -| [ "Qed" lconstr(c) ] -> [ mk_lterm c None ] -END - -GEXTEND Gram - GLOBAL: lcpattern; - lcpattern: [[ k = ssrtermkind; c = lconstr -> - let pattern = mk_term k c None in - if loc_of_cpattern pattern <> Some !@loc && k = '(' - then mk_term 'x' c None - else pattern ]]; -END - -ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_rpattern -| [ rpattern(pat) ] -> [ pat ] -END - -TACTIC EXTEND ssrinstoftpat -| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof arg) ] -END - -(* We wipe out all the keywords generated by the grammar rules we defined. *) -(* The user is supposed to Require Import ssreflect or Require ssreflect *) -(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) -(* consequence the extended ssreflect grammar. *) -let () = CLexer.set_keyword_state frozen_lexer ;; diff --git a/plugins/ssrmatching/g_ssrmatching.mlg b/plugins/ssrmatching/g_ssrmatching.mlg new file mode 100644 index 0000000000..3f0794fdd4 --- /dev/null +++ b/plugins/ssrmatching/g_ssrmatching.mlg @@ -0,0 +1,120 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* { mk_rpattern (T (mk_lterm c None)) } + | [ "in" lconstr(c) ] -> { mk_rpattern (In_T (mk_lterm c None)) } + | [ lconstr(x) "in" lconstr(c) ] -> + { mk_rpattern (X_In_T (mk_lterm x None, mk_lterm c None)) } + | [ "in" lconstr(x) "in" lconstr(c) ] -> + { mk_rpattern (In_X_In_T (mk_lterm x None, mk_lterm c None)) } + | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> + { mk_rpattern (E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) } + | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> + { mk_rpattern (E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) } +END + +{ + +let pr_ssrterm _ _ _ = pr_ssrterm + +} + +ARGUMENT EXTEND cpattern + PRINTED BY { pr_ssrterm } + INTERPRETED BY { interp_ssrterm } + GLOBALIZED BY { glob_cpattern } SUBSTITUTED BY { subst_ssrterm } + RAW_PRINTED BY { pr_ssrterm } + GLOB_PRINTED BY { pr_ssrterm } +| [ "Qed" constr(c) ] -> { mk_lterm c None } +END + +{ + +let input_ssrtermkind strm = match Util.stream_nth 0 strm with + | Tok.KEYWORD "(" -> '(' + | Tok.KEYWORD "@" -> '@' + | _ -> ' ' +let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind + +} + +GRAMMAR EXTEND Gram + GLOBAL: cpattern; + cpattern: [[ k = ssrtermkind; c = constr -> { + let pattern = mk_term k c None in + if loc_of_cpattern pattern <> Some loc && k = '(' + then mk_term 'x' c None + else pattern } ]]; +END + +ARGUMENT EXTEND lcpattern + TYPED AS cpattern + PRINTED BY { pr_ssrterm } + INTERPRETED BY { interp_ssrterm } + GLOBALIZED BY { glob_cpattern } SUBSTITUTED BY { subst_ssrterm } + RAW_PRINTED BY { pr_ssrterm } + GLOB_PRINTED BY { pr_ssrterm } +| [ "Qed" lconstr(c) ] -> { mk_lterm c None } +END + +GRAMMAR EXTEND Gram + GLOBAL: lcpattern; + lcpattern: [[ k = ssrtermkind; c = lconstr -> { + let pattern = mk_term k c None in + if loc_of_cpattern pattern <> Some loc && k = '(' + then mk_term 'x' c None + else pattern } ]]; +END + +ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY { pr_rpattern } +| [ rpattern(pat) ] -> { pat } +END + +TACTIC EXTEND ssrinstoftpat +| [ "ssrinstancesoftpat" cpattern(arg) ] -> { Proofview.V82.tactic (ssrinstancesof arg) } +END + +{ + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +let () = CLexer.set_keyword_state frozen_lexer ;; + +} diff --git a/plugins/syntax/g_numeral.ml4 b/plugins/syntax/g_numeral.ml4 deleted file mode 100644 index 55f61a58f9..0000000000 --- a/plugins/syntax/g_numeral.ml4 +++ /dev/null @@ -1,38 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) -(* mt () - | Warning n -> str "(warning after " ++ str n ++ str ")" - | Abstract n -> str "(abstract after " ++ str n ++ str ")" - -ARGUMENT EXTEND numnotoption - PRINTED BY pr_numnot_option -| [ ] -> [ Nop ] -| [ "(" "warning" "after" bigint(waft) ")" ] -> [ Warning waft ] -| [ "(" "abstract" "after" bigint(n) ")" ] -> [ Abstract n ] -END - -VERNAC COMMAND EXTEND NumeralNotation CLASSIFIED AS SIDEFF - | [ "Numeral" "Notation" reference(ty) reference(f) reference(g) ":" - ident(sc) numnotoption(o) ] -> - [ vernac_numeral_notation (Locality.make_module_locality atts.locality) ty f g (Id.to_string sc) o ] -END diff --git a/plugins/syntax/g_numeral.mlg b/plugins/syntax/g_numeral.mlg new file mode 100644 index 0000000000..5dbc9eea7a --- /dev/null +++ b/plugins/syntax/g_numeral.mlg @@ -0,0 +1,42 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* mt () + | Warning n -> str "(warning after " ++ str n ++ str ")" + | Abstract n -> str "(abstract after " ++ str n ++ str ")" + +} + +ARGUMENT EXTEND numnotoption + PRINTED BY { pr_numnot_option } +| [ ] -> { Nop } +| [ "(" "warning" "after" bigint(waft) ")" ] -> { Warning waft } +| [ "(" "abstract" "after" bigint(n) ")" ] -> { Abstract n } +END + +VERNAC COMMAND EXTEND NumeralNotation CLASSIFIED AS SIDEFF + | [ "Numeral" "Notation" reference(ty) reference(f) reference(g) ":" + ident(sc) numnotoption(o) ] -> + { vernac_numeral_notation (Locality.make_module_locality atts.locality) ty f g (Id.to_string sc) o } +END -- cgit v1.2.3