From 32170384190168856efeac5bcf90edf1170b54d6 Mon Sep 17 00:00:00 2001 From: herbelin Date: Wed, 29 May 2002 10:48:37 +0000 Subject: Nouveau modèle d'analyse syntaxique et d'interprétation des tactiques et commandes vernaculaires (cf dev/changements.txt pour plus de précisions) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@2722 85f007b7-540e-0410-9357-904b9bb8a0f7 --- contrib/interface/Centaur.v | 8 +- contrib/interface/blast.ml | 11 +- contrib/interface/blast.mli | 4 +- contrib/interface/centaur.ml | 776 ------------ contrib/interface/centaur.ml4 | 946 ++++++++++++++ contrib/interface/dad.ml | 102 +- contrib/interface/dad.mli | 4 +- contrib/interface/debug_tac.ml | 474 ------- contrib/interface/debug_tac.ml4 | 570 +++++++++ contrib/interface/debug_tac.mli | 6 +- contrib/interface/name_to_ast.ml | 62 +- contrib/interface/name_to_ast.mli | 2 +- contrib/interface/parse.ml | 131 +- contrib/interface/pbp.ml | 471 +++++-- contrib/interface/pbp.mli | 4 +- contrib/interface/showproof.ml | 207 ++-- contrib/interface/xlate.ml | 1393 +++++++++++++++++++-- contrib/interface/xlate.mli | 8 +- kernel/safe_typing.ml | 7 +- kernel/term.ml | 3 + kernel/term.mli | 1 + kernel/typeops.ml | 7 + kernel/typeops.mli | 4 + kernel/univ.ml | 21 +- lib/options.ml | 3 +- lib/options.mli | 4 +- lib/pp_control.ml | 6 +- lib/pp_control.mli | 4 +- lib/util.ml | 2 + lib/util.mli | 2 + library/declare.ml | 5 +- library/declare.mli | 14 +- library/goptions.ml | 55 +- library/goptions.mli | 51 +- library/impargs.ml | 2 +- library/impargs.mli | 2 +- library/library.ml | 29 +- library/library.mli | 9 +- library/nametab.ml | 32 +- library/nametab.mli | 14 +- parsing/ast.ml | 535 +++++--- parsing/ast.mli | 88 +- parsing/astterm.ml | 17 +- parsing/astterm.mli | 10 +- parsing/coqast.ml | 147 +++ parsing/coqast.mli | 7 + parsing/coqlib.ml | 22 +- parsing/coqlib.mli | 10 - parsing/egrammar.ml | 166 ++- parsing/egrammar.mli | 23 +- parsing/esyntax.ml | 105 +- parsing/esyntax.mli | 13 +- parsing/extend.ml | 243 ++++ parsing/extend.mli | 97 +- parsing/g_basevernac.ml4 | 509 ++++---- parsing/g_cases.ml4 | 2 +- parsing/g_constr.ml4 | 73 +- parsing/g_ltac.ml4 | 311 +++-- parsing/g_natsyntax.ml | 30 +- parsing/g_prim.ml4 | 162 ++- parsing/g_proofs.ml4 | 238 ++-- parsing/g_rsyntax.ml | 11 +- parsing/g_tactic.ml4 | 568 ++++----- parsing/g_vernac.ml4 | 573 +++++---- parsing/g_zsyntax.ml | 38 +- parsing/genarg.ml | 181 +++ parsing/genarg.mli | 208 ++++ parsing/pcoq.ml4 | 511 ++++---- parsing/pcoq.mli | 224 ++-- parsing/ppconstr.ml | 148 +++ parsing/ppconstr.mli | 33 + parsing/pptactic.ml | 607 +++++++++ parsing/pptactic.mli | 35 + parsing/prettyp.ml | 47 +- parsing/prettyp.mli | 9 +- parsing/printer.ml | 81 +- parsing/printer.mli | 4 +- parsing/q_coqast.ml4 | 476 ++++++- parsing/q_util.ml4 | 63 + parsing/q_util.mli | 28 + pretyping/classops.ml | 57 +- pretyping/classops.mli | 18 +- pretyping/detyping.ml | 70 +- pretyping/indrec.mli | 1 + pretyping/rawterm.ml | 39 +- pretyping/rawterm.mli | 38 + pretyping/tacred.ml | 61 +- pretyping/tacred.mli | 12 +- proofs/clenv.ml | 64 +- proofs/clenv.mli | 12 +- proofs/evar_refiner.ml | 3 + proofs/evar_refiner.mli | 2 + proofs/logic.ml | 1 - proofs/logic.mli | 1 - proofs/pfedit.ml | 30 +- proofs/pfedit.mli | 21 +- proofs/tacmach.ml | 151 ++- proofs/tacmach.mli | 80 +- proofs/tactic_debug.ml | 5 +- proofs/tactic_debug.mli | 5 +- tactics/auto.ml | 352 ++---- tactics/auto.mli | 47 +- tactics/autorewrite.ml | 46 +- tactics/autorewrite.mli | 4 +- tactics/contradiction.mli | 19 + tactics/dhyp.ml | 119 +- tactics/dhyp.mli | 11 + tactics/elim.ml | 110 +- tactics/elim.mli | 15 +- tactics/equality.ml | 203 +-- tactics/equality.mli | 41 +- tactics/extraargs.mli | 21 + tactics/extratactics.mli | 19 + tactics/hiddentac.ml | 104 +- tactics/hiddentac.mli | 91 +- tactics/hipattern.ml | 21 +- tactics/inv.ml | 105 +- tactics/inv.mli | 13 +- tactics/leminv.ml | 99 +- tactics/leminv.mli | 15 + tactics/refine.ml | 17 - tactics/refine.mli | 2 + tactics/setoid_replace.ml | 105 +- tactics/setoid_replace.mli | 5 + tactics/tacinterp.ml | 1738 ++++++++++++++++++++++++++ tactics/tacinterp.mli | 115 ++ tactics/tactics.ml | 831 ++++--------- tactics/tactics.mli | 129 +- tactics/wcclausenv.ml | 7 +- tactics/wcclausenv.mli | 2 +- toplevel/cerrors.ml | 2 +- toplevel/class.ml | 10 + toplevel/class.mli | 4 + toplevel/command.ml | 128 +- toplevel/command.mli | 22 +- toplevel/coqtop.ml | 5 +- toplevel/discharge.ml | 1 + toplevel/himsg.ml | 5 - toplevel/metasyntax.ml | 259 ++-- toplevel/metasyntax.mli | 12 +- toplevel/protectedtoplevel.ml | 8 +- toplevel/record.ml | 37 +- toplevel/record.mli | 4 +- toplevel/recordobj.ml | 2 + toplevel/recordobj.mli | 5 +- toplevel/toplevel.ml | 18 +- toplevel/toplevel.mli | 2 +- toplevel/vernac.ml | 47 +- toplevel/vernac.mli | 3 +- toplevel/vernacentries.ml | 2449 ++++++++++++------------------------- toplevel/vernacentries.mli | 36 +- toplevel/vernacinterp.ml | 84 +- toplevel/vernacinterp.mli | 41 +- 153 files changed, 12508 insertions(+), 7862 deletions(-) delete mode 100644 contrib/interface/centaur.ml create mode 100644 contrib/interface/centaur.ml4 delete mode 100644 contrib/interface/debug_tac.ml create mode 100644 contrib/interface/debug_tac.ml4 create mode 100644 parsing/extend.ml create mode 100644 parsing/genarg.ml create mode 100644 parsing/genarg.mli create mode 100644 parsing/ppconstr.ml create mode 100644 parsing/ppconstr.mli create mode 100644 parsing/pptactic.ml create mode 100644 parsing/pptactic.mli create mode 100644 parsing/q_util.ml4 create mode 100644 parsing/q_util.mli create mode 100644 tactics/contradiction.mli create mode 100644 tactics/extraargs.mli create mode 100644 tactics/extratactics.mli create mode 100644 tactics/leminv.mli create mode 100644 tactics/tacinterp.ml create mode 100644 tactics/tacinterp.mli diff --git a/contrib/interface/Centaur.v b/contrib/interface/Centaur.v index 654cf123f..d27929f86 100644 --- a/contrib/interface/Centaur.v +++ b/contrib/interface/Centaur.v @@ -1,3 +1,4 @@ +(* Declare ML Module "ctast". Declare ML Module "paths". Declare ML Module "name_to_ast". @@ -15,7 +16,7 @@ Declare ML Module "history". Declare ML Module "centaur". (* Require Export Illustrations. *) (* Require Export AddDad. *) - +(* Grammar vernac vernac : ast := goal_cmd [ "Goal" "Cmd" numarg($n) "with" tacarg($tac) "." ] -> [(GOAL_CMD $n (TACTIC $tac))] @@ -33,7 +34,7 @@ Grammar vernac vernac : ast := | compute_in_goal [ "COMPUTE_IN_GOAL" numarg($n) constrarg($c) "." ] -> [(CHECK_IN_GOAL "COMPUTE" $n $c)] | centaur_reset [ "Centaur" "Reset" identarg($id) "." ] -> [(Centaur_Reset $id)] -(* | show_dad_rules [ "Show" "Dad" "Rules" "." ] -> [(Show_dad_rules)] *) +(*| show_dad_rules [ "Show" "Dad" "Rules" "." ] -> [(Show_dad_rules)]*) | start_pcoq [ "Start" "Pcoq" "Mode" "." ] -> [ (START_PCOQ) ] | start_pcoq [ "Start" "Pcoq" "Debug" "Mode" "." ] -> [ (START_PCOQ_DEBUG) ]. Grammar vernac ne_id_list : ast list := @@ -83,4 +84,5 @@ Grammar vernac vernac : ast := | text_proof_flag_on [ "Text" "Mode" "Off" "." ] -> [(TEXT_MODE (AST "off"))]. - +*) +*) diff --git a/contrib/interface/blast.ml b/contrib/interface/blast.ml index f6a47f986..4c57760de 100755 --- a/contrib/interface/blast.ml +++ b/contrib/interface/blast.ml @@ -191,7 +191,7 @@ and e_my_find_search db_list local_db hdc concl = tclTHEN (unify_e_resolve (term,cl)) (e_trivial_fail_db db_list local_db) | Unfold_nth c -> unfold_constr c - | Extern tacast -> Tacticals.conclPattern concl + | Extern tacast -> Auto.conclPattern concl (out_some p) tacast in (free_try tac,fmt_autotactic t)) @@ -567,7 +567,6 @@ let vire_extvar s = let blast gls = let leaf g = { status = Incomplete_proof; - subproof = None; goal = g; ref = None } in try (let (sgl,v) as res = !blast_tactic gls in @@ -592,17 +591,19 @@ let blast gls = ;; let blast_tac display_function = function - | ((Integer n)::_) as l -> + | (n::_) as l -> (function g -> let exp_ast = (blast g) in - (display_function (ast_to_ct exp_ast); + (display_function exp_ast; tclIDTAC g)) | _ -> failwith "expecting other arguments";; let blast_tac_txt = - blast_tac (function x -> msgnl(Printer.gentacpr (ct_to_ast x)));; + blast_tac (function x -> msgnl(Pptactic.pr_raw_tactic x));; +(* Obsolte ? overwriting_add_tactic "Blast1" blast_tac_txt;; +*) (* Grammar tactic ne_numarg_list : list := diff --git a/contrib/interface/blast.mli b/contrib/interface/blast.mli index e8bcf95c6..21c29bc98 100644 --- a/contrib/interface/blast.mli +++ b/contrib/interface/blast.mli @@ -1,5 +1,5 @@ -val blast_tac : (Ctast.t -> 'a) -> - Proof_type.tactic_arg list -> +val blast_tac : (Tacexpr.raw_tactic_expr -> 'a) -> + int list -> Proof_type.goal Tacmach.sigma -> Proof_type.goal list Proof_type.sigma * Proof_type.validation;; diff --git a/contrib/interface/centaur.ml b/contrib/interface/centaur.ml deleted file mode 100644 index 6f5281d56..000000000 --- a/contrib/interface/centaur.ml +++ /dev/null @@ -1,776 +0,0 @@ - -(*Toplevel loop for the communication between Coq and Centaur *) -open Names;; -open Nameops;; -open Util;; -open Ast;; -open Term;; -open Pp;; -open Libobject;; -open Library;; -open Vernacinterp;; -open Evd;; -open Proof_trees;; -open Termast;; -open Tacmach;; -open Pfedit;; -open Proof_type;; -open Parsing;; -open Environ;; -open Declare;; -open Declarations;; -open Rawterm;; -open Reduction;; -open Classops;; -open Vernacinterp;; -open Vernac;; -open Command;; -open Protectedtoplevel;; -open Coqast;; -open Line_oriented_parser;; -open Xlate;; -open Vtp;; -open Ascent;; -open Translate;; -open Name_to_ast;; -open Pbp;; -open Blast;; -open Dad;; -open Debug_tac;; -open Search;; -open Astterm;; -open Nametab;; -open Showproof;; -open Showproof_ct;; - - -let text_proof_flag = ref "en";; - -let current_proof_name = ref "";; - -let current_goal_index = ref 0;; - -set_flags := (function () -> - if List.mem "G_natsyntax" (Mltop.get_loaded_modules()) then - (g_nat_syntax_flag := true; ()) - else ());; - -let guarded_force_eval_stream (s : std_ppcmds) = - let l = ref [] in - let f elt = l:= elt :: !l in - (try Stream.iter f s with - | _ -> f (Stream.next (str "error guarded_force_eval_stream"))); - Stream.of_list (List.rev !l);; - - -let rec string_of_path p = - match p with [] -> "\n" - | i::p -> (string_of_int i)^" "^ (string_of_path p) -;; -let print_path p = - output_results_nl (str "Path:" ++ str (string_of_path p)) -;; - -let kill_proof_node index = - let paths = History.historical_undo !current_proof_name index in - let _ = List.iter - (fun path -> (traverse_to path; - Pfedit.mutate weak_undo_pftreestate; - traverse_to [])) - paths in - History.border_length !current_proof_name;; - - -(*Message functions, the text of these messages is recognized by the protocols *) -(*of CtCoq *) -let ctf_header message_name request_id = - fnl () ++ str "message" ++ fnl() ++ str message_name ++ fnl() ++ - int request_id ++ fnl();; - -let ctf_acknowledge_command request_id command_count opt_exn = - let goal_count, goal_index = - if refining() then - let g_count = - List.length - (fst (frontier (proof_of_pftreestate (get_pftreestate ())))) in - g_count, (min g_count !current_goal_index) - else - (0, 0) in - (ctf_header "acknowledge" request_id ++ - int command_count ++ fnl() ++ - int goal_count ++ fnl () ++ - int goal_index ++ fnl () ++ - str !current_proof_name ++ fnl() ++ - (match opt_exn with - Some e -> Cerrors.explain_exn e - | None -> mt ()) ++ fnl() ++ str "E-n-d---M-e-s-s-a-g-e" ++ fnl ());; - -let ctf_undoResults = ctf_header "undo_results";; - -let ctf_TextMessage = ctf_header "text_proof";; - -let ctf_SearchResults = ctf_header "search_results";; - -let ctf_OtherGoal = ctf_header "other_goal";; - -let ctf_Location = ctf_header "location";; - -let ctf_StateMessage = ctf_header "state";; - -let ctf_PathGoalMessage () = - fnl () ++ str "message" ++ fnl () ++ str "single_goal" ++ fnl ();; - -let ctf_GoalReqIdMessage = ctf_header "single_goal_state";; - -let ctf_NewStateMessage = ctf_header "fresh_state";; - -let ctf_SavedMessage () = fnl () ++ str "message" ++ fnl () ++ - str "saved" ++ fnl();; - -let ctf_KilledMessage req_id ngoals = - ctf_header "killed" req_id ++ int ngoals ++ fnl ();; - -let ctf_AbortedAllMessage () = - fnl() ++ str "message" ++ fnl() ++ str "aborted_all" ++ fnl();; - -let ctf_AbortedMessage request_id na = - ctf_header "aborted_proof" request_id ++ str na ++ fnl () ++ - str "E-n-d---M-e-s-s-a-g-e" ++ fnl ();; - -let ctf_UserErrorMessage request_id stream = - let stream = guarded_force_eval_stream stream in - ctf_header "user_error" request_id ++ stream ++ fnl() ++ - str "E-n-d---M-e-s-s-a-g-e" ++ fnl();; - -let ctf_ResetInitialMessage () = - fnl () ++ str "message" ++ fnl () ++ str "reset_initial" ++ fnl ();; - -let ctf_ResetIdentMessage request_id s = - ctf_header "reset_ident" request_id ++ str s ++ fnl () ++ - str "E-n-d---M-e-s-s-a-g-e" ++ fnl();; - -type vtp_tree = - | P_rl of ct_RULE_LIST - | P_r of ct_RULE - | P_s_int of ct_SIGNED_INT_LIST - | P_pl of ct_PREMISES_LIST - | P_cl of ct_COMMAND_LIST - | P_t of ct_TACTIC_COM - | P_text of ct_TEXT - | P_ids of ct_ID_LIST;; - -let print_tree t = - (match t with - | P_rl x -> fRULE_LIST x - | P_r x -> fRULE x - | P_s_int x -> fSIGNED_INT_LIST x - | P_pl x -> fPREMISES_LIST x - | P_cl x -> fCOMMAND_LIST x - | P_t x -> fTACTIC_COM x - | P_text x -> fTEXT x - | P_ids x -> fID_LIST x); - print_string "e\nblabla\n";; - - - -let break_happened = ref false;; - -let output_results stream vtp_tree = - let _ = Sys.signal Sys.sigint - (Sys.Signal_handle(fun i -> (break_happened := true;()))) in - msg stream; - match vtp_tree with - Some t -> print_tree t - | None -> ();; - -let output_results_nl stream = - let _ = Sys.signal Sys.sigint - (Sys.Signal_handle(fun i -> break_happened := true;())) - in - msgnl stream;; - - -let rearm_break () = - let _ = Sys.signal Sys.sigint (Sys.Signal_handle(fun i -> raise Sys.Break)) - in ();; - -let check_break () = - if (!break_happened) then - begin - break_happened := false; - raise Sys.Break - end - else ();; - -let print_past_goal index = - let path = History.get_path_for_rank !current_proof_name index in - try traverse_to path; - let pf = proof_of_pftreestate (get_pftreestate ()) in - output_results (ctf_PathGoalMessage ()) - (Some (P_r (translate_goal pf.goal))) - with - | Invalid_argument s -> - ((try traverse_to [] with _ -> ()); - error "No focused proof (No proof-editing in progress)") - | e -> (try traverse_to [] with _ -> ()); raise e -;; - -let show_nth n = - try - let pf = proof_of_pftreestate (get_pftreestate()) in - if (!text_proof_flag<>"off") then - (if n=0 - then output_results (ctf_TextMessage !global_request_id) - (Some (P_text (show_proof !text_proof_flag []))) - else - let path = History.get_nth_open_path !current_proof_name n in - output_results (ctf_TextMessage !global_request_id) - (Some (P_text (show_proof !text_proof_flag path)))) - else - output_results (ctf_GoalReqIdMessage !global_request_id) - (let goal = List.nth (fst (frontier pf)) - (n - 1) in - (Some (P_r (translate_goal goal)))) - with - | Invalid_argument s -> - error "No focused proof (No proof-editing in progress)";; - -(* The rest of the file contains commands that are changed from the plain - Coq distribution *) - -let ctv_SEARCH_LIST = ref ([] : ct_PREMISE list);; - -let filter_by_module_from_varg_list (l:vernac_arg list) = - let dir_list, b = Vernacentries.inside_outside l in - Search.filter_by_module_from_list (dir_list, b);; - -let add_search (global_reference:global_reference) assumptions cstr = - try - let env = Global.env() in - let id_string = - string_of_qualid (Nametab.shortest_qualid_of_global env - global_reference) in - let ast = - try - CT_premise (CT_ident id_string, translate_constr false assumptions cstr) - with Not_found -> - CT_premise (CT_ident id_string, - CT_coerce_ID_to_FORMULA( - CT_ident ("Error printing" ^ id_string))) in - ctv_SEARCH_LIST:= ast::!ctv_SEARCH_LIST - with e -> msgnl (str "add_search raised an exception"); raise e;; - -let make_error_stream node_string = - str "The syntax of " ++ str node_string ++ - str " is inconsistent with the vernac interpreter entry";; - -let ctf_EmptyGoalMessage id = - fnl () ++ str "Empty Goal is a no-op. Fun oh fun." ++ fnl ();; - - -let print_check (ast, judg) = - let {uj_val=value; uj_type=typ} = judg in - let value_ct_ast = - (try translate_constr false (Global.env()) value - with UserError(f,str) -> - raise(UserError(f, - Ast.print_ast - (ast_of_constr true (Global.env()) value) ++ - fnl () ++ str ))) in - let type_ct_ast = - (try translate_constr false (Global.env()) typ - with UserError(f,str) -> - raise(UserError(f, Ast.print_ast (ast_of_constr true (Global.env()) - value) ++ fnl() ++ str))) in - ((ctf_SearchResults !global_request_id), - (Some (P_pl - (CT_premises_list - [CT_coerce_TYPED_FORMULA_to_PREMISE - (CT_typed_formula(value_ct_ast,type_ct_ast) - )]))));; - -let ct_print_eval ast red_fun env evd judg = -((if refining() then traverse_to []); -let {uj_val=value; uj_type=typ} = judg in -let nvalue = red_fun value -(* // Attention , ici il faut peut tre utiliser des environnemenst locaux *) -and ntyp = nf_betaiota typ in -(ctf_SearchResults !global_request_id, - Some (P_pl - (CT_premises_list - [CT_eval_result - (xlate_formula ast, - translate_constr false env nvalue, - translate_constr false env ntyp)]))));; - - - -(* The following function is copied from globpr in env/printer.ml *) -let globcv x = - let env = Global.env() in - match x with - | Node(_,"MUTIND", (Path(_,sp))::(Num(_,tyi))::_) -> - let env = Global.env() in - convert_qualid - (Nametab.shortest_qualid_of_global env (IndRef(sp,tyi))) - | Node(_,"MUTCONSTRUCT",(Path(_,sp))::(Num(_,tyi))::(Num(_,i))::_) -> - convert_qualid - (Nametab.shortest_qualid_of_global env - (ConstructRef ((sp, tyi), i))) - | _ -> failwith "globcv : unexpected value";; - -let pbp_tac_pcoq = - pbp_tac (function (x:Ctast.t) -> - output_results - (ctf_header "pbp_results" !global_request_id) - (Some (P_t(xlate_tactic x))));; - -let blast_tac_pcoq = - blast_tac (function (x:Ctast.t) -> - output_results - (ctf_header "pbp_results" !global_request_id) - (Some (P_t(xlate_tactic x))));; - - -let dad_tac_pcoq = - dad_tac(function x -> - output_results - (ctf_header "pbp_results" !global_request_id) - (Some (P_t(xlate_tactic x))));; - -let search_output_results () = - output_results - (ctf_SearchResults !global_request_id) - (Some (P_pl (CT_premises_list - (List.rev !ctv_SEARCH_LIST))));; - - -let debug_tac2_pcoq = function - [Tacexp ast] -> - (fun g -> - let the_goal = ref (None : goal sigma option) in - let the_ast = ref ast in - let the_path = ref ([] : int list) in - try - let result = report_error ast the_goal the_ast the_path [] g in - (errorlabstrm "DEBUG TACTIC" - (str "no error here " ++ fnl () ++ pr_goal (sig_it g) ++ - fnl () ++ str "the tactic is" ++ fnl () ++ - Printer.gentacpr ast); - result) - with - e -> - match !the_goal with - None -> raise e - | Some g -> - (output_results - (ctf_Location !global_request_id) - (Some (P_s_int - (CT_signed_int_list - (List.map - (fun n -> CT_coerce_INT_to_SIGNED_INT - (CT_int n)) - (clean_path 0 ast - (List.rev !the_path))))))); - (output_results - (ctf_OtherGoal !global_request_id) - (Some (P_r (translate_goal (sig_it g))))); - raise e) - | _ -> error "wrong arguments for debug_tac2_pcoq";; - -let rec selectinspect n env = - match env with - [] -> [] - | a::tl -> - if n = 0 then - [] - else - match a with - (sp, Lib.Leaf lobj) -> a::(selectinspect (n -1 ) tl) - | _ -> (selectinspect n tl);; - -open Term;; - -let inspect n = - let env = Global.env() in - let add_search2 x y = add_search x env y in - let l = selectinspect n (Lib.contents_after None) in - ctv_SEARCH_LIST := []; - List.iter - (fun a -> - try - (match a with - sp, Lib.Leaf lobj -> - (match sp, object_tag lobj with - _, "VARIABLE" -> - let ((_, _, v), _) = get_variable (basename sp) in - add_search2 (Nametab.locate (qualid_of_sp sp)) v - | sp, ("CONSTANT"|"PARAMETER") -> - let {const_type=typ} = Global.lookup_constant sp in - add_search2 (Nametab.locate (qualid_of_sp sp)) typ - | sp, "MUTUALINDUCTIVE" -> - add_search2 (Nametab.locate (qualid_of_sp sp)) - (Pretyping.understand Evd.empty (Global.env()) - (RRef(dummy_loc, IndRef(sp,0)))) - | _ -> failwith ("unexpected value 1 for "^ - (string_of_id (basename sp)))) - | _ -> failwith "unexpected value") - with e -> ()) - l; - output_results - (ctf_SearchResults !global_request_id) - (Some - (P_pl (CT_premises_list (List.rev !ctv_SEARCH_LIST))));; - -let ct_int_to_TARG n = - CT_coerce_ID_OR_INT_to_TARG - (CT_coerce_INT_to_ID_OR_INT (CT_int n));; - -let pair_list_to_ct l = - CT_user_tac(CT_ident "pair_int_list", - CT_targ_list - (List.map (fun (a,b) -> - CT_coerce_TACTIC_COM_to_TARG - (CT_user_tac - (CT_ident "pair_int", - CT_targ_list - [ct_int_to_TARG a; ct_int_to_TARG b]))) - l));; - -let logical_kill n = - let path = History.get_path_for_rank !current_proof_name n in - begin - traverse_to path; - Pfedit.mutate weak_undo_pftreestate; - (let kept_cmds, undone_cmds, remaining_goals, current_goal = - History.logical_undo !current_proof_name n in - output_results (ctf_undoResults !global_request_id) - (Some - (P_t - (CT_user_tac - (CT_ident "log_undo_result", - CT_targ_list - [CT_coerce_TACTIC_COM_to_TARG (pair_list_to_ct kept_cmds); - CT_coerce_TACTIC_COM_to_TARG(pair_list_to_ct undone_cmds); - ct_int_to_TARG remaining_goals; - ct_int_to_TARG current_goal]))))); - traverse_to [] - end;; - -let command_changes = [ - ("TEXT_MODE", - (function - | [VARG_AST (Str(_,x))] -> - (match x with - "fr" -> (function () -> text_proof_flag := "fr") - | "en" -> (function () -> text_proof_flag := "en") - | "off" -> (function () -> text_proof_flag := "off") - | s -> errorlabstrm "TEXT_MODE" (make_error_stream - ("Unexpected flag " ^ s))) - | _ -> errorlabstrm "TEXT_MODE" - (make_error_stream "Unexpected argument"))); - - ("StartProof", - (function - | (VARG_STRING kind) :: - ((VARG_IDENTIFIER s) :: - ((VARG_CONSTR c) :: [])) -> - let stre = - match kind with - | "THEOREM" -> NeverDischarge - | "LEMMA" -> NeverDischarge - | "FACT" -> make_strength_1 () - | "REMARK" -> make_strength_0 () - | "DEFINITION" -> NeverDischarge - | "LET" -> make_strength_2 () - | "LETTOP" -> NotDeclare - | "LOCAL" -> make_strength_0 () - | _ -> errorlabstrm "StartProof" (make_error_stream "StartProof") in - fun () -> - begin - if kind = "LETTOP" && not(refining ()) then - errorlabstrm "StartProof" - (str - "Let declarations can only be used in proof editing mode" - ); - let str = (string_of_id s) in - start_proof_com (Some s) stre c; - History.start_proof str; - current_proof_name := str; - current_goal_index := 1 - end - | _ -> errorlabstrm "StartProof" (make_error_stream "StartProof"))); - - ("GOAL", - (function - | (VARG_CONSTR c) :: [] -> - (fun () -> - if not (refining ()) then - begin - start_proof_com None NeverDischarge c; - History.start_proof "Unnamed_thm"; - current_proof_name := "Unnamed_thm"; - current_goal_index := 1 - end) - | [] -> - (function () -> output_results_nl(ctf_EmptyGoalMessage "")) - | _ -> errorlabstrm "Goal" (make_error_stream "Goal"))); - - ("SOLVE", - (function - | [VARG_NUMBER n; VARG_TACTIC tcom] -> - (fun () -> - if not (refining ()) then - error "Unknown command of the non proof-editing mode"; - solve_nth n (Tacinterp.hide_interp tcom); - let old_n_count = History.border_length !current_proof_name in - let pf = proof_of_pftreestate (get_pftreestate ()) in - let n_goals = (List.length (fst (frontier pf))) + 1 - old_n_count in - begin - current_goal_index := n; - History.push_command !current_proof_name n n_goals - end) - | _ -> errorlabstrm "SOLVE" (make_error_stream "SOLVE"))); - - ("GOAL_CMD", - (function - | (VARG_NUMBER n) :: - ((VARG_TACTIC tac) :: []) -> - (function () -> - let path = History.get_nth_open_path !current_proof_name n in - solve_nth n (Tacinterp.hide_interp tac); - traverse_to path; - Pfedit.mutate weak_undo_pftreestate; - traverse_to []) - | _ -> errorlabstrm "GOAL_CMD" (make_error_stream "GOAL_CMD"))); - - ("KILL_NODE", - (function - | (VARG_NUMBER n) :: [] -> - (function () -> - let ngoals = kill_proof_node n in - output_results_nl - (ctf_KilledMessage !global_request_id ngoals)) - | _ -> errorlabstrm "KILL_NODE" (make_error_stream "KILL_NODE"))); - ("KILL_SUB_PROOF", - (function - | [VARG_NUMBER n] -> - (function () -> logical_kill n) - | _ -> errorlabstrm "KILL_SUB_PROOF" (make_error_stream "KILL_SUB_PROOF"))); - - ("RESUME", - (function [VARG_IDENTIFIER id] -> - (fun () -> - let str = (string_of_id id) in - resume_proof id; - current_proof_name := str) - | _ -> errorlabstrm "RESUME" (make_error_stream "RESUME"))); - - ("BeginSilent", - (function - | [] -> - (function - () -> - errorlabstrm "Begin Silent" - (str "not available in Centaur mode")) - | _ -> errorlabstrm "BeginSilent" (make_error_stream "BeginSilent"))); - - ("EndSilent", - (function - | [] -> - (function - () -> - errorlabstrm "End Silent" - (str "not available in Centaur mode")) - | _ -> errorlabstrm "EndSilent" (make_error_stream "EndSilent"))); - - ("ABORT", - (function - | (VARG_IDENTIFIER id) :: [] -> - (function - () -> - delete_proof id; - current_proof_name := ""; - output_results_nl - (ctf_AbortedMessage !global_request_id (string_of_id id))) - | [] -> - (function - () -> delete_current_proof (); - current_proof_name := ""; - output_results_nl - (ctf_AbortedMessage !global_request_id "")) - | _ -> errorlabstrm "ABORT" (make_error_stream "ABORT"))); - ("SEARCH", - function - (VARG_QUALID qid)::l -> - (fun () -> - ctv_SEARCH_LIST:=[]; - let global_ref = Nametab.global dummy_loc qid in - filtered_search - (filter_by_module_from_varg_list l) - add_search (Nametab.locate qid); - search_output_results()) - | _ -> failwith "bad form of arguments"); - - ("SearchRewrite", - function - (VARG_CONSTR c)::l -> - (fun () -> - ctv_SEARCH_LIST:=[]; - let _,pat = interp_constrpattern Evd.empty (Global.env()) c in - raw_search_rewrite - (filter_by_module_from_varg_list l) - add_search pat; - search_output_results()) - | _ -> failwith "bad form of arguments"); - - ("SearchPattern", - function - (VARG_CONSTR c)::l -> - (fun () -> - ctv_SEARCH_LIST := []; - let _,pat = interp_constrpattern Evd.empty (Global.env()) c in - raw_pattern_search - (filter_by_module_from_varg_list l) - add_search pat; - search_output_results()) - | _ -> failwith "bad form of arguments"); - - ("PrintId", - (function - | [VARG_QUALID qid] -> - (function () -> - let results = xlate_vernac_list (Ctast.ast_to_ct (name_to_ast qid)) in - output_results - (fnl () ++ str "message" ++ fnl () ++ str "PRINT_VALUE" ++ fnl ()) - (Some (P_cl results))) - | _ -> errorlabstrm "PrintId" (make_error_stream "PrintId"))); - - ("Check", - (function - | (VARG_STRING kind) :: ((VARG_CONSTR c) :: g) -> - let evmap, env = - Vernacentries.get_current_context_of_args g in - let f = - match kind with - | "CHECK" -> print_check - | "PRINTTYPE" -> - errorlabstrm "PrintType" (str "Not yet supported in CtCoq") - | _ -> errorlabstrm "CHECK" (make_error_stream "CHECK") in - (function () -> - let a,b = f (c, judgment_of_rawconstr evmap env c) in - output_results a b) - | _ -> errorlabstrm "CHECK" (make_error_stream "CHECK"))); - - ("Eval", - (function - | VARG_TACTIC_ARG(Redexp redexp):: VARG_CONSTR c :: g -> - let evmap, env = Vernacentries.get_current_context_of_args g in - let redfun = - ct_print_eval (Ctast.ast_to_ct c) (Tacred.reduction_of_redexp redexp env evmap) env in - fun () -> - let strm, vtp = redfun evmap (judgment_of_rawconstr evmap env c) in - output_results strm vtp - | _ -> errorlabstrm "Eval" (make_error_stream "Eval"))); - - ("Centaur_Reset", - (function - | (VARG_IDENTIFIER c) :: [] -> - if refining () then - output_results (ctf_AbortedAllMessage ()) None; - current_proof_name := ""; - (match string_of_id c with - | "CtCoqInitialState" -> - (function - () -> - current_proof_name := ""; - Vernacentries.abort_refine Lib.reset_initial (); - output_results (ctf_ResetInitialMessage ()) None) - | _ -> - (function - () -> - current_proof_name := ""; - Vernacentries.abort_refine Lib.reset_name c; - output_results - (ctf_ResetIdentMessage - !global_request_id (string_of_id c)) None)) - | _ -> errorlabstrm "Centaur_Reset" (make_error_stream "Centaur_Reset"))); - ("Show_dad_rules", - (function - | [] -> - (fun () -> - let results = dad_rule_names() in - output_results - (ctf_header "dad_rule_names" !global_request_id) - (Some (P_ids - (CT_id_list - (List.map (fun s -> CT_ident s) results))))) - | _ -> - errorlabstrm - "Show_dad_rules" (make_error_stream "Show_dad_rules"))); - ("INSPECT", - (function - | [VARG_NUMBER n] -> - (function () -> inspect n) - | _ -> errorlabstrm "INSPECT" (make_error_stream "INSPECT"))) - -];; - -let non_debug_changes = [ - ("SHOW", - (function - | [VARG_NUMBER n] -> (function () -> show_nth n) - | _ -> errorlabstrm "SHOW" (make_error_stream "SHOW")))];; - -let command_creations = [ - ("Comments", - function l -> (fun () -> message ("Comments ok\n"))); - ("CommentsBold", - function l -> (fun () -> message ("CommentsBold ok\n"))); - ("Title", - function l -> (fun () -> message ("Title ok\n"))); - ("Author", - function l -> (fun () -> message ("Author ok\n"))); - ("Note", - function l -> (fun () -> message ("Note ok\n"))); - ("NL", - function l -> (fun () -> message ("Newline ok\n")))];; - - - -let start_pcoq_mode debug = - begin - start_dad(); - set_xlate_mut_stuff (fun x ->Ctast.ast_to_ct (globcv (Ctast.ct_to_ast x))); - declare_in_coq(); - add_tactic "PcoqPbp" pbp_tac_pcoq; - add_tactic "PcoqBlast" blast_tac_pcoq; - add_tactic "Dad" dad_tac_pcoq; - add_tactic "CtDebugTac" debug_tac2_pcoq; - add_tactic "CtDebugTac2" debug_tac2_pcoq; -(* The following ones are added to enable rich comments in pcoq *) - add_tactic "Image" (fun _ -> tclIDTAC); - List.iter (fun (a,b) -> vinterp_add a b) command_creations; - List.iter (fun (a,b) -> overwriting_vinterp_add a b) command_changes; - if not debug then - List.iter (fun (a,b) -> overwriting_vinterp_add a b) non_debug_changes; - end;; - -vinterp_add "START_PCOQ" - (function _ -> - (function () -> - start_pcoq_mode false; - set_acknowledge_command ctf_acknowledge_command; - set_start_marker "CENTAUR_RESERVED_TOKEN_start_command"; - set_end_marker "CENTAUR_RESERVED_TOKEN_end_command"; - raise Vernacinterp.ProtectedLoop));; - -vinterp_add "START_PCOQ_DEBUG" - (function _ -> - (function () -> - start_pcoq_mode true; - set_acknowledge_command ctf_acknowledge_command; - set_start_marker "--->"; - set_end_marker "<---"; - raise Vernacinterp.ProtectedLoop));; - diff --git a/contrib/interface/centaur.ml4 b/contrib/interface/centaur.ml4 new file mode 100644 index 000000000..e1ef0ecc1 --- /dev/null +++ b/contrib/interface/centaur.ml4 @@ -0,0 +1,946 @@ +(*i camlp4deps: "parsing/grammar.cma" i*) + +(*Toplevel loop for the communication between Coq and Centaur *) +open Names;; +open Nameops;; +open Util;; +open Ast;; +open Term;; +open Pp;; +open Libobject;; +open Library;; +open Vernacinterp;; +open Evd;; +open Proof_trees;; +open Termast;; +open Tacmach;; +open Pfedit;; +open Proof_type;; +open Parsing;; +open Environ;; +open Declare;; +open Declarations;; +open Rawterm;; +open Reduction;; +open Classops;; +open Vernacinterp;; +open Vernac;; +open Command;; +open Protectedtoplevel;; +open Coqast;; +open Line_oriented_parser;; +open Xlate;; +open Vtp;; +open Ascent;; +open Translate;; +open Name_to_ast;; +open Pbp;; +open Blast;; +open Dad;; +open Debug_tac;; +open Search;; +open Astterm;; +open Nametab;; +open Showproof;; +open Showproof_ct;; +open Tacexpr;; +open Vernacexpr;; + +let pcoq_started = ref None;; + +let if_pcoq f a = + if !pcoq_started <> None then f a else error "Pcoq is not started";; + +let text_proof_flag = ref "en";; + +(* +let current_proof_name = ref "";; +*) +let current_proof_name () = string_of_id (get_current_proof_name ()) + +let current_goal_index = ref 0;; + +set_flags := (function () -> + if List.mem "G_natsyntax" (Mltop.get_loaded_modules()) then + (g_nat_syntax_flag := true; ()) + else ());; + +let guarded_force_eval_stream (s : std_ppcmds) = + let l = ref [] in + let f elt = l:= elt :: !l in + (try Stream.iter f s with + | _ -> f (Stream.next (str "error guarded_force_eval_stream"))); + Stream.of_list (List.rev !l);; + + +let rec string_of_path p = + match p with [] -> "\n" + | i::p -> (string_of_int i)^" "^ (string_of_path p) +;; +let print_path p = + output_results_nl (str "Path:" ++ str (string_of_path p)) +;; + +let kill_proof_node index = + let paths = History.historical_undo (current_proof_name()) index in + let _ = List.iter + (fun path -> (traverse_to path; + Pfedit.mutate weak_undo_pftreestate; + traverse_to [])) + paths in + History.border_length (current_proof_name());; + + +(*Message functions, the text of these messages is recognized by the protocols *) +(*of CtCoq *) +let ctf_header message_name request_id = + fnl () ++ str "message" ++ fnl() ++ str message_name ++ fnl() ++ + int request_id ++ fnl();; + +let ctf_acknowledge_command request_id command_count opt_exn = + let goal_count, goal_index = + if refining() then + let g_count = + List.length + (fst (frontier (proof_of_pftreestate (get_pftreestate ())))) in + g_count, (min g_count !current_goal_index) + else + (0, 0) in + (ctf_header "acknowledge" request_id ++ + int command_count ++ fnl() ++ + int goal_count ++ fnl () ++ + int goal_index ++ fnl () ++ + str (current_proof_name()) ++ fnl() ++ + (match opt_exn with + Some e -> Cerrors.explain_exn e + | None -> mt ()) ++ fnl() ++ str "E-n-d---M-e-s-s-a-g-e" ++ fnl ());; + +let ctf_undoResults = ctf_header "undo_results";; + +let ctf_TextMessage = ctf_header "text_proof";; + +let ctf_SearchResults = ctf_header "search_results";; + +let ctf_OtherGoal = ctf_header "other_goal";; + +let ctf_Location = ctf_header "location";; + +let ctf_StateMessage = ctf_header "state";; + +let ctf_PathGoalMessage () = + fnl () ++ str "message" ++ fnl () ++ str "single_goal" ++ fnl ();; + +let ctf_GoalReqIdMessage = ctf_header "single_goal_state";; + +let ctf_NewStateMessage = ctf_header "fresh_state";; + +let ctf_SavedMessage () = fnl () ++ str "message" ++ fnl () ++ + str "saved" ++ fnl();; + +let ctf_KilledMessage req_id ngoals = + ctf_header "killed" req_id ++ int ngoals ++ fnl ();; + +let ctf_AbortedAllMessage () = + fnl() ++ str "message" ++ fnl() ++ str "aborted_all" ++ fnl();; + +let ctf_AbortedMessage request_id na = + ctf_header "aborted_proof" request_id ++ str na ++ fnl () ++ + str "E-n-d---M-e-s-s-a-g-e" ++ fnl ();; + +let ctf_UserErrorMessage request_id stream = + let stream = guarded_force_eval_stream stream in + ctf_header "user_error" request_id ++ stream ++ fnl() ++ + str "E-n-d---M-e-s-s-a-g-e" ++ fnl();; + +let ctf_ResetInitialMessage () = + fnl () ++ str "message" ++ fnl () ++ str "reset_initial" ++ fnl ();; + +let ctf_ResetIdentMessage request_id s = + ctf_header "reset_ident" request_id ++ str s ++ fnl () ++ + str "E-n-d---M-e-s-s-a-g-e" ++ fnl();; + +type vtp_tree = + | P_rl of ct_RULE_LIST + | P_r of ct_RULE + | P_s_int of ct_SIGNED_INT_LIST + | P_pl of ct_PREMISES_LIST + | P_cl of ct_COMMAND_LIST + | P_t of ct_TACTIC_COM + | P_text of ct_TEXT + | P_ids of ct_ID_LIST;; + +let print_tree t = + (match t with + | P_rl x -> fRULE_LIST x + | P_r x -> fRULE x + | P_s_int x -> fSIGNED_INT_LIST x + | P_pl x -> fPREMISES_LIST x + | P_cl x -> fCOMMAND_LIST x + | P_t x -> fTACTIC_COM x + | P_text x -> fTEXT x + | P_ids x -> fID_LIST x); + print_string "e\nblabla\n";; + + + +let break_happened = ref false;; + +let output_results stream vtp_tree = + let _ = Sys.signal Sys.sigint + (Sys.Signal_handle(fun i -> (break_happened := true;()))) in + msg stream; + match vtp_tree with + Some t -> print_tree t + | None -> ();; + +let output_results_nl stream = + let _ = Sys.signal Sys.sigint + (Sys.Signal_handle(fun i -> break_happened := true;())) + in + msgnl stream;; + + +let rearm_break () = + let _ = Sys.signal Sys.sigint (Sys.Signal_handle(fun i -> raise Sys.Break)) + in ();; + +let check_break () = + if (!break_happened) then + begin + break_happened := false; + raise Sys.Break + end + else ();; + +let print_past_goal index = + let path = History.get_path_for_rank (current_proof_name()) index in + try traverse_to path; + let pf = proof_of_pftreestate (get_pftreestate ()) in + output_results (ctf_PathGoalMessage ()) + (Some (P_r (translate_goal pf.goal))) + with + | Invalid_argument s -> + ((try traverse_to [] with _ -> ()); + error "No focused proof (No proof-editing in progress)") + | e -> (try traverse_to [] with _ -> ()); raise e +;; + +let show_nth n = + try + let pf = proof_of_pftreestate (get_pftreestate()) in + if (!text_proof_flag<>"off") then + (if n=0 + then output_results (ctf_TextMessage !global_request_id) + (Some (P_text (show_proof !text_proof_flag []))) + else + let path = History.get_nth_open_path (current_proof_name()) n in + output_results (ctf_TextMessage !global_request_id) + (Some (P_text (show_proof !text_proof_flag path)))) + else + output_results (ctf_GoalReqIdMessage !global_request_id) + (let goal = List.nth (fst (frontier pf)) + (n - 1) in + (Some (P_r (translate_goal goal)))) + with + | Invalid_argument s -> + error "No focused proof (No proof-editing in progress)";; + +(* The rest of the file contains commands that are changed from the plain + Coq distribution *) + +let ctv_SEARCH_LIST = ref ([] : ct_PREMISE list);; + +(* +let filter_by_module_from_varg_list l = + let dir_list, b = Vernacentries.interp_search_restriction l in + Search.filter_by_module_from_list (dir_list, b);; +*) + +let add_search (global_reference:global_reference) assumptions cstr = + try + let env = Global.env() in + let id_string = + string_of_qualid (Nametab.shortest_qualid_of_global env + global_reference) in + let ast = + try + CT_premise (CT_ident id_string, translate_constr false assumptions cstr) + with Not_found -> + CT_premise (CT_ident id_string, + CT_coerce_ID_to_FORMULA( + CT_ident ("Error printing" ^ id_string))) in + ctv_SEARCH_LIST:= ast::!ctv_SEARCH_LIST + with e -> msgnl (str "add_search raised an exception"); raise e;; + +(* +let make_error_stream node_string = + str "The syntax of " ++ str node_string ++ + str " is inconsistent with the vernac interpreter entry";; +*) + +let ctf_EmptyGoalMessage id = + fnl () ++ str "Empty Goal is a no-op. Fun oh fun." ++ fnl ();; + + +let print_check judg = + let {uj_val=value; uj_type=typ} = judg in + let value_ct_ast = + (try translate_constr false (Global.env()) value + with UserError(f,str) -> + raise(UserError(f, + Ast.print_ast + (ast_of_constr true (Global.env()) value) ++ + fnl () ++ str ))) in + let type_ct_ast = + (try translate_constr false (Global.env()) typ + with UserError(f,str) -> + raise(UserError(f, Ast.print_ast (ast_of_constr true (Global.env()) + value) ++ fnl() ++ str))) in + ((ctf_SearchResults !global_request_id), + (Some (P_pl + (CT_premises_list + [CT_coerce_TYPED_FORMULA_to_PREMISE + (CT_typed_formula(value_ct_ast,type_ct_ast) + )]))));; + +let ct_print_eval ast red_fun env judg = +((if refining() then traverse_to []); +let {uj_val=value; uj_type=typ} = judg in +let nvalue = red_fun value +(* // Attention , ici il faut peut tre utiliser des environnemenst locaux *) +and ntyp = nf_betaiota typ in +(ctf_SearchResults !global_request_id, + Some (P_pl + (CT_premises_list + [CT_eval_result + (xlate_formula ast, + translate_constr false env nvalue, + translate_constr false env ntyp)]))));; + + + +(* The following function is copied from globpr in env/printer.ml *) +let globcv x = + let env = Global.env() in + match x with + | Node(_,"MUTIND", (Path(_,sp))::(Num(_,tyi))::_) -> + let env = Global.env() in + convert_qualid + (Nametab.shortest_qualid_of_global env (IndRef(sp,tyi))) + | Node(_,"MUTCONSTRUCT",(Path(_,sp))::(Num(_,tyi))::(Num(_,i))::_) -> + convert_qualid + (Nametab.shortest_qualid_of_global env + (ConstructRef ((sp, tyi), i))) + | _ -> failwith "globcv : unexpected value";; + +let pbp_tac_pcoq = + pbp_tac (function (x:raw_tactic_expr) -> + output_results + (ctf_header "pbp_results" !global_request_id) + (Some (P_t(xlate_tactic x))));; + +let blast_tac_pcoq = + blast_tac (function (x:raw_tactic_expr) -> + output_results + (ctf_header "pbp_results" !global_request_id) + (Some (P_t(xlate_tactic x))));; + + +let dad_tac_pcoq = + dad_tac(function x -> + output_results + (ctf_header "pbp_results" !global_request_id) + (Some (P_t(xlate_tactic x))));; + +let search_output_results () = + output_results + (ctf_SearchResults !global_request_id) + (Some (P_pl (CT_premises_list + (List.rev !ctv_SEARCH_LIST))));; + + +let debug_tac2_pcoq tac = + (fun g -> + let the_goal = ref (None : goal sigma option) in + let the_ast = ref tac in + let the_path = ref ([] : int list) in + try + let result = report_error tac the_goal the_ast the_path [] g in + (errorlabstrm "DEBUG TACTIC" + (str "no error here " ++ fnl () ++ pr_goal (sig_it g) ++ + fnl () ++ str "the tactic is" ++ fnl () ++ + Pptactic.pr_raw_tactic tac); + result) + with + e -> + match !the_goal with + None -> raise e + | Some g -> + (output_results + (ctf_Location !global_request_id) + (Some (P_s_int + (CT_signed_int_list + (List.map + (fun n -> CT_coerce_INT_to_SIGNED_INT + (CT_int n)) + (clean_path tac + (List.rev !the_path))))))); + (output_results + (ctf_OtherGoal !global_request_id) + (Some (P_r (translate_goal (sig_it g))))); + raise e);; + +let rec selectinspect n env = + match env with + [] -> [] + | a::tl -> + if n = 0 then + [] + else + match a with + (sp, Lib.Leaf lobj) -> a::(selectinspect (n -1 ) tl) + | _ -> (selectinspect n tl);; + +open Term;; + +let inspect n = + let env = Global.env() in + let add_search2 x y = add_search x env y in + let l = selectinspect n (Lib.contents_after None) in + ctv_SEARCH_LIST := []; + List.iter + (fun a -> + try + (match a with + sp, Lib.Leaf lobj -> + (match sp, object_tag lobj with + _, "VARIABLE" -> + let ((_, _, v), _) = get_variable (basename sp) in + add_search2 (Nametab.locate (qualid_of_sp sp)) v + | sp, ("CONSTANT"|"PARAMETER") -> + let {const_type=typ} = Global.lookup_constant sp in + add_search2 (Nametab.locate (qualid_of_sp sp)) typ + | sp, "MUTUALINDUCTIVE" -> + add_search2 (Nametab.locate (qualid_of_sp sp)) + (Pretyping.understand Evd.empty (Global.env()) + (RRef(dummy_loc, IndRef(sp,0)))) + | _ -> failwith ("unexpected value 1 for "^ + (string_of_id (basename sp)))) + | _ -> failwith "unexpected value") + with e -> ()) + l; + output_results + (ctf_SearchResults !global_request_id) + (Some + (P_pl (CT_premises_list (List.rev !ctv_SEARCH_LIST))));; + +let ct_int_to_TARG n = + CT_coerce_ID_OR_INT_to_TARG + (CT_coerce_INT_to_ID_OR_INT (CT_int n));; + +let pair_list_to_ct l = + CT_user_tac(CT_ident "pair_int_list", + CT_targ_list + (List.map (fun (a,b) -> + CT_coerce_TACTIC_COM_to_TARG + (CT_user_tac + (CT_ident "pair_int", + CT_targ_list + [ct_int_to_TARG a; ct_int_to_TARG b]))) + l));; + +(* Annule toutes les commandes qui s'appliquent sur les sous-buts du + but auquel a t applique la n-ime tactique *) +let logical_kill n = + let path = History.get_path_for_rank (current_proof_name()) n in + begin + traverse_to path; + Pfedit.mutate weak_undo_pftreestate; + (let kept_cmds, undone_cmds, remaining_goals, current_goal = + History.logical_undo (current_proof_name()) n in + output_results (ctf_undoResults !global_request_id) + (Some + (P_t + (CT_user_tac + (CT_ident "log_undo_result", + CT_targ_list + [CT_coerce_TACTIC_COM_to_TARG (pair_list_to_ct kept_cmds); + CT_coerce_TACTIC_COM_to_TARG(pair_list_to_ct undone_cmds); + ct_int_to_TARG remaining_goals; + ct_int_to_TARG current_goal]))))); + traverse_to [] + end;; + +let simulate_solve n tac = + let path = History.get_nth_open_path (current_proof_name()) n in + solve_nth n (Tacinterp.hide_interp tac); + traverse_to path; + Pfedit.mutate weak_undo_pftreestate; + traverse_to [] + +let kill_node_verbose n = + let ngoals = kill_proof_node n in + output_results_nl (ctf_KilledMessage !global_request_id ngoals) + +let set_text_mode s = text_proof_flag := s + +VERNAC COMMAND EXTEND TextMode +| [ "Text" "Mode" "fr" ] -> [ set_text_mode "fr" ] +| [ "Text" "Mode" "en" ] -> [ set_text_mode "en" ] +| [ "Text" "Mode" "Off" ] -> [ set_text_mode "off" ] +END + +VERNAC COMMAND EXTEND OutputGoal + [ "Goal" ] -> [ output_results_nl(ctf_EmptyGoalMessage "") ] +END + +VERNAC COMMAND EXTEND OutputGoal + [ "Goal" "Cmd" natural(n) "with" tactic(tac) ] -> [ simulate_solve n tac ] +END + +VERNAC COMMAND EXTEND KillProof +| [ "Kill" "Proof" "after" natural(n) ] -> [ kill_node_verbose n ] +| [ "Kill" "Proof" "at" natural(n) ] -> [ kill_node_verbose n ] +END + +VERNAC COMMAND EXTEND KillSubProof + [ "Kill" "SubProof" natural(n) ] -> [ logical_kill n ] +END + +let start_proof_hook () = + History.start_proof (current_proof_name()); + current_goal_index := 1 + +let solve_hook n = + let name = current_proof_name () in + let old_n_count = History.border_length name in + let pf = proof_of_pftreestate (get_pftreestate ()) in + let n_goals = (List.length (fst (frontier pf))) + 1 - old_n_count in + begin + current_goal_index := n; + History.push_command name n n_goals + end + +let abort_hook s = output_results_nl (ctf_AbortedMessage !global_request_id s) + +let pcoq_search s l = + ctv_SEARCH_LIST:=[]; + begin match s with + | SearchPattern c -> + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + raw_pattern_search (filter_by_module_from_list l) add_search pat + | SearchRewrite c -> + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + raw_search_rewrite (filter_by_module_from_list l) add_search pat; + | SearchHead locqid -> + filtered_search + (filter_by_module_from_list l) add_search (Nametab.global locqid) + end; + search_output_results() + +let pcoq_print_name (_,qid) = + let results = xlate_vernac_list (name_to_ast qid) in + output_results + (fnl () ++ str "message" ++ fnl () ++ str "PRINT_VALUE" ++ fnl ()) + (Some (P_cl results)) + +let pcoq_print_check j = + let a,b = print_check j in output_results a b + +let pcoq_print_eval redfun env c j = + let strm, vtp = ct_print_eval (Ctast.ast_to_ct c) redfun env j in + output_results strm vtp;; + +open Vernacentries + +let pcoq_show_goal = function + | Some n -> show_nth n + | None -> + if !pcoq_started = Some true (* = debug *) then show_open_subgoals () + else errorlabstrm "show_goal" + (str "Show must be followed by an integer in Centaur mode");; + +let pcoq_hook = { + start_proof = start_proof_hook; + solve = solve_hook; + abort = abort_hook; + search = pcoq_search; + print_name = pcoq_print_name; + print_check = pcoq_print_check; + print_eval = pcoq_print_eval; + show_goal = pcoq_show_goal +} + +(* +let command_changes = [ + ("TEXT_MODE", + (function + | [Coqast.VARG_AST (Str(_,x))] -> + (fun () -> set_text_mode x))); + + ("StartProof", + (function + | (Coqast.VARG_STRING kind) :: + ((Coqast.VARG_IDENTIFIER s) :: + ((Coqast.VARG_CONSTR c) :: [])) -> + let stre = + match kind with + | "THEOREM" -> NeverDischarge + | "LEMMA" -> NeverDischarge + | "FACT" -> make_strength_1 () + | "REMARK" -> make_strength_0 () + | "DEFINITION" -> NeverDischarge + | "LET" -> make_strength_2 () + | "LETTOP" -> NotDeclare + | "LOCAL" -> make_strength_0 () + | _ -> errorlabstrm "StartProof" (make_error_stream "StartProof") in + fun () -> + begin + if kind = "LETTOP" && not(refining ()) then + errorlabstrm "StartProof" + (str + "Let declarations can only be used in proof editing mode" + ); + let str = (string_of_id s) in + start_proof_com (Some s) stre c; + start_proof_hook str; + end + | _ -> errorlabstrm "StartProof" (make_error_stream "StartProof"))); + + ("GOAL", + (function + | (Coqast.VARG_CONSTR c) :: [] -> + (fun () -> + if not (refining ()) then + begin + start_proof_com None NeverDischarge c; + start_proof_hook "Unnamed_thm" + end) + | [] -> + (function () -> output_results_nl(ctf_EmptyGoalMessage "")) + | _ -> errorlabstrm "Goal" (make_error_stream "Goal"))); + + ("SOLVE", + (function + | [Coqast.VARG_NUMBER n; Coqast.VARG_TACTIC tcom] -> + (fun () -> + if not (refining ()) then + error "Unknown command of the non proof-editing mode"; + solve_nth n (Tacinterp.hide_interp tcom); +(* pcoq *) + solve_hook n +(**) + | _ -> errorlabstrm "SOLVE" (make_error_stream "SOLVE"))); + +(* SIMULE SOLVE SANS EFFET *) + ("GOAL_CMD", + (function + | (Coqast.VARG_NUMBER n) :: + ((Coqast.VARG_TACTIC tac) :: []) -> + (function () -> + let path = History.get_nth_open_path !current_proof_name n in + solve_nth n (Tacinterp.hide_interp tac); + traverse_to path; + Pfedit.mutate weak_undo_pftreestate; + traverse_to []) + | _ -> errorlabstrm "GOAL_CMD" (make_error_stream "GOAL_CMD"))); + +(* Revient l'tat avant l'application de la n-ime tactique *) + ("KILL_NODE", + (function + | (Coqast.VARG_NUMBER n) :: [] -> + (function () -> + let ngoals = kill_proof_node n in + output_results_nl + (ctf_KilledMessage !global_request_id ngoals)) + | _ -> errorlabstrm "KILL_NODE" (make_error_stream "KILL_NODE"))); +(* Annule toutes les commandes qui s'appliquent sur les sous-buts du + but auquel a t applique la n-ime tactique *) + ("KILL_SUB_PROOF", + (function + | [Coqast.VARG_NUMBER n] -> + (function () -> logical_kill n) + | _ -> errorlabstrm "KILL_SUB_PROOF" (make_error_stream "KILL_SUB_PROOF"))); + + ("RESUME", + (function [Coqast.VARG_IDENTIFIER id] -> + (fun () -> + let str = (string_of_id id) in + resume_proof id; +(* Pcoq *) + current_proof_name := str) +(**) + | _ -> errorlabstrm "RESUME" (make_error_stream "RESUME"))); + +(* NoOp... *) + ("BeginSilent", + (function + | [] -> + (function + () -> + errorlabstrm "Begin Silent" + (str "not available in Centaur mode")) + | _ -> errorlabstrm "BeginSilent" (make_error_stream "BeginSilent"))); + + ("EndSilent", + (function + | [] -> + (function + () -> + errorlabstrm "End Silent" + (str "not available in Centaur mode")) + | _ -> errorlabstrm "EndSilent" (make_error_stream "EndSilent"))); + + ("ABORT", + (function + | (Coqast.VARG_IDENTIFIER id) :: [] -> + (function + () -> + delete_proof id; +(* Pcoq *) + current_proof_name := ""; + output_results_nl + (ctf_AbortedMessage !global_request_id (string_of_id id))) +(**) + | [] -> + (function + () -> delete_current_proof (); +(* Pcoq *) + current_proof_name := ""; + output_results_nl + (ctf_AbortedMessage !global_request_id "")) +(**) + | _ -> errorlabstrm "ABORT" (make_error_stream "ABORT"))); + ("SEARCH", + function + (Coqast.VARG_QUALID qid)::l -> + (fun () -> + ctv_SEARCH_LIST:=[]; + let global_ref = Nametab.global dummy_loc qid in + filtered_search + (filter_by_module_from_varg_list l) + add_search (Nametab.locate qid); + search_output_results()) + | _ -> failwith "bad form of arguments"); + + ("SearchRewrite", + function + (Coqast.VARG_CONSTR c)::l -> + (fun () -> + ctv_SEARCH_LIST:=[]; + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + raw_search_rewrite + (filter_by_module_from_varg_list l) + add_search pat; + search_output_results()) + | _ -> failwith "bad form of arguments"); + + ("SearchPattern", + function + (Coqast.VARG_CONSTR c)::l -> + (fun () -> + ctv_SEARCH_LIST := []; + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + raw_pattern_search + (filter_by_module_from_varg_list l) + add_search pat; + search_output_results()) + | _ -> failwith "bad form of arguments"); + + ("PrintId", + (function + | [Coqast.VARG_QUALID qid] -> + (function () -> + let results = xlate_vernac_list (Ctast.ast_to_ct (name_to_ast qid)) in + output_results + (fnl () ++ str "message" ++ fnl () ++ str "PRINT_VALUE" ++ fnl ()) + (Some (P_cl results))) + | _ -> errorlabstrm "PrintId" (make_error_stream "PrintId"))); + + ("Check", + (function + | (Coqast.VARG_STRING kind) :: ((Coqast.VARG_CONSTR c) :: g) -> + let evmap, env = + Vernacentries.get_current_context_of_args g in + let f = + match kind with + | "CHECK" -> print_check + | "PRINTTYPE" -> + errorlabstrm "PrintType" (str "Not yet supported in CtCoq") + | _ -> errorlabstrm "CHECK" (make_error_stream "CHECK") in + (function () -> + let a,b = f (c, judgment_of_rawconstr evmap env c) in + output_results a b) + | _ -> errorlabstrm "CHECK" (make_error_stream "CHECK"))); + + ("Eval", + (function + | Coqast.VARG_TACTIC_ARG(Redexp redexp):: Coqast.VARG_CONSTR c :: g -> + let evmap, env = Vernacentries.get_current_context_of_args g in + let redfun = + ct_print_eval (Ctast.ast_to_ct c) (Tacred.reduction_of_redexp redexp env evmap) env in + fun () -> + let strm, vtp = redfun evmap (judgment_of_rawconstr evmap env c) in + output_results strm vtp + | _ -> errorlabstrm "Eval" (make_error_stream "Eval"))); + + ("Centaur_Reset", + (function + | (Coqast.VARG_IDENTIFIER c) :: [] -> + if refining () then + output_results (ctf_AbortedAllMessage ()) None; + current_proof_name := ""; + (match string_of_id c with + | "CtCoqInitialState" -> + (function + () -> + current_proof_name := ""; + Vernacentries.abort_refine Lib.reset_initial (); + output_results (ctf_ResetInitialMessage ()) None) + | _ -> + (function + () -> + current_proof_name := ""; + Vernacentries.abort_refine Lib.reset_name c; + output_results + (ctf_ResetIdentMessage + !global_request_id (string_of_id c)) None)) + | _ -> errorlabstrm "Centaur_Reset" (make_error_stream "Centaur_Reset"))); + ("Show_dad_rules", + (function + | [] -> + (fun () -> + let results = dad_rule_names() in + output_results + (ctf_header "dad_rule_names" !global_request_id) + (Some (P_ids + (CT_id_list + (List.map (fun s -> CT_ident s) results))))) + | _ -> + errorlabstrm + "Show_dad_rules" (make_error_stream "Show_dad_rules"))); + ("INSPECT", + (function + | [Coqast.VARG_NUMBER n] -> + (function () -> inspect n) + | _ -> errorlabstrm "INSPECT" (make_error_stream "INSPECT"))) + +];; +*) +(* +let non_debug_changes = [ + ("SHOW", + (function + | [Coqast.VARG_NUMBER n] -> (function () -> show_nth n) + | _ -> errorlabstrm "SHOW" (make_error_stream "SHOW")))];; +*) + +(* Moved to Vernacentries... +let command_creations = [ + ("Comments", + function l -> (fun () -> message ("Comments ok\n"))); +(* Dead code + ("CommentsBold", + function l -> (fun () -> message ("CommentsBold ok\n"))); + ("Title", + function l -> (fun () -> message ("Title ok\n"))); + ("Author", + function l -> (fun () -> message ("Author ok\n"))); + ("Note", + function l -> (fun () -> message ("Note ok\n"))); + ("NL", + function l -> (fun () -> message ("Newline ok\n"))) +*) +];; +*) + +TACTIC EXTEND Pbp +| [ "Pbp" ident_opt(idopt) natural_list(nl) ] -> + [ if_pcoq pbp_tac_pcoq idopt nl ] +END +TACTIC EXTEND Blast +| [ "Blast" ne_natural_list(nl) ] -> + [ if_pcoq blast_tac_pcoq nl ] +END +TACTIC EXTEND Dad +| [ "Dad" natural_list(nl1) "to" natural_list(nl2) ] -> + [ if_pcoq dad_tac_pcoq nl1 nl2 ] +END + +TACTIC EXTEND CtDebugTac +| [ "DebugTac" tactic(t) ] -> [ if_pcoq debug_tac2_pcoq t ] +END + +TACTIC EXTEND CtDebugTac2 +| [ "DebugTac2" tactic(t) ] -> [ if_pcoq debug_tac2_pcoq t ] +END + + +let start_pcoq_mode debug = + begin + pcoq_started := Some debug; + start_dad(); + set_xlate_mut_stuff (fun x ->Ctast.ast_to_ct (globcv (Ctast.ct_to_ast x))); + declare_in_coq(); +(* + add_tactic "PcoqPbp" pbp_tac_pcoq; + add_tactic "PcoqBlast" blast_tac_pcoq; + add_tactic "Dad" dad_tac_pcoq; + add_tactic "CtDebugTac" debug_tac2_pcoq; + add_tactic "CtDebugTac2" debug_tac2_pcoq; +*) +(* The following ones are added to enable rich comments in pcoq *) +(* TODO ... + add_tactic "Image" (fun _ -> tclIDTAC); +*) +(* "Comments" moved to Vernacentries, other obsolete ? + List.iter (fun (a,b) -> vinterp_add a b) command_creations; +*) +(* Now hooks in Vernacentries + List.iter (fun (a,b) -> overwriting_vinterp_add a b) command_changes; + if not debug then + List.iter (fun (a,b) -> overwriting_vinterp_add a b) non_debug_changes; +*) + set_pcoq_hook pcoq_hook; + end;; + +(* +vinterp_add "START_PCOQ" + (function _ -> + (function () -> + start_pcoq_mode false; + set_acknowledge_command ctf_acknowledge_command; + set_start_marker "CENTAUR_RESERVED_TOKEN_start_command"; + set_end_marker "CENTAUR_RESERVED_TOKEN_end_command"; + raise Vernacinterp.ProtectedLoop));; + +vinterp_add "START_PCOQ_DEBUG" + (function _ -> + (function () -> + start_pcoq_mode true; + set_acknowledge_command ctf_acknowledge_command; + set_start_marker "--->"; + set_end_marker "<---"; + raise Vernacinterp.ProtectedLoop));; +*) +let start_pcoq () = + start_pcoq_mode false; + set_acknowledge_command ctf_acknowledge_command; + set_start_marker "CENTAUR_RESERVED_TOKEN_start_command"; + set_end_marker "CENTAUR_RESERVED_TOKEN_end_command"(*; + raise Vernacexpr.ProtectedLoop*) + +let start_pcoq_debug () = + start_pcoq_mode true; + set_acknowledge_command ctf_acknowledge_command; + set_start_marker "--->"; + set_end_marker "<---"(*; + raise Vernacexpr.ProtectedLoop;;*) + +VERNAC COMMAND EXTEND StartPcoq + [ "Start" "Pcoq" "Mode" ] -> [ start_pcoq () ] +END + +VERNAC COMMAND EXTEND StartPcoqDebug +| [ "Start" "Pcoq" "Debug" "Mode" ] -> [ start_pcoq_debug () ] +END diff --git a/contrib/interface/dad.ml b/contrib/interface/dad.ml index ceb0f5953..9d90782e4 100644 --- a/contrib/interface/dad.ml +++ b/contrib/interface/dad.ml @@ -25,6 +25,10 @@ open Pp;; open Paths;; +open Genarg;; +open Tacexpr;; +open Rawterm;; + (* In a first approximation, drag-and-drop rules are like in CtCoq 1/ a pattern, 2,3/ Two paths: start and end positions, @@ -37,7 +41,8 @@ open Paths;; *) -type dad_rule = Ctast.t * int list * int list * int * int list * Ctast.t;; +type dad_rule = + Ctast.t * int list * int list * int * int list * raw_atomic_tactic_expr;; (* This value will be used systematically when constructing objects of type Ctast.t, the value is stupid and meaningless, but it is needed @@ -65,13 +70,20 @@ let rec get_subterm (depth:int) (path: int list) (constr:constr) = let map_subst (env :env) (subst:(int * Term.constr) list) = let rec map_subst_aux = function - Node(_, "META", [Num(_, i)]) -> + Coqast.Node(_, "META", [Coqast.Num(_, i)]) -> let constr = List.assoc i subst in - Ctast.ast_to_ct (ast_of_constr false env constr) - | Node(loc, s, l) -> Node(loc, s, List.map map_subst_aux l) + ast_of_constr false env constr + | Coqast.Node(loc, s, l) -> Coqast.Node(loc, s, List.map map_subst_aux l) | ast -> ast in map_subst_aux;; +let map_subst_tactic env subst = function + | TacExtend ("Rewrite" as x,[b;cbl]) -> + let c,bl = out_gen rawwit_constr_with_bindings cbl in + assert (bl = NoBindings); + let c = (map_subst env subst c,NoBindings) in + TacExtend (x,[b;in_gen rawwit_constr_with_bindings c]) + | _ -> failwith "map_subst_tactic: unsupported tactic" (* This function is really the one that is important. *) let rec find_cmd (l:(string * dad_rule) list) env constr p p1 p2 = @@ -93,7 +105,7 @@ let rec find_cmd (l:(string * dad_rule) list) env constr p p1 p2 = (ct_to_ast pat) in let subst = matches constr_pat term_to_match in if (is_prefix p_f (p_r@p1)) & (is_prefix p_l (p_r@p2)) then - map_subst env subst cmd + TacAtom (zz, map_subst_tactic env subst cmd) else failwith "internal" with @@ -103,7 +115,7 @@ let rec find_cmd (l:(string * dad_rule) list) env constr p p1 p2 = let dad_rule_list = ref ([]: (string * dad_rule) list);; - +(* (* \\ This function is also used in pbp. *) let rec tactic_args_to_ints = function [] -> [] @@ -130,6 +142,12 @@ let dad_tac display_function = function (find_cmd (!dad_rule_list) (pf_env g) (pf_concl g) p_a p1prime p2prime)); tclIDTAC g);; +*) +let dad_tac display_function p1 p2 g = + let (p_a, p1prime, p2prime) = decompose_path (p1,p2) in + (display_function + (find_cmd (!dad_rule_list) (pf_env g) (pf_concl g) p_a p1prime p2prime)); + tclIDTAC g;; (* Now we enter dad rule list management. *) @@ -233,34 +251,40 @@ let rec sort_list = function | a::l -> add_in_list_sorting a (sort_list l);; let mk_dad_meta n = Node(zz,"META",[Num(zz, n)]);; -let mk_rewrite1 ast = - Node(zz,"TACTICLIST", - [Node(zz,"RewriteLR", - [Node(zz, "COMMAND", ast);Node(zz,"BINDINGS",[])])]);; - -let mk_rewrite2 ast = - Node(zz,"TACTICLIST", - [Node(zz,"RewriteRL", - [Node(zz, "COMMAND", ast);Node(zz,"BINDINGS",[])])]);; - - - -let start_dad () = -vinterp_add "AddDadRule" - (function - | [VARG_STRING name; VARG_CONSTR pat; VARG_NUMBERLIST p1; - VARG_NUMBERLIST p2; VARG_TACTIC com] -> - (function () -> - let pr = match decompose_path (p1, p2) with pr, _, _ -> pr in - (add_dad_rule name (Ctast.ast_to_ct pat) p1 p2 (List.length pr) pr (Ctast.ast_to_ct com))) - | _ -> errorlabstrm "AddDadRule1" (str "AddDadRule2")); +let mk_rewrite lr ast = + let b = in_gen rawwit_bool lr in + let cb = in_gen rawwit_constr_with_bindings (Ctast.ct_to_ast ast,NoBindings) in + TacExtend ("Rewrite",[b;cb]) + +open Vernacexpr + +let dad_status = ref false;; + +let start_dad () = dad_status := true;; + +let add_dad_rule_fn name pat p1 p2 tac = + let pr = match decompose_path (p1, p2) with pr, _, _ -> pr in + add_dad_rule name (Ctast.ast_to_ct pat) p1 p2 (List.length pr) pr tac;; + +(* To be parsed by camlp4 + +(*i camlp4deps: "parsing/grammar.cma" i*) + +VERNAC COMMAND EXTEND AddDadRule + [ "Add" "Dad" "Rule" string(name) constr(pat) + "From" natural_list(p1) "To" natural_list(p2) tactic(tac) ] -> + [ add_dad_rule_fn name pat p1 p2 tac ] +END + +*) + add_dad_rule "distributivity-inv" (Node(zz,"APPLIST",[Nvar(zz,"mult");Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,4)]);Node(zz,"META",[Num(zz,3)])]);Node(zz,"META",[Num(zz,2)])])) [2; 2] [2; 1] 1 [2] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "distributivity1-r" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"APPLIST",[Nvar(zz,"mult");Node(zz,"META",[Num(zz,4)]);Node(zz,"META",[Num(zz,2)])]);Node(zz,"APPLIST",[Nvar(zz,"mult");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])])) @@ -268,7 +292,7 @@ add_dad_rule "distributivity1-r" [] 0 [] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "distributivity1-l" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"APPLIST",[Nvar(zz,"mult");Node(zz,"META",[Num(zz,4)]);Node(zz,"META",[Num(zz,2)])]);Node(zz,"APPLIST",[Nvar(zz,"mult");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])])) @@ -276,7 +300,7 @@ add_dad_rule "distributivity1-l" [] 0 [] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "mult_plus_distr");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "associativity" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,4)]);Node(zz,"META",[Num(zz,3)])]);Node(zz,"META",[Num(zz,2)])])) @@ -284,7 +308,7 @@ add_dad_rule "associativity" [] 0 [] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "plus_assoc_r");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "plus_assoc_r");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "minus-identity-lr" (Node(zz,"APPLIST",[Nvar(zz,"minus");Node(zz,"META",[Num(zz,2)]);Node(zz,"META",[Num(zz,2)])])) @@ -292,7 +316,7 @@ add_dad_rule "minus-identity-lr" [2; 2] 1 [2] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "minus_n_n");(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "minus_n_n");(mk_dad_meta 2) ]))); add_dad_rule "minus-identity-rl" (Node(zz,"APPLIST",[Nvar(zz,"minus");Node(zz,"META",[Num(zz,2)]);Node(zz,"META",[Num(zz,2)])])) @@ -300,7 +324,7 @@ add_dad_rule "minus-identity-rl" [2; 1] 1 [2] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "minus_n_n");(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "minus_n_n");(mk_dad_meta 2) ]))); add_dad_rule "plus-sym-rl" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])) @@ -308,7 +332,7 @@ add_dad_rule "plus-sym-rl" [2; 1] 1 [2] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "plus_sym");(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "plus_sym");(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "plus-sym-lr" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])) @@ -316,7 +340,7 @@ add_dad_rule "plus-sym-lr" [2; 2] 1 [2] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "plus_sym");(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "plus_sym");(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "absorb-0-r-rl" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,2)]);Nvar(zz,"O")])) @@ -324,7 +348,7 @@ add_dad_rule "absorb-0-r-rl" [1] 0 [] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "plus_n_O");(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "plus_n_O");(mk_dad_meta 2) ]))); add_dad_rule "absorb-0-r-lr" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,2)]);Nvar(zz,"O")])) @@ -332,7 +356,7 @@ add_dad_rule "absorb-0-r-lr" [2; 2] 0 [] -(mk_rewrite2 ([Node(zz, "APPLIST", [Nvar(zz, "plus_n_O");(mk_dad_meta 2) ])])); +(mk_rewrite false (Node(zz, "APPLIST", [Nvar(zz, "plus_n_O");(mk_dad_meta 2) ]))); add_dad_rule "plus-permute-lr" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,4)]);Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])])) @@ -340,7 +364,7 @@ add_dad_rule "plus-permute-lr" [2; 2; 2; 1] 1 [2] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "plus_permute");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])])); +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "plus_permute");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ]))); add_dad_rule "plus-permute-rl" (Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,4)]);Node(zz,"APPLIST",[Nvar(zz,"plus");Node(zz,"META",[Num(zz,3)]);Node(zz,"META",[Num(zz,2)])])])) @@ -348,7 +372,7 @@ add_dad_rule "plus-permute-rl" [2; 1] 1 [2] -(mk_rewrite1 ([Node(zz, "APPLIST", [Nvar(zz, "plus_permute");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])]));; +(mk_rewrite true (Node(zz, "APPLIST", [Nvar(zz, "plus_permute");(mk_dad_meta 4) ;(mk_dad_meta 3) ;(mk_dad_meta 2) ])));; vinterp_add "StartDad" (function diff --git a/contrib/interface/dad.mli b/contrib/interface/dad.mli index 9f8f7354b..dc2b2734c 100644 --- a/contrib/interface/dad.mli +++ b/contrib/interface/dad.mli @@ -4,7 +4,7 @@ open Tacmach;; val dad_rule_names : unit -> string list;; val start_dad : unit -> unit;; -val dad_tac : (Ctast.t -> 'a) -> tactic_arg list -> goal sigma -> +val dad_tac : (Tacexpr.raw_tactic_expr -> 'a) -> int list -> int list -> goal sigma -> goal list sigma * validation;; val add_dad_rule : string -> Ctast.t -> (int list) -> (int list) -> - int -> (int list) -> Ctast.t -> unit;; + int -> (int list) -> Tacexpr.raw_atomic_tactic_expr -> unit;; diff --git a/contrib/interface/debug_tac.ml b/contrib/interface/debug_tac.ml deleted file mode 100644 index be469108f..000000000 --- a/contrib/interface/debug_tac.ml +++ /dev/null @@ -1,474 +0,0 @@ -open Ast;; -open Coqast;; -open Tacmach;; -open Proof_trees;; -open Pp;; -open Printer;; -open Util;; -open Proof_type;; - -(* Compacting and uncompacting proof commands *) - -type report_tree = - Report_node of bool *int * report_tree list - | Mismatch of int * int - | Tree_fail of report_tree - | Failed of int;; - -type report_card = - Ngoals of int - | Goals_mismatch of int - | Recursive_fail of report_tree - | Fail;; - -type card_holder = report_card ref;; -type report_holder = report_tree list ref;; - -(* This tactical receives an integer and a tactic and checks that the - tactic produces that number of goals. It never fails but signals failure - by updating the boolean reference given as third argument to false. - It is especially suited for use in checked_thens below. *) - -let check_subgoals_count : card_holder -> int -> bool ref -> tactic -> tactic = - fun card_holder count flag t g -> - try - let (gls, v) as result = t g in - let len = List.length (sig_it gls) in - card_holder := - (if len = count then - (flag := true; - Ngoals count) - else - (flag := false; - Goals_mismatch len)); - result - with - e -> card_holder := Fail; - flag := false; - tclIDTAC g;; - -let no_failure = function - [Report_node(true,_,_)] -> true - | _ -> false;; - -let check_subgoals_count2 - : card_holder -> int -> bool ref -> (report_holder -> tactic) -> tactic = - fun card_holder count flag t g -> - let new_report_holder = ref ([] : report_tree list) in - let (gls, v) as result = t new_report_holder g in - let succeeded = no_failure !new_report_holder in - let len = List.length (sig_it gls) in - card_holder := - (if (len = count) & succeeded then - (flag := true; - Ngoals count) - else - (flag := false; - Recursive_fail (List.hd !new_report_holder))); - result;; - -let traceable = function - Node(_, "TACTICLIST", a::b::tl) -> true - | _ -> false;; - -let rec collect_status = function - Report_node(true,_,_)::tl -> collect_status tl - | [] -> true - | _ -> false;; - -(* This tactical receives a tactic and executes it, reporting information - about success in the report holder and a boolean reference. *) - -let count_subgoals : card_holder -> bool ref -> tactic -> tactic = - fun card_holder flag t g -> - try - let (gls, _) as result = t g in - card_holder := (Ngoals(List.length (sig_it gls))); - flag := true; - result - with - e -> card_holder := Fail; - flag := false; - tclIDTAC g;; - -let count_subgoals2 - : card_holder -> bool ref -> (report_holder -> tactic) -> tactic = - fun card_holder flag t g -> - let new_report_holder = ref([] : report_tree list) in - let (gls, v) as result = t new_report_holder g in - let succeeded = no_failure !new_report_holder in - if succeeded then - (flag := true; - card_holder := Ngoals (List.length (sig_it gls))) - else - (flag := false; - card_holder := Recursive_fail(List.hd !new_report_holder)); - result;; - -let rec local_interp : Coqast.t -> report_holder -> tactic = function - Node(_, "TACTICLIST", [a;Node(_, "TACLIST", l)]) -> - (fun report_holder -> checked_thens report_holder a l) - | Node(_, "TACTICLIST", a::((Node(_, "TACLIST", l))as b)::c::tl) -> - local_interp(ope ("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) - | Node(_, "TACTICLIST", [a;b]) -> - (fun report_holder -> checked_then report_holder a b) - | Node(_, "TACTICLIST", a::b::c::tl) -> - local_interp(ope ("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) - | ast -> - (fun report_holder g -> - try - let (gls, _) as result = Tacinterp.interp ast g in - report_holder := (Report_node(true, List.length (sig_it gls), [])) - ::!report_holder; - result - with e -> (report_holder := (Failed 1)::!report_holder; - tclIDTAC g)) - - -(* This tactical receives a tactic and a list of tactics as argument. - It applies the first tactic and then maps the list of tactics to - various produced sub-goals. This tactic will never fail, but reports - are added in the report_holder in the following way: - - In case of partial success, a new report_tree is added to the report_holder - - In case of failure of the first tactic, with no more indications - then Failed 0 is added to the report_holder, - - In case of partial failure of the first tactic then (Failed n) is added to - the report holder. - - In case of success of the first tactic, but count mismatch, then - Mismatch n is added to the report holder. *) - -and checked_thens: report_holder -> Coqast.t -> Coqast.t list -> tactic = - (fun report_holder t1 l g -> - let flag = ref true in - let traceable_t1 = traceable t1 in - let card_holder = ref Fail in - let new_holder = ref ([]:report_tree list) in - let tac_t1 = - if traceable_t1 then - (check_subgoals_count2 card_holder (List.length l) - flag (local_interp t1)) - else - (check_subgoals_count card_holder (List.length l) - flag (Tacinterp.interp t1)) in - let (gls, _) as result = - tclTHEN_i tac_t1 - (fun i -> - if !flag then - (fun g -> - let tac_i = (List.nth l i) in - if traceable tac_i then - local_interp tac_i new_holder g - else - try - let (gls,_) as result = Tacinterp.interp tac_i g in - let len = List.length (sig_it gls) in - new_holder := - (Report_node(true, len, []))::!new_holder; - result - with - e -> (new_holder := (Failed 1)::!new_holder; - tclIDTAC g)) - else - tclIDTAC) g in - let new_goal_list = sig_it gls in - (if !flag then - report_holder := - (Report_node(collect_status !new_holder, - (List.length new_goal_list), - List.rev !new_holder))::!report_holder - else - report_holder := - (match !card_holder with - Goals_mismatch(n) -> Mismatch(n, List.length l) - | Recursive_fail tr -> Tree_fail tr - | Fail -> Failed 1 - | _ -> errorlabstrm "check_thens" - (str "this case should not happen in check_thens")):: - !report_holder); - result) - -(* This tactical receives two tactics as argument, it executes the - first tactic and applies the second one to all the produced goals, - reporting information about the success of all tactics in the report - holder. It never fails. *) - -and checked_then: report_holder -> Coqast.t -> Coqast.t -> tactic = - (fun report_holder t1 t2 g -> - let flag = ref true in - let card_holder = ref Fail in - let tac_t1 = - if traceable t1 then - (count_subgoals2 card_holder flag (local_interp t1)) - else - (count_subgoals card_holder flag (Tacinterp.interp t1)) in - let new_tree_holder = ref ([] : report_tree list) in - let (gls, _) as result = - tclTHEN tac_t1 - (fun (g:goal sigma) -> - if !flag then - if traceable t2 then - local_interp t2 new_tree_holder g - else - try - let (gls, _) as result = Tacinterp.interp t2 g in - new_tree_holder := - (Report_node(true, List.length (sig_it gls),[])):: - !new_tree_holder; - result - with - e -> - (new_tree_holder := ((Failed 1)::!new_tree_holder); - tclIDTAC g) - else - tclIDTAC g) g in - (if !flag then - report_holder := - (Report_node(collect_status !new_tree_holder, - List.length (sig_it gls), - List.rev !new_tree_holder))::!report_holder - else - report_holder := - (match !card_holder with - Recursive_fail tr -> Tree_fail tr - | Fail -> Failed 1 - | _ -> error "this case should not happen in check_then")::!report_holder); - result);; - -(* This tactic applies the given tactic only to those subgoals designated - by the list of integers given as extra arguments. - *) - -let on_then : tactic_arg list -> tactic = function - (Tacexp t1)::(Tacexp t2)::l -> - tclTHEN_i (Tacinterp.interp t1) - (fun i -> - if List.mem (Integer (i + 1)) l then - (Tacinterp.interp t2) - else - tclIDTAC) - | _ -> error "bad arguments for on_then";; - -(* Analyzing error reports *) - -let rec select_success n = function - [] -> [] - | Report_node(true,_,_)::tl -> (Num((0,0),n))::select_success (n+1) tl - | _::tl -> select_success (n+1) tl;; - -let rec expand_tactic = function - Node(loc1, "TACTICLIST", [a;Node(loc2,"TACLIST", l)]) -> - Node(loc1, "TACTICLIST", - [expand_tactic a; - Node(loc2, "TACLIST", List.map expand_tactic l)]) - | Node(loc1, "TACTICLIST", a::((Node(loc2, "TACLIST", l))as b)::c::tl) -> - expand_tactic (Node(loc1, "TACTICLIST", - (Node(loc1, "TACTICLIST", [a;b]))::c::tl)) - | Node(loc1, "TACTICLIST", [a;b]) -> - Node(loc1, "TACTICLIST",[expand_tactic a;expand_tactic b]) - | Node(loc1, "TACTICLIST", a::b::c::tl) -> - expand_tactic (Node(loc1, "TACTICLIST", - (Node(loc1, "TACTICLIST", [a;b]))::c::tl)) - | any -> any;; - -let rec reconstruct_success_tac ast = - match ast with - Node(_, "TACTICLIST", [a;Node(_,"TACLIST",l)]) -> - (function - Report_node(true, n, l) -> ast - | Report_node(false, n, rl) -> - ope("TACTICLIST",[a;ope("TACLIST", - List.map2 reconstruct_success_tac l rl)]) - | Failed n -> ope("Idtac",[]) - | Tree_fail r -> reconstruct_success_tac a r - | Mismatch (n,p) -> a) - | Node(_, "TACTICLIST", [a;b]) -> - (function - Report_node(true, n, l) -> ast - | Report_node(false, n, rl) -> - let selected_indices = select_success 1 rl in - ope("OnThen", a::b::selected_indices) - | Failed n -> ope("Idtac",[]) - | Tree_fail r -> reconstruct_success_tac a r - | _ -> error "this error case should not happen in a THEN tactic") - | _ -> - (function - Report_node(true, n, l) -> ast - | Failed n -> ope("Idtac",[]) - | _ -> - errorlabstrm - "this error case should not happen on an unknown tactic" - (str "error in reconstruction with " ++ fnl () ++ - (gentacpr ast)));; - - -let rec path_to_first_error = function -| Report_node(true, _, l) -> - let rec find_first_error n = function - | (Report_node(true, _, _))::tl -> find_first_error (n + 1) tl - | it::tl -> n, it - | [] -> error "no error detected" in - let p, t = find_first_error 1 l in - p::(path_to_first_error t) -| _ -> [];; - -let rec flatten_then_list tail = function - | Node(_, "TACTICLIST", [a;b]) -> - flatten_then_list ((flatten_then b)::tail) a - | ast -> ast::tail -and flatten_then = function - Node(_, "TACTICLIST", [a;b]) -> - ope("TACTICLIST", flatten_then_list [flatten_then b] a) - | Node(_, "TACLIST", l) -> - ope("TACLIST", List.map flatten_then l) - | Node(_, "OnThen", t1::t2::l) -> - ope("OnThen", (flatten_then t1)::(flatten_then t2)::l) - | ast -> ast;; - -let debug_tac = function - [(Tacexp ast)] -> - (fun g -> - let report = ref ([] : report_tree list) in - let result = local_interp ast report g in - let clean_ast = expand_tactic ast in - let report_tree = - try List.hd !report with - Failure "hd" -> (msgnl (str "report is empty"); Failed 1) in - let success_tac = - reconstruct_success_tac clean_ast report_tree in - let compact_success_tac = - flatten_then success_tac in - msgnl (fnl () ++ - str "========= Successful tactic =============" ++ - fnl () ++ - gentacpr compact_success_tac ++ fnl () ++ - str "========= End of successful tactic ============"); - result) - | _ -> error "wrong arguments for debug_tac";; - -add_tactic "DebugTac" debug_tac;; - -hide_tactic "OnThen" on_then;; - -let rec clean_path p ast l = - match ast, l with - Node(_, "TACTICLIST", ([_;_] as tacs)), fst::tl -> - fst::(clean_path 0 (List.nth tacs (fst - 1)) tl) - | Node(_, "TACTICLIST", tacs), 2::tl -> - let rank = (List.length tacs) - p in - rank::(clean_path 0 (List.nth tacs (rank - 1)) tl) - | Node(_, "TACTICLIST", tacs), 1::tl -> - clean_path (p+1) ast tl - | Node(_, "TACLIST", tacs), fst::tl -> - fst::(clean_path 0 (List.nth tacs (fst - 1)) tl) - | _, [] -> [] - | _, _ -> failwith "this case should not happen in clean_path";; - - - -let rec report_error - : Coqast.t -> goal sigma option ref -> Coqast.t ref -> int list ref -> - int list -> tactic = - fun ast the_goal the_ast returned_path path -> - match ast with - Node(loc1, "TACTICLIST", [a;(Node(loc2, "TACLIST", l)) as tail]) -> - let the_card_holder = ref Fail in - let the_flag = ref false in - let the_exn = ref (Failure "") in - tclTHENS - (fun g -> - let result = - check_subgoals_count - the_card_holder - (List.length l) - the_flag - (fun g2 -> - try - (report_error a the_goal the_ast returned_path (1::path) g2) - with - e -> (the_exn := e; raise e)) - g in - if !the_flag then - result - else - (match !the_card_holder with - Fail -> - the_ast := ope("TACTICLIST", [!the_ast; tail]); - raise !the_exn - | Goals_mismatch p -> - the_ast := ast; - returned_path := path; - error ("Wrong number of tactics: expected " ^ - (string_of_int (List.length l)) ^ " received " ^ - (string_of_int p)) - | _ -> error "this should not happen")) - (let rec fold_num n = function - [] -> [] - | t::tl -> (report_error t the_goal the_ast returned_path (n::2::path)):: - (fold_num (n + 1) tl) in - fold_num 1 l) - | Node(_, "TACTICLIST", a::((Node(_, "TACLIST", l)) as b)::c::tl) -> - report_error(ope("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) - the_goal the_ast returned_path path - | Node(_, "TACTICLIST", [a;b]) -> - let the_count = ref 1 in - tclTHEN - (fun g -> - try - report_error a the_goal the_ast returned_path (1::path) g - with - e -> - (the_ast := ope("TACTICLIST", [!the_ast; b]); - raise e)) - (fun g -> - try - let result = - report_error b the_goal the_ast returned_path (2::path) g in - the_count := !the_count + 1; - result - with - e -> - if !the_count > 1 then - msgnl - (str "in branch no " ++ int !the_count ++ - str " after tactic " ++ gentacpr a); - raise e) - | Node(_, "TACTICLIST", a::b::c::tl) -> - report_error (ope("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) - the_goal the_ast returned_path path - | ast -> - (fun g -> - try - Tacinterp.interp ast g - with - e -> - (the_ast := ast; - the_goal := Some g; - returned_path := path; - raise e));; - -let strip_some = function - Some n -> n - | None -> failwith "No optional value";; - -let descr_first_error = function - [(Tacexp ast)] -> - (fun g -> - let the_goal = ref (None : goal sigma option) in - let the_ast = ref ast in - let the_path = ref ([] : int list) in - try - let result = report_error ast the_goal the_ast the_path [] g in - msgnl (str "no Error here"); - result - with - e -> - (msgnl (str "Execution of this tactic raised message " ++ fnl () ++ - fnl () ++ Cerrors.explain_exn e ++ fnl () ++ - fnl () ++ str "on goal" ++ fnl () ++ - pr_goal (sig_it (strip_some !the_goal)) ++ fnl () ++ - str "faulty tactic is" ++ fnl () ++ fnl () ++ - gentacpr (flatten_then !the_ast) ++ fnl ()); - tclIDTAC g)) - | _ -> error "wrong arguments for descr_first_error";; - -add_tactic "DebugTac2" descr_first_error;; diff --git a/contrib/interface/debug_tac.ml4 b/contrib/interface/debug_tac.ml4 new file mode 100644 index 000000000..b4db22803 --- /dev/null +++ b/contrib/interface/debug_tac.ml4 @@ -0,0 +1,570 @@ +(*i camlp4deps: "parsing/grammar.cma" i*) + +open Ast;; +open Coqast;; +open Tacmach;; +open Tacticals;; +open Proof_trees;; +open Pp;; +open Pptactic;; +open Util;; +open Proof_type;; +open Tacexpr;; +open Genarg;; + +(* Compacting and uncompacting proof commands *) + +type report_tree = + Report_node of bool *int * report_tree list + | Mismatch of int * int + | Tree_fail of report_tree + | Failed of int;; + +type report_card = + Ngoals of int + | Goals_mismatch of int + | Recursive_fail of report_tree + | Fail;; + +type card_holder = report_card ref;; +type report_holder = report_tree list ref;; + +(* This tactical receives an integer and a tactic and checks that the + tactic produces that number of goals. It never fails but signals failure + by updating the boolean reference given as third argument to false. + It is especially suited for use in checked_thens below. *) + +let check_subgoals_count : card_holder -> int -> bool ref -> tactic -> tactic = + fun card_holder count flag t g -> + try + let (gls, v) as result = t g in + let len = List.length (sig_it gls) in + card_holder := + (if len = count then + (flag := true; + Ngoals count) + else + (flag := false; + Goals_mismatch len)); + result + with + e -> card_holder := Fail; + flag := false; + tclIDTAC g;; + +let no_failure = function + [Report_node(true,_,_)] -> true + | _ -> false;; + +let check_subgoals_count2 + : card_holder -> int -> bool ref -> (report_holder -> tactic) -> tactic = + fun card_holder count flag t g -> + let new_report_holder = ref ([] : report_tree list) in + let (gls, v) as result = t new_report_holder g in + let succeeded = no_failure !new_report_holder in + let len = List.length (sig_it gls) in + card_holder := + (if (len = count) & succeeded then + (flag := true; + Ngoals count) + else + (flag := false; + Recursive_fail (List.hd !new_report_holder))); + result;; + +(* +let traceable = function + Node(_, "TACTICLIST", a::b::tl) -> true + | _ -> false;; +*) +let traceable = function + | TacThen _ | TacThens _ -> true + | _ -> false;; + +let rec collect_status = function + Report_node(true,_,_)::tl -> collect_status tl + | [] -> true + | _ -> false;; + +(* This tactical receives a tactic and executes it, reporting information + about success in the report holder and a boolean reference. *) + +let count_subgoals : card_holder -> bool ref -> tactic -> tactic = + fun card_holder flag t g -> + try + let (gls, _) as result = t g in + card_holder := (Ngoals(List.length (sig_it gls))); + flag := true; + result + with + e -> card_holder := Fail; + flag := false; + tclIDTAC g;; + +let count_subgoals2 + : card_holder -> bool ref -> (report_holder -> tactic) -> tactic = + fun card_holder flag t g -> + let new_report_holder = ref([] : report_tree list) in + let (gls, v) as result = t new_report_holder g in + let succeeded = no_failure !new_report_holder in + if succeeded then + (flag := true; + card_holder := Ngoals (List.length (sig_it gls))) + else + (flag := false; + card_holder := Recursive_fail(List.hd !new_report_holder)); + result;; + +let rec local_interp : raw_tactic_expr -> report_holder -> tactic = function +(* + Node(_, "TACTICLIST", [a;Node(_, "TACLIST", l)]) -> + (fun report_holder -> checked_thens report_holder a l) + | Node(_, "TACTICLIST", a::((Node(_, "TACLIST", l))as b)::c::tl) -> + local_interp(ope ("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) + | Node(_, "TACTICLIST", [a;b]) -> + (fun report_holder -> checked_then report_holder a b) + | Node(_, "TACTICLIST", a::b::c::tl) -> + local_interp(ope ("TACTICLIST", (ope("TACTICLIST", [a;b]))::c::tl)) + | ast -> + (fun report_holder g -> + try + let (gls, _) as result = Tacinterp.interp ast g in + report_holder := (Report_node(true, List.length (sig_it gls), [])) + ::!report_holder; + result + with e -> (report_holder := (Failed 1)::!report_holder; + tclIDTAC g)) +*) + TacThens (a,l) -> + (fun report_holder -> checked_thens report_holder a l) + | TacThen (a,b) -> + (fun report_holder -> checked_then report_holder a b) + | t -> + (fun report_holder g -> + try + let (gls, _) as result = Tacinterp.interp t g in + report_holder := (Report_node(true, List.length (sig_it gls), [])) + ::!report_holder; + result + with e -> (report_holder := (Failed 1)::!report_holder; + tclIDTAC g)) + + +(* This tactical receives a tactic and a list of tactics as argument. + It applies the first tactic and then maps the list of tactics to + various produced sub-goals. This tactic will never fail, but reports + are added in the report_holder in the following way: + - In case of partial success, a new report_tree is added to the report_holder + - In case of failure of the first tactic, with no more indications + then Failed 0 is added to the report_holder, + - In case of partial failure of the first tactic then (Failed n) is added to + the report holder. + - In case of success of the first tactic, but count mismatch, then + Mismatch n is added to the report holder. *) + +and checked_thens: report_holder -> raw_tactic_expr -> raw_tactic_expr list -> tactic = + (fun report_holder t1 l g -> + let flag = ref true in + let traceable_t1 = traceable t1 in + let card_holder = ref Fail in + let new_holder = ref ([]:report_tree list) in + let tac_t1 = + if traceable_t1 then + (check_subgoals_count2 card_holder (List.length l) + flag (local_interp t1)) + else + (check_subgoals_count card_holder (List.length l) + flag (Tacinterp.interp t1)) in + let (gls, _) as result = + tclTHEN_i tac_t1 + (fun i -> + if !flag then + (fun g -> + let tac_i = (List.nth l i) in + if traceable tac_i then + local_interp tac_i new_holder g + else + try + let (gls,_) as result = Tacinterp.interp tac_i g in + let len = List.length (sig_it gls) in + new_holder := + (Report_node(true, len, []))::!new_holder; + result + with + e -> (new_holder := (Failed 1)::!new_holder; + tclIDTAC g)) + else + tclIDTAC) g in + let new_goal_list = sig_it gls in + (if !flag then + report_holder := + (Report_node(collect_status !new_holder, + (List.length new_goal_list), + List.rev !new_holder))::!report_holder + else + report_holder := + (match !card_holder with + Goals_mismatch(n) -> Mismatch(n, List.length l) + | Recursive_fail tr -> Tree_fail tr + | Fail -> Failed 1 + | _ -> errorlabstrm "check_thens" + (str "this case should not happen in check_thens")):: + !report_holder); + result) + +(* This tactical receives two tactics as argument, it executes the + first tactic and applies the second one to all the produced goals, + reporting information about the success of all tactics in the report + holder. It never fails. *) + +and checked_then: report_holder -> raw_tactic_expr -> raw_tactic_expr -> tactic = + (fun report_holder t1 t2 g -> + let flag = ref true in + let card_holder = ref Fail in + let tac_t1 = + if traceable t1 then + (count_subgoals2 card_holder flag (local_interp t1)) + else + (count_subgoals card_holder flag (Tacinterp.interp t1)) in + let new_tree_holder = ref ([] : report_tree list) in + let (gls, _) as result = + tclTHEN tac_t1 + (fun (g:goal sigma) -> + if !flag then + if traceable t2 then + local_interp t2 new_tree_holder g + else + try + let (gls, _) as result = Tacinterp.interp t2 g in + new_tree_holder := + (Report_node(true, List.length (sig_it gls),[])):: + !new_tree_holder; + result + with + e -> + (new_tree_holder := ((Failed 1)::!new_tree_holder); + tclIDTAC g) + else + tclIDTAC g) g in + (if !flag then + report_holder := + (Report_node(collect_status !new_tree_holder, + List.length (sig_it gls), + List.rev !new_tree_holder))::!report_holder + else + report_holder := + (match !card_holder with + Recursive_fail tr -> Tree_fail tr + | Fail -> Failed 1 + | _ -> error "this case should not happen in check_then")::!report_holder); + result);; + +(* This tactic applies the given tactic only to those subgoals designated + by the list of integers given as extra arguments. + *) + +let on_then = function [t1;t2;l] -> + let t1 = out_gen wit_tactic t1 in + let t2 = out_gen wit_tactic t2 in + let l = out_gen (wit_list0 wit_int) l in + tclTHEN_i (Tacinterp.interp t1) + (fun i -> + if List.mem (i + 1) l then + (Tacinterp.interp t2) + else + tclIDTAC) + | _ -> anomaly "bad arguments for on_then";; + +let mkOnThen t1 t2 selected_indices = + let a = in_gen rawwit_tactic t1 in + let b = in_gen rawwit_tactic t2 in + let l = in_gen (wit_list0 rawwit_int) selected_indices in + TacAtom (dummy_loc, TacExtend ("OnThen", [a;b;l]));; + +(* Analyzing error reports *) + +(* +let rec select_success n = function + [] -> [] + | Report_node(true,_,_)::tl -> (Num((0,0),n))::select_success (n+1) tl + | _::tl -> select_success (n+1) tl;; +*) +let rec select_success n = function + [] -> [] + | Report_node(true,_,_)::tl -> n::select_success (n+1) tl + | _::tl -> select_success (n+1) tl;; + +(* +let rec expand_tactic = function + Node(loc1, "TACTICLIST", [a;Node(loc2,"TACLIST", l)]) -> + Node(loc1, "TACTICLIST", + [expand_tactic a; + Node(loc2, "TACLIST", List.map expand_tactic l)]) + | Node(loc1, "TACTICLIST", a::((Node(loc2, "TACLIST", l))as b)::c::tl) -> + expand_tactic (Node(loc1, "TACTICLIST", + (Node(loc1, "TACTICLIST", [a;b]))::c::tl)) + | Node(loc1, "TACTICLIST", [a;b]) -> + Node(loc1, "TACTICLIST",[expand_tactic a;expand_tactic b]) + | Node(loc1, "TACTICLIST", a::b::c::tl) -> + expand_tactic (Node(loc1, "TACTICLIST", + (Node(loc1, "TACTICLIST", [a;b]))::c::tl)) + | any -> any;; +*) +(* Useless: already in binary form... +let rec expand_tactic = function + TacThens (a,l) -> TacThens (expand_tactic a, List.map expand_tactic l) + | TacThen (a,b) -> TacThen (expand_tactic a, expand_tactic b) + | any -> any;; +*) + +(* +let rec reconstruct_success_tac ast = + match ast with + Node(_, "TACTICLIST", [a;Node(_,"TACLIST",l)]) -> + (function + Report_node(true, n, l) -> ast + | Report_node(false, n, rl) -> + ope("TACTICLIST",[a;ope("TACLIST", + List.map2 reconstruct_success_tac l rl)]) + | Failed n -> ope("Idtac",[]) + | Tree_fail r -> reconstruct_success_tac a r + | Mismatch (n,p) -> a) + | Node(_, "TACTICLIST", [a;b]) -> + (function + Report_node(true, n, l) -> ast + | Report_node(false, n, rl) -> + let selected_indices = select_success 1 rl in + ope("OnThen", a::b::selected_indices) + | Failed n -> ope("Idtac",[]) + | Tree_fail r -> reconstruct_success_tac a r + | _ -> error "this error case should not happen in a THEN tactic") + | _ -> + (function + Report_node(true, n, l) -> ast + | Failed n -> ope("Idtac",[]) + | _ -> + errorlabstrm + "this error case should not happen on an unknown tactic" + (str "error in reconstruction with " ++ fnl () ++ + (gentacpr ast)));; +*) +let rec reconstruct_success_tac tac = + match tac with + TacThens (a,l) -> + (function + Report_node(true, n, l) -> tac + | Report_node(false, n, rl) -> + TacThens (a,List.map2 reconstruct_success_tac l rl) + | Failed n -> TacId + | Tree_fail r -> reconstruct_success_tac a r + | Mismatch (n,p) -> a) + | TacThen (a,b) -> + (function + Report_node(true, n, l) -> tac + | Report_node(false, n, rl) -> + let selected_indices = select_success 1 rl in + TacAtom (Ast.dummy_loc,TacExtend ("OnThen", + [in_gen rawwit_tactic a; + in_gen rawwit_tactic b; + in_gen (wit_list0 rawwit_int) selected_indices])) + | Failed n -> TacId + | Tree_fail r -> reconstruct_success_tac a r + | _ -> error "this error case should not happen in a THEN tactic") + | _ -> + (function + Report_node(true, n, l) -> tac + | Failed n -> TacId + | _ -> + errorlabstrm + "this error case should not happen on an unknown tactic" + (str "error in reconstruction with " ++ fnl () ++ + (pr_raw_tactic tac)));; + + +let rec path_to_first_error = function +| Report_node(true, _, l) -> + let rec find_first_error n = function + | (Report_node(true, _, _))::tl -> find_first_error (n + 1) tl + | it::tl -> n, it + | [] -> error "no error detected" in + let p, t = find_first_error 1 l in + p::(path_to_first_error t) +| _ -> [];; + +(* +let rec flatten_then_list tail = function + | Node(_, "TACTICLIST", [a;b]) -> + flatten_then_list ((flatten_then b)::tail) a + | ast -> ast::tail +and flatten_then = function + Node(_, "TACTICLIST", [a;b]) -> + ope("TACTICLIST", flatten_then_list [flatten_then b] a) + | Node(_, "TACLIST", l) -> + ope("TACLIST", List.map flatten_then l) + | Node(_, "OnThen", t1::t2::l) -> + ope("OnThen", (flatten_then t1)::(flatten_then t2)::l) + | ast -> ast;; +*) + +let debug_tac = function + [(Tacexp ast)] -> + (fun g -> + let report = ref ([] : report_tree list) in + let result = local_interp ast report g in + let clean_ast = (* expand_tactic *) ast in + let report_tree = + try List.hd !report with + Failure "hd" -> (msgnl (str "report is empty"); Failed 1) in + let success_tac = + reconstruct_success_tac clean_ast report_tree in + let compact_success_tac = (* flatten_then *) success_tac in + msgnl (fnl () ++ + str "========= Successful tactic =============" ++ + fnl () ++ + pr_raw_tactic compact_success_tac ++ fnl () ++ + str "========= End of successful tactic ============"); + result) + | _ -> error "wrong arguments for debug_tac";; + +(* TODO ... used ? +add_tactic "DebugTac" debug_tac;; +*) + +(* +hide_tactic "OnThen" on_then;; +*) +Refiner.add_tactic "OnThen" on_then;; + +(* +let rec clean_path p ast l = + match ast, l with + Node(_, "TACTICLIST", ([_;_] as tacs)), fst::tl -> + fst::(clean_path 0 (List.nth tacs (fst - 1)) tl) + | Node(_, "TACTICLIST", tacs), 2::tl -> + let rank = (List.length tacs) - p in + rank::(clean_path 0 (List.nth tacs (rank - 1)) tl) + | Node(_, "TACTICLIST", tacs), 1::tl -> + clean_path (p+1) ast tl + | Node(_, "TACLIST", tacs), fst::tl -> + fst::(clean_path 0 (List.nth tacs (fst - 1)) tl) + | _, [] -> [] + | _, _ -> failwith "this case should not happen in clean_path";; +*) +let rec clean_path tac l = + match tac, l with + | TacThen (a,b), fst::tl -> + fst::(clean_path (if fst = 1 then a else b) tl) + | TacThens (a,l), 1::tl -> + 1::(clean_path a tl) + | TacThens (a,tacs), 2::fst::tl -> + 2::fst::(clean_path (List.nth tacs (fst - 1)) tl) + | _, [] -> [] + | _, _ -> failwith "this case should not happen in clean_path";; + +let rec report_error + : raw_tactic_expr -> goal sigma option ref -> raw_tactic_expr ref -> int list ref -> + int list -> tactic = + fun tac the_goal the_ast returned_path path -> + match tac with + TacThens (a,l) -> + let the_card_holder = ref Fail in + let the_flag = ref false in + let the_exn = ref (Failure "") in + tclTHENS + (fun g -> + let result = + check_subgoals_count + the_card_holder + (List.length l) + the_flag + (fun g2 -> + try + (report_error a the_goal the_ast returned_path (1::path) g2) + with + e -> (the_exn := e; raise e)) + g in + if !the_flag then + result + else + (match !the_card_holder with + Fail -> + the_ast := TacThens (!the_ast, l); + raise !the_exn + | Goals_mismatch p -> + the_ast := tac; + returned_path := path; + error ("Wrong number of tactics: expected " ^ + (string_of_int (List.length l)) ^ " received " ^ + (string_of_int p)) + | _ -> error "this should not happen")) + (let rec fold_num n = function + [] -> [] + | t::tl -> (report_error t the_goal the_ast returned_path (n::2::path)):: + (fold_num (n + 1) tl) in + fold_num 1 l) + | TacThen (a,b) -> + let the_count = ref 1 in + tclTHEN + (fun g -> + try + report_error a the_goal the_ast returned_path (1::path) g + with + e -> + (the_ast := TacThen (!the_ast, b); + raise e)) + (fun g -> + try + let result = + report_error b the_goal the_ast returned_path (2::path) g in + the_count := !the_count + 1; + result + with + e -> + if !the_count > 1 then + msgnl + (str "in branch no " ++ int !the_count ++ + str " after tactic " ++ pr_raw_tactic a); + raise e) + | tac -> + (fun g -> + try + Tacinterp.interp tac g + with + e -> + (the_ast := tac; + the_goal := Some g; + returned_path := path; + raise e));; + +let strip_some = function + Some n -> n + | None -> failwith "No optional value";; + +let descr_first_error tac = + (fun g -> + let the_goal = ref (None : goal sigma option) in + let the_ast = ref tac in + let the_path = ref ([] : int list) in + try + let result = report_error tac the_goal the_ast the_path [] g in + msgnl (str "no Error here"); + result + with + e -> + (msgnl (str "Execution of this tactic raised message " ++ fnl () ++ + fnl () ++ Cerrors.explain_exn e ++ fnl () ++ + fnl () ++ str "on goal" ++ fnl () ++ + pr_goal (sig_it (strip_some !the_goal)) ++ fnl () ++ + str "faulty tactic is" ++ fnl () ++ fnl () ++ + pr_raw_tactic ((*flatten_then*) !the_ast) ++ fnl ()); + tclIDTAC g)) + +(* TODO ... used ?? +add_tactic "DebugTac2" descr_first_error;; +*) + +(* +TACTIC EXTEND DebugTac2 + [ ??? ] -> [ descr_first_error tac ] +END +*) diff --git a/contrib/interface/debug_tac.mli b/contrib/interface/debug_tac.mli index a582b029e..05cef23aa 100644 --- a/contrib/interface/debug_tac.mli +++ b/contrib/interface/debug_tac.mli @@ -1,6 +1,6 @@ -val report_error : Coqast.t -> +val report_error : Tacexpr.raw_tactic_expr -> Proof_type.goal Proof_type.sigma option ref -> - Coqast.t ref -> int list ref -> int list -> Tacmach.tactic;; + Tacexpr.raw_tactic_expr ref -> int list ref -> int list -> Tacmach.tactic;; -val clean_path : int -> Coqast.t -> int list -> int list;; +val clean_path : Tacexpr.raw_tactic_expr -> int list -> int list;; diff --git a/contrib/interface/name_to_ast.ml b/contrib/interface/name_to_ast.ml index a7d5644f0..0f3dcdf67 100644 --- a/contrib/interface/name_to_ast.ml +++ b/contrib/interface/name_to_ast.ml @@ -17,7 +17,7 @@ open Util;; open Pp;; open Declare;; open Nametab - +open Vernacexpr;; (* This function converts the parameter binders of an inductive definition, in particular you have to be careful to handle each element in the @@ -26,9 +26,7 @@ open Nametab let convert_env = let convert_binder env (na, _, c) = match na with - | Name id -> - ope("BINDER", - [ast_of_constr true env c;nvar id]) + | Name id -> (id, ast_of_constr true env c) | Anonymous -> failwith "anomaly: Anonymous variables in inductives" in let rec cvrec env = function [] -> [] @@ -56,9 +54,9 @@ let impl_args_to_string = function let implicit_args_id_to_ast_list id l ast_list = (match impl_args_to_string l with None -> ast_list - | Some(s) -> (string ("For " ^ (string_of_id id))):: - (string s):: - ast_list);; + | Some(s) -> CommentString ("For " ^ (string_of_id id)):: + CommentString s:: + ast_list);; (* This function construct an ast to enumerate the implicit positions for an inductive type and its constructors. It is obtained directly from @@ -66,7 +64,7 @@ let implicit_args_id_to_ast_list id l ast_list = let implicit_args_to_ast_list sp mipv = let implicit_args_descriptions = - let ast_list = ref ([]:Coqast.t list) in + let ast_list = ref [] in (Array.iteri (fun i mip -> let imps = inductive_implicits_list(sp, i) in @@ -82,7 +80,7 @@ let implicit_args_to_ast_list sp mipv = !ast_list) in match implicit_args_descriptions with [] -> [] - | _ -> [ope("COMMENT", List.rev implicit_args_descriptions)];; + | _ -> [VernacComments (List.rev implicit_args_descriptions)];; let convert_qualid qid = let d, id = Nametab.repr_qualid qid in @@ -97,10 +95,11 @@ let convert_qualid qid = let convert_constructors envpar names types = let array_idC = array_map2 - (fun n t -> ope("BINDER", - [ast_of_constr true envpar t; nvar n])) + (fun n t -> + let coercion_flag = false (* arbitrary *) in + (n, coercion_flag, ast_of_constr true envpar t)) names types in - Node((0,0), "BINDERLIST", Array.to_list array_idC);; + Array.to_list array_idC;; (* this function converts one inductive type in a possibly multiple inductive definition *) @@ -109,11 +108,11 @@ let convert_one_inductive sp tyi = let (ref, params, arity, cstrnames, cstrtypes) = build_inductive sp tyi in let env = Global.env () in let envpar = push_rel_context params env in - ope("VERNACARGLIST", - [convert_qualid (Nametab.shortest_qualid_of_global env (IndRef (sp, tyi))); - ope("CONSTR", [ast_of_constr true envpar arity]); - ope("BINDERLIST", convert_env(List.rev params)); - convert_constructors envpar cstrnames cstrtypes]);; + let sp = sp_of_global env (IndRef (sp, tyi)) in + (basename sp, + convert_env(List.rev params), + (ast_of_constr true envpar arity), + convert_constructors envpar cstrnames cstrtypes);; (* This function converts a Mutual inductive definition to a Coqast.t. It is obtained directly from print_mutual in pretty.ml. However, all @@ -121,22 +120,21 @@ let convert_one_inductive sp tyi = let mutual_to_ast_list sp mib = let mipv = (Global.lookup_mind sp).mind_packets in - let _, ast_list = + let _, l = Array.fold_right (fun mi (n,l) -> (n+1, (convert_one_inductive sp n)::l)) mipv (0, []) in - (ope("MUTUALINDUCTIVE", - [string (if mib.mind_finite then "Inductive" else "CoInductive"); - ope("VERNACARGLIST", ast_list)]):: - (implicit_args_to_ast_list sp mipv));; + VernacInductive (mib.mind_finite, l) + :: (implicit_args_to_ast_list sp mipv);; let constr_to_ast v = ast_of_constr true (Global.env()) v;; let implicits_to_ast_list implicits = - (match (impl_args_to_string implicits) with - None -> [] - | Some s -> [ope("COMMENT", [string s])]);; + match (impl_args_to_string implicits) with + | None -> [] + | Some s -> [VernacComments [CommentString s]];; +(* let make_variable_ast name typ implicits = (ope("VARIABLE", [string "VARIABLE"; @@ -145,7 +143,13 @@ let make_variable_ast name typ implicits = [(constr_to_ast (body_of_type typ)); nvar name])])]))::(implicits_to_ast_list implicits) ;; +*) +let make_variable_ast name typ implicits = + (VernacAssumption + (AssumptionVariable, [name, constr_to_ast (body_of_type typ)])) + ::(implicits_to_ast_list implicits);; +(* let make_definition_ast name c typ implicits = (ope("DEFINITION", [string "DEFINITION"; @@ -155,6 +159,12 @@ let make_definition_ast name c typ implicits = [(constr_to_ast c); (constr_to_ast (body_of_type typ))])])])):: (implicits_to_ast_list implicits);; +*) +let make_definition_ast name c typ implicits = + VernacDefinition (Definition, name, None, + (ope("CAST", [(constr_to_ast c);(constr_to_ast (body_of_type typ))])), + None, (fun _ _ -> ())) + ::(implicits_to_ast_list implicits);; (* This function is inspired by print_constant *) let constant_to_ast_list sp = @@ -238,5 +248,5 @@ let name_to_ast (qid:Nametab.qualid) = (Nametab.pr_qualid qid ++ spc () ++ str "not a defined object") in - ope("vernac_list", l);; + VernacList (List.map (fun x -> (dummy_loc,x)) l) diff --git a/contrib/interface/name_to_ast.mli b/contrib/interface/name_to_ast.mli index 894b27392..4a68e0134 100644 --- a/contrib/interface/name_to_ast.mli +++ b/contrib/interface/name_to_ast.mli @@ -1,2 +1,2 @@ -val name_to_ast : Nametab.qualid -> Coqast.t;; +val name_to_ast : Nametab.qualid -> Vernacexpr.vernac_expr;; val convert_qualid : Nametab.qualid -> Coqast.t;; diff --git a/contrib/interface/parse.ml b/contrib/interface/parse.ml index 389fa8cda..5bce60f22 100644 --- a/contrib/interface/parse.ml +++ b/contrib/interface/parse.ml @@ -18,6 +18,8 @@ open Line_parser;; open Pcoq;; +open Vernacexpr;; + type parsed_tree = | P_cl of ct_COMMAND_LIST | P_c of ct_COMMAND @@ -55,6 +57,7 @@ let ctf_FileErrorMessage reqid pps = int reqid ++ fnl () ++ pps ++ fnl () ++ str "E-n-d---M-e-s-s-a-g-e" ++ fnl ();; +(* (*In the code for CoqV6.2, the require_module call is encapsulated in a function "without_mes_ambig". Here I have supposed that this function has no effect on parsing *) @@ -62,23 +65,29 @@ let try_require_module import specif names = try Library.require_module (if specif = "UNSPECIFIED" then None else Some (specif = "SPECIFICATION")) - (List.map (fun name -> Nametab.make_short_qualid (Names.id_of_string name)) + (List.map + (fun name -> + (dummy_loc,Nametab.make_short_qualid (Names.id_of_string name))) names) (import = "IMPORT") with | e -> msgnl (str "Reinterning of " ++ prlist str names ++ str " failed");; - +*) +(* let try_require_module_from_file import specif name fname = try Library.require_module_from_file (if specif = "UNSPECIFIED" then None - else Some (specif = "SPECIFICATION")) (Nametab.make_short_qualid (Names.id_of_string name)) fname (import = "IMPORT") + else Some (specif = "SPECIFICATION")) (Some (Names.id_of_string name)) fname (import = "IMPORT") with | e -> msgnl (str "Reinterning of " ++ str name ++ str " failed");; - +*) +(* let execute_when_necessary ast = (match ast with | Node (_, "GRAMMAR", ((Nvar (_, s)) :: ((Node (_, "ASTLIST", al)) :: []))) -> Metasyntax.add_grammar_obj s (List.map Ctast.ct_to_ast al) +(* Obsolete | Node (_, "TOKEN", ((Str (_, s)) :: [])) -> Metasyntax.add_token_obj s +*) | Node (_, "Require", ((Str (_, import)) :: ((Str (_, specif)) :: l))) -> @@ -92,6 +101,23 @@ let execute_when_necessary ast = ((Nvar (_, mname)) :: ((Str (_, file_name)) :: []))))) -> try_require_module_from_file import specif mname file_name | _ -> ()); ast;; +*) + +let execute_when_necessary v = + (match v with + | VernacGrammar _ -> Vernacentries.interp v + | VernacRequire (_,_,l) -> + (try + Vernacentries.interp v + with _ -> + let l=prlist_with_sep spc (fun (_,qid) -> Nametab.pr_qualid qid) l in + msgnl (str "Reinterning of " ++ l ++ str " failed")) + | VernacRequireFrom (_,_,name,_) -> + (try + Vernacentries.interp v + with _ -> + msgnl (str "Reinterning of " ++ Nameops.pr_id name ++ str " failed")) + | _ -> ()); v;; let parse_to_dot = let rec dot st = match Stream.next st with @@ -106,13 +132,13 @@ let rec discard_to_dot stream = let rec decompose_string_aux s n = try let index = String.index_from s n '\n' in - (Ctast.str (String.sub s n (index - n))):: + (String.sub s n (index - n)):: (decompose_string_aux s (index + 1)) - with Not_found -> [Ctast.str(String.sub s n ((String.length s) - n))];; + with Not_found -> [String.sub s n ((String.length s) - n)];; let decompose_string s n = match decompose_string_aux s n with - (Ctast.Str(_,""))::tl -> tl + ""::tl -> tl | a -> a;; let make_string_list file_chan fst_pos snd_pos = @@ -161,13 +187,21 @@ let rec get_substring_list string_list fst_pos snd_pos = (* When parsing a list of commands, we try to recover error messages for each individual command. *) +type parse_result = + | ParseOK of Vernacexpr.vernac_expr located option + | ParseError of string * string list + +let embed_string s = + CT_coerce_STRING_OPT_to_VARG (CT_coerce_STRING_to_STRING_OPT (CT_string s)) + +let make_parse_error_item s l = + CT_user_vernac (CT_ident s, CT_varg_list (List.map embed_string l)) + let parse_command_list reqid stream string_list = let rec parse_whole_stream () = let this_pos = Stream.count stream in let first_ast = - try option_app - Ctast.ast_to_ct (Gram.Entry.parse - Pcoq.main_entry (Gram.parsable stream)) + try ParseOK (Gram.Entry.parse Pcoq.main_entry (Gram.parsable stream)) with | (Stdpp.Exc_located(l, Stream.Error txt)) as e -> begin @@ -176,31 +210,48 @@ let parse_command_list reqid stream string_list = discard_to_dot stream; msgnl (str "debug" ++ fnl () ++ int this_pos ++ fnl () ++ int (Stream.count stream)); +(* Some( Node(l, "PARSING_ERROR", List.map Ctast.str (get_substring_list string_list this_pos (Stream.count stream)))) - with End_of_file -> None +*) + ParseError ("PARSE_ERROR", + get_substring_list string_list this_pos + (Stream.count stream)) + with End_of_file -> ParseOK None end | e-> begin discard_to_dot stream; +(* Some(Node((0,0), "PARSING_ERROR2", List.map Ctast.str (get_substring_list string_list this_pos (Stream.count stream)))) +*) + ParseError ("PARSE_ERROR2", + get_substring_list string_list this_pos (Stream.count stream)) end in match first_ast with - | Some ast -> + | ParseOK (Some (loc,ast)) -> let ast0 = (execute_when_necessary ast) in (try xlate_vernac ast with e -> +(* xlate_vernac (Node((0,0), "PARSING_ERROR2", List.map Ctast.str (get_substring_list string_list this_pos (Stream.count stream)))))::parse_whole_stream() - | None -> [] in +*) + make_parse_error_item "PARSING_ERROR2" + (get_substring_list string_list this_pos + (Stream.count stream)))::parse_whole_stream() + | ParseOK None -> [] + | ParseError (s,l) -> + (make_parse_error_item s l)::parse_whole_stream() + in match parse_whole_stream () with | first_one::tail -> (P_cl (CT_command_list(first_one, tail))) | [] -> raise (UserError ("parse_string", (str "empty text.")));; @@ -219,25 +270,25 @@ let parse_string_action reqid phylum char_stream string_list = P_c (xlate_vernac (execute_when_necessary - (Ctast.ast_to_ct (Gram.Entry.parse Pcoq.Vernac_.vernac_eoi (Gram.parsable char_stream))))) + (Gram.Entry.parse Pcoq.Vernac_.vernac_eoi (Gram.parsable char_stream)))) | "TACTIC_COM" -> P_t - (xlate_tactic (Ctast.ast_to_ct(Gram.Entry.parse Pcoq.Tactic.tactic_eoi - (Gram.parsable char_stream)))) + (xlate_tactic (Gram.Entry.parse Pcoq.Tactic.tactic_eoi + (Gram.parsable char_stream))) | "FORMULA" -> P_f (xlate_formula (Ctast.ast_to_ct (Gram.Entry.parse Pcoq.Constr.constr_eoi (Gram.parsable char_stream)))) - | "ID" -> P_id (xlate_id - (Ctast.ast_to_ct (Gram.Entry.parse Pcoq.Prim.ident - (Gram.parsable char_stream)))) + | "ID" -> P_id (xlate_ident + (Gram.Entry.parse Pcoq.Prim.ident + (Gram.parsable char_stream))) | "STRING" -> P_s - (xlate_string (Ctast.ast_to_ct (Gram.Entry.parse Pcoq.Prim.string - (Gram.parsable char_stream)))) + (CT_string (Gram.Entry.parse Pcoq.Prim.string + (Gram.parsable char_stream))) | "INT" -> - P_i (xlate_int (Ctast.ast_to_ct (Gram.Entry.parse Pcoq.Prim.number - (Gram.parsable char_stream)))) + P_i (CT_int (Gram.Entry.parse Pcoq.Prim.natural + (Gram.parsable char_stream))) | _ -> error "parse_string_action : bad phylum" in print_parse_results reqid msg with @@ -273,10 +324,7 @@ let parse_file_action reqid file_name = let this_pos = Stream.count stream in match try - let this_ast = - option_app Ctast.ast_to_ct - (Gram.Entry.parse Pcoq.main_entry (Gram.parsable stream)) in - this_ast + ParseOK(Gram.Entry.parse Pcoq.main_entry (Gram.parsable stream)) with | Stdpp.Exc_located(l,Stream.Error txt) -> msgnl (ctf_SyntaxWarningMessage reqid @@ -287,21 +335,21 @@ let parse_file_action reqid file_name = (try begin discard_to_dot (); - Some( Node(l, "PARSING_ERROR", + ParseError ("PARSING_ERROR", (make_string_list file_chan_err this_pos - (Stream.count stream)))) + (Stream.count stream))) end - with End_of_file -> None) + with End_of_file -> ParseOK None) | e -> begin Gram.Entry.parse parse_to_dot (Gram.parsable stream); - Some(Node((0,0), "PARSING_ERROR2", - (make_string_list file_chan this_pos - (Stream.count stream)))) + ParseError ("PARSING_ERROR2", + (make_string_list file_chan this_pos + (Stream.count stream))) end with - | Some ast -> + | ParseOK (Some (_,ast)) -> let ast0=(execute_when_necessary ast) in let term = (try xlate_vernac ast @@ -311,12 +359,13 @@ let parse_file_action reqid file_name = " " ^ (string_of_int (Stream.count stream)) ^ "\n"); - xlate_vernac - (Node((0,0), "PARSING_ERROR2", - (make_string_list file_chan_err this_pos - (Stream.count stream))))) in + make_parse_error_item "PARSING_ERROR2" + (make_string_list file_chan_err this_pos + (Stream.count stream))) in term::parse_whole_file () - | None -> [] in + | ParseOK None -> [] + | ParseError (s,l) -> + (make_parse_error_item s l)::parse_whole_file () in parse_whole_file () with | first_one :: tail -> print_parse_results reqid @@ -379,9 +428,9 @@ let load_syntax_action reqid module_name = msg (str "loading " ++ str module_name ++ str "... "); try (let qid = Nametab.make_short_qualid (Names.id_of_string module_name) in - read_module qid; + read_module (dummy_loc,qid); msg (str "opening... "); - import_module false qid; + import_module false (dummy_loc,qid); msgnl (str "done" ++ fnl ()); ()) with diff --git a/contrib/interface/pbp.ml b/contrib/interface/pbp.ml index 2e92dd7c9..5dd9d071b 100644 --- a/contrib/interface/pbp.ml +++ b/contrib/interface/pbp.ml @@ -14,6 +14,7 @@ open Environ;; open Proof_trees;; open Proof_type;; open Tacmach;; +open Tacexpr;; open Typing;; open Pp;; @@ -24,37 +25,79 @@ let get_hyp_by_name g name = let env = pf_env g in try (let judgment = Pretyping.understand_judgment - evd env (RVar(dummy_loc, id_of_string name)) in + evd env (RVar(dummy_loc, name)) in ("hyp",judgment.uj_type)) (* je sais, c'est pas beau, mais je ne sais pas trop me servir de look_up... Loc *) - with _ -> (let parse_ast = Pcoq.parse_string Pcoq.Constr.constr in - let parse s = Astterm.interp_constr Evd.empty (Global.env()) - (parse_ast s) in - ("cste",type_of (Global.env()) Evd.empty (parse name))) + with _ -> (let ast = Termast.ast_of_qualid (Nametab.make_short_qualid name)in + let c = Astterm.interp_constr evd env ast in + ("cste",type_of (Global.env()) Evd.empty c)) ;; +type pbp_atom = + | PbpTryAssumption of identifier option + | PbpTryClear of identifier list + | PbpGeneralize of identifier * identifier list + | PbpLApply of identifier (* = CutAndApply *) + | PbpIntros of identifier list + | PbpSplit + (* Existential *) + | PbpExists of identifier + (* Or *) + | PbpLeft + | PbpRight + (* Not *) + | PbpReduce + (* Head *) + | PbpApply of identifier + | PbpElim of identifier * identifier list;; + +(* Invariant: In PbpThens ([a1;...;an],[t1;...;tp]), all tactics + [a1]..[an-1] are atomic (or try of an atomic) tactic and produce + exactly one goal, and [an] produces exactly p subgoals + + In [PbpThen [a1;..an]], all tactics are (try of) atomic tactics and + produces exactly one subgoal, except the last one which may complete the + goal + + Convention: [PbpThen []] is Idtac and [PbpThen t] is a coercion + from atomic to composed tactic +*) + +type pbp_sequence = + | PbpThens of pbp_atom list * pbp_sequence list + | PbpThen of pbp_atom list + +(* This flattens sequences of tactics producing just one subgoal *) +let chain_tactics tl1 = function + | PbpThens (tl2, tl3) -> PbpThens (tl1@tl2, tl3) + | PbpThen tl2 -> PbpThen (tl1@tl2) + type pbp_rule = (identifier list * - string list * + identifier list * bool * - string option * + identifier option * (types, constr) kind_of_term * int list * (identifier list -> - string list -> + identifier list -> bool -> - string option -> (types, constr) kind_of_term -> int list -> Ctast.t)) -> - Ctast.t option;; + identifier option -> (types, constr) kind_of_term -> int list -> pbp_sequence)) -> + pbp_sequence option;; let zz = (0,0);; +(* let make_named_intro s = Node(zz, "Intros", [Node(zz,"INTROPATTERN", [Node(zz,"LISTPATTERN", [Node(zz, "IDENTIFIER", [Nvar(zz, s)])])])]);; +*) +let make_named_intro id = PbpIntros [id] +(* let get_name_from_intro = function Node(a, "Intros", [Node(b,"INTROPATTERN", @@ -63,21 +106,24 @@ let get_name_from_intro = function [Nvar(e, s)])])])]) -> Some s | _ -> None;; - +*) +(* let make_clears = function [] -> Node(zz, "Idtac",[]) | str_list -> Node(zz, "TRY", [Node(zz,"Clear", [Node(zz, "CLAUSE", List.map (function s -> Nvar(zz,s)) str_list)]) ]);; +*) +let make_clears str_list = PbpThen [PbpTryClear str_list] let add_clear_names_if_necessary tactic clear_names = match clear_names with [] -> tactic - | l -> Node(zz, "TACTICLIST", [make_clears l; tactic]);; + | l -> chain_tactics [PbpTryClear l] tactic;; let make_final_cmd f optname clear_names constr path = - let tactic = f optname constr path in +(* let tactic = f optname constr path in*) add_clear_names_if_necessary (f optname constr path) clear_names;; let (rem_cast:pbp_rule) = function @@ -94,18 +140,17 @@ let (forall_intro: pbp_rule) = function (2::path), f) -> let x' = next_global_ident_away x avoid in - Some(Node(zz, "TACTICLIST", - [make_named_intro (string_of_id x'); f (x'::avoid) - clear_names clear_flag None (kind_of_term body) path])) + Some(chain_tactics [make_named_intro x'] + (f (x'::avoid) + clear_names clear_flag None (kind_of_term body) path)) | _ -> None;; let (imply_intro2: pbp_rule) = function avoid, clear_names, clear_flag, None, Prod(Anonymous, _, body), 2::path, f -> let h' = next_global_ident_away (id_of_string "H") avoid in - Some(Node(zz, "TACTICLIST", - [make_named_intro (string_of_id h'); - f (h'::avoid) clear_names clear_flag None (kind_of_term body) path])) + Some(chain_tactics [make_named_intro h'] + (f (h'::avoid) clear_names clear_flag None (kind_of_term body) path)) | _ -> None;; @@ -113,72 +158,90 @@ let (imply_intro1: pbp_rule) = function avoid, clear_names, clear_flag, None, Prod(Anonymous, prem, body), 1::path, f -> let h' = next_global_ident_away (id_of_string "H") avoid in - let str_h' = (string_of_id h') in - Some(Node(zz, "TACTICLIST", - [make_named_intro str_h'; - f (h'::avoid) clear_names clear_flag (Some str_h') - (kind_of_term prem) path])) + let str_h' = h' in + Some(chain_tactics [make_named_intro str_h'] + (f (h'::avoid) clear_names clear_flag (Some str_h') + (kind_of_term prem) path)) | _ -> None;; +let make_pbp_pattern x = + Coqast.Node(zz,"APPLIST", + [Coqast.Nvar (zz, id_of_string "PBP_META"); + Coqast.Nvar (zz, id_of_string ("Value_for_" ^ (string_of_id x)))]) + +let rec make_then = function + | [] -> TacId + | [t] -> t + | t1::t2::l -> make_then (TacThen (t1,t2)::l) + +let make_pbp_atomic_tactic = function + | PbpTryAssumption None -> TacTry (TacAtom (zz, TacAssumption)) + | PbpTryAssumption (Some a) -> + TacTry (TacAtom (zz, TacExact (Coqast.Nvar (zz,a)))) + | PbpExists x -> + TacAtom (zz, TacSplit (ImplicitBindings [make_pbp_pattern x])) + | PbpGeneralize (h,args) -> + let l = Coqast.Nvar (zz, h)::List.map make_pbp_pattern args in + TacAtom (zz, TacGeneralize [Coqast.Node (zz, "APPLIST", l)]) + | PbpLeft -> TacAtom (zz, TacLeft NoBindings) + | PbpRight -> TacAtom (zz, TacRight NoBindings) + | PbpReduce -> TacAtom (zz, TacReduce (Red false, [])) + | PbpIntros l -> + let l = List.map (fun id -> IntroIdentifier id) l in + TacAtom (zz, TacIntroPattern l) + | PbpLApply h -> TacAtom (zz, TacLApply (Coqast.Nvar (zz, h))) + | PbpApply h -> TacAtom (zz, TacApply (Coqast.Nvar(zz, h),NoBindings)) + | PbpElim (hyp_name, names) -> + let bind = List.map (fun s -> (NamedHyp s,make_pbp_pattern s)) names in + TacAtom + (zz, TacElim ((Coqast.Nvar(zz,hyp_name),ExplicitBindings bind),None)) + | PbpTryClear l -> + TacTry (TacAtom (zz, TacClear (List.map (fun s -> AN (zz,s)) l))) + | PbpSplit -> TacAtom (zz, TacSplit NoBindings);; + +let rec make_pbp_tactic = function + | PbpThen tl -> make_then (List.map make_pbp_atomic_tactic tl) + | PbpThens (l,tl) -> + TacThens + (make_then (List.map make_pbp_atomic_tactic l), + List.map make_pbp_tactic tl) let (forall_elim: pbp_rule) = function avoid, clear_names, clear_flag, Some h, Prod(Name x, _, body), 2::path, f -> let h' = next_global_ident_away (id_of_string "H") avoid in let clear_names' = if clear_flag then h::clear_names else clear_names in - let str_h' = (string_of_id h') in - Some(Node(zz, "TACTICLIST", - [Node(zz,"Generalize",[Node - (zz, "COMMAND", - [Node - (zz, "APPLIST", - [Nvar (zz, h); - Node(zz,"APPLIST", [Nvar (zz, "PBP_META"); - Nvar(zz, "Value_for_" ^ (string_of_id x))])])])]); - make_named_intro str_h'; - f (h'::avoid) clear_names' true (Some str_h') (kind_of_term body) path])) + Some + (chain_tactics [PbpGeneralize (h,[x]); make_named_intro h'] + (f (h'::avoid) clear_names' true (Some h') (kind_of_term body) path)) | _ -> None;; - + + let (imply_elim1: pbp_rule) = function avoid, clear_names, clear_flag, Some h, Prod(Anonymous, prem, body), 1::path, f -> let clear_names' = if clear_flag then h::clear_names else clear_names in let h' = next_global_ident_away (id_of_string "H") avoid in let str_h' = (string_of_id h') in - Some(Node - (zz, "TACTICLIST", - [Node - (zz, "CutAndApply", - [Node (zz, "COMMAND", [Nvar (zz, h)])]); - Node(zz, "TACLIST", - [Node - (zz, "TACTICLIST", - [Node (zz, "Intro", [Nvar (zz, str_h')]); - make_clears (h::clear_names)]); - Node (zz, "TACTICLIST", - [f avoid clear_names' false None - (kind_of_term prem) path])])])) + Some(PbpThens + ([PbpLApply h], + [PbpThen [PbpIntros [h']]; + make_clears (h::clear_names); + f avoid clear_names' false None (kind_of_term prem) path])) | _ -> None;; + let (imply_elim2: pbp_rule) = function avoid, clear_names, clear_flag, Some h, Prod(Anonymous, prem, body), 2::path, f -> let clear_names' = if clear_flag then h::clear_names else clear_names in let h' = next_global_ident_away (id_of_string "H") avoid in - let str_h' = (string_of_id h') in - Some(Node - (zz, "TACTICLIST", - [Node - (zz, "CutAndApply", - [Node (zz, "COMMAND", [Nvar (zz, h)])]); - Node(zz, "TACLIST", - [Node - (zz, "TACTICLIST", - [Node (zz, "Intro", [Nvar (zz, str_h')]); - f (h'::avoid) clear_names' false (Some str_h') - (kind_of_term body) path]); - Node (zz, "TACTICLIST", - [make_clears clear_names])])])) + Some(PbpThens + ([PbpLApply h], + [chain_tactics [PbpIntros [h']] + (f (h'::avoid) clear_names' false (Some h') + (kind_of_term body) path); + make_clears clear_names])) | _ -> None;; let reference dir s = @@ -217,18 +280,11 @@ let (and_intro: pbp_rule) = function let cont_cmd = f avoid clear_names false None (kind_of_term cont_term) path in let clear_cmd = make_clears clear_names in - let cmds = List.map - (function tac -> - Node (zz, "TACTICLIST", [tac])) + let cmds = (if a = 1 then [cont_cmd;clear_cmd] else [clear_cmd;cont_cmd]) in - Some - (Node - (zz, "TACTICLIST", - [Node (zz, "Split", [Node (zz, "BINDINGS", [])]); - Node - (zz, "TACLIST", cmds)])) + Some (PbpThens ([PbpSplit],cmds)) else None | _ -> None;; @@ -239,17 +295,7 @@ let (ex_intro: pbp_rule) = function or (is_matching_local (sigconstr ()) oper) or (is_matching_local (sigTconstr ()) oper) -> (match kind_of_term c2 with - Lambda(Name x, _, body) -> - Some(Node(zz, "Split", - [Node(zz, "BINDINGS", - [Node(zz, "BINDING", - [Node(zz, "COMMAND", - [Node(zz, "APPLIST", - [Nvar(zz, "PBP_META"); - Nvar(zz, - ("Value_for_" ^ - (string_of_id x)))]) - ])])])])) + Lambda(Name x, _, body) -> Some (PbpThen [PbpExists x]) | _ -> None) | _ -> None;; @@ -261,12 +307,10 @@ let (or_intro: pbp_rule) = function (is_matching_local (sumconstr ()) or_oper)) & (a = 1 or a = 2) then let cont_term = if a = 1 then c1 else c2 in - let fst_cmd = Node(zz, (if a = 1 then "Left" else "Right"), - [Node(zz, "BINDINGS",[])]) in + let fst_cmd = if a = 1 then PbpLeft else PbpRight in let cont_cmd = f avoid clear_names false None (kind_of_term cont_term) path in - Some(Node(zz, "TACTICLIST", - [fst_cmd;cont_cmd])) + Some(chain_tactics [fst_cmd] cont_cmd) else None | _ -> None;; @@ -279,19 +323,16 @@ let (not_intro: pbp_rule) = function if(is_matching_local (notconstr ()) not_oper) or (is_matching_local (notTconstr ()) not_oper) then let h' = next_global_ident_away (id_of_string "H") avoid in - let str_h' = (string_of_id h') in - Some(Node(zz,"TACTICLIST", - [Node(zz,"Reduce",[Node(zz,"REDEXP",[Node(zz,"Red",[])]); - Node(zz,"CLAUSE",[])]); - make_named_intro str_h'; - f (h'::avoid) clear_names false (Some str_h') - (kind_of_term c1) path])) + Some(chain_tactics [PbpReduce;make_named_intro h'] + (f (h'::avoid) clear_names false (Some h') + (kind_of_term c1) path)) else None | _ -> None;; +(* let elim_with_bindings hyp_name names = Node(zz,"Elim", [Node(zz, "COMMAND", [Nvar(zz,hyp_name)]); @@ -307,6 +348,9 @@ let elim_with_bindings hyp_name names = [Nvar(zz,"PBP_META"); Nvar(zz, "value_for_" ^ s)])])])) names)]);; +*) +let elim_with_bindings hyp_name names = + PbpElim (hyp_name, names);; let auxiliary_goals clear_names clear_flag this_name n_aux = let clear_cmd = @@ -338,13 +382,13 @@ let auxiliary_goals clear_names clear_flag this_name n_aux = let rec down_prods: (types, constr) kind_of_term * (int list) * int -> - string list * (int list) * int * (types, constr) kind_of_term * + identifier list * (int list) * int * (types, constr) kind_of_term * (int list) = function Prod(Name x, _, body), 2::path, k -> let res_sl, res_il, res_i, res_cstr, res_p = down_prods (kind_of_term body, path, k+1) in - (string_of_id x)::res_sl, (k::res_il), res_i, res_cstr, res_p + x::res_sl, (k::res_il), res_i, res_cstr, res_p | Prod(Anonymous, _, body), 2::path, k -> let res_sl, res_il, res_i, res_cstr, res_p = down_prods (kind_of_term body, path, k+1) in @@ -400,6 +444,7 @@ let (mk_db_indices: int list -> int -> int list) = answer is true, then the built command takes advantage of the power of head tactics. *) +(* let (head_tactic_patt: pbp_rule) = function avoid, clear_names, clear_flag, Some h, cstr, path, f -> (match down_prods (cstr, path, 0) with @@ -505,7 +550,98 @@ let (head_tactic_patt: pbp_rule) = function clear_names) | _ -> None) | _ -> None;; - +*) + +let (head_tactic_patt: pbp_rule) = function + avoid, clear_names, clear_flag, Some h, cstr, path, f -> + (match down_prods (cstr, path, 0) with + | (str_list, _, nprems, + App(oper,[|c1|]), 2::1::path) + when + (is_matching_local (notconstr ()) oper) or + (is_matching_local (notTconstr ()) oper) -> + Some(chain_tactics [elim_with_bindings h str_list] + (f avoid clear_names false None (kind_of_term c1) path)) + | (str_list, _, nprems, + App(oper, [|c1; c2|]), 2::a::path) + when ((is_matching_local (andconstr()) oper) or + (is_matching_local (prodconstr()) oper)) & (a = 1 or a = 2) -> + let h1 = next_global_ident_away (id_of_string "H") avoid in + let h2 = next_global_ident_away (id_of_string "H") (h1::avoid) in + Some(PbpThens + ([elim_with_bindings h str_list; + make_named_intro h1; + make_named_intro h2], + let cont_body = + if a = 1 then c1 else c2 in + let cont_tac = + f (h2::h1::avoid) (h::clear_names) + false (Some (if 1 = a then h1 else h2)) + (kind_of_term cont_body) path in + cont_tac::(auxiliary_goals + clear_names clear_flag + h nprems))) + | (str_list, _, nprems, App(oper,[|c1; c2|]), 2::a::path) + when ((is_matching_local (exconstr ()) oper) or + (is_matching_local (exTconstr ()) oper) or + (is_matching_local (sigconstr ()) oper) or + (is_matching_local (sigTconstr()) oper)) & a = 2 -> + (match (kind_of_term c2),path with + Lambda(Name x, _,body), (2::path) -> + Some(PbpThens + ([elim_with_bindings h str_list], + let x' = next_global_ident_away x avoid in + let cont_body = + Prod(Name x', c1, + mkProd(Anonymous, body, + mkVar(dummy_id))) in + let cont_tac + = f avoid (h::clear_names) false None + cont_body (2::1::path) in + cont_tac::(auxiliary_goals + clear_names clear_flag + h nprems))) + | _ -> None) + | (str_list, _, nprems, App(oper,[|c1; c2|]), 2::a::path) + when ((is_matching_local (orconstr ()) oper) or + (is_matching_local (sumboolconstr ()) oper) or + (is_matching_local (sumconstr ()) oper)) & + (a = 1 or a = 2) -> + Some(PbpThens + ([elim_with_bindings h str_list], + let cont_body = + if a = 1 then c1 else c2 in + (* h' is the name for the new intro *) + let h' = next_global_ident_away (id_of_string "H") avoid in + let cont_tac = + chain_tactics + [make_named_intro h'] + (f + (* h' should not be used again *) + (h'::avoid) + (* the disjunct itself can be discarded *) + (h::clear_names) false (Some h') + (kind_of_term cont_body) path) in + let snd_tac = + chain_tactics + [make_named_intro h'] + (make_clears (h::clear_names)) in + let tacs1 = + if a = 1 then + [cont_tac; snd_tac] + else + [snd_tac; cont_tac] in + tacs1@(auxiliary_goals (h::clear_names) + false dummy_id nprems))) + | (str_list, int_list, nprems, c, []) + when (check_apply c (mk_db_indices int_list nprems)) & + (match c with Prod(_,_,_) -> false + | _ -> true) & + (List.length int_list) + nprems > 0 -> + Some(add_clear_names_if_necessary (PbpThen [PbpApply h]) clear_names) + | _ -> None) + | _ -> None;; + let pbp_rules = ref [rem_cast;head_tactic_patt;forall_intro;imply_intro2; forall_elim; imply_intro1; imply_elim1; imply_elim2; @@ -541,11 +677,15 @@ let rec clean_trace flag = command requested by the point-and-shoot strategy. Default is Try Assumption--Try Exact. *) +(* let default_ast optname constr path = match optname with None -> Node(zz, "TRY", [Node(zz, "Assumption",[])]) | Some(a) -> Node(zz, "TRY", [Node(zz, "Exact",[Node(zz,"COMMAND",[Nvar(zz,a)])])]);; +*) + +let default_ast optname constr path = PbpThen [PbpTryAssumption optname] (* This is the main proof by pointing function. *) (* avoid: les noms a ne pas utiliser *) @@ -564,8 +704,9 @@ let rec pbpt final_cmd avoid clear_names clear_flag opt_name constr path = in try_all_rules (!pbp_rules);; (* these are the optimisation functions. *) -(* This function takes care of flattening successive intro commands. *) +(* This function takes care of flattening successive then commands. *) +(* let rec optim1 = function Node(a,"TACTICLIST",b) -> @@ -591,10 +732,25 @@ let rec optim1 = | [t] -> t | args -> Node(zz,"TACTICLIST", args)) | t -> t;; - +*) + +(* Invariant: in [flatten_sequence t], occurrences of [PbpThenCont(l,t)] enjoy + that t is some [PbpAtom t] *) +(* +let rec flatten_sequence = + function + PbpThens (tl1,tl2) -> PbpThens (tl1,List.map flatten_sequence tl2) + | PbpThen (tl1,t1) as x -> + (match flatten_sequence t1 with + | PbpThenCont (tl2,t2) -> PbpThenCont (tl1@tl2,t2) + | PbpThens (tl2,tl3) -> PbpThens (tl1@tl2,tl3) + | PbpAtom _ -> x) + | PbpAtom t as x -> x;; +*) (* This optimization function takes care of compacting successive Intro commands together. *) +(* let rec optim2 = function Node(a,"TACTICLIST",b) -> @@ -624,7 +780,42 @@ let rec optim2 = | l -> (put_ids l)::t1::(group_intros [] others)) in Node(a,"TACTICLIST",group_intros [] b) | t -> t;; +*) +(* +let rec optim2 = function + | TacThens (t,tl) -> TacThens (optim2 t, List.map optim2 tl) + | t -> + let get_ids = + List.map (function IntroIdentifier _ as x -> x + | _ -> failwith "get_ids expected an identifier") in + let rec group_intros names = function + [] -> (match names with + [] -> [] + | l -> [TacIntroPattern l]) + | (TacIntroPattern ids)::others -> + group_intros (names@(get_ids ids)) others + | t1::others -> + (match names with + [] -> t1::(group_intros [] others) + | l -> (TacIntroPattern l)::t1::(group_intros [] others)) in + make_then (group_intros [] (flatten_sequence t)) +*) +let rec group_intros names = function + [] -> (match names with + [] -> [] + | l -> [PbpIntros l]) + | (PbpIntros ids)::others -> group_intros (names@ids) others + | t1::others -> + (match names with + [] -> t1::(group_intros [] others) + | l -> (PbpIntros l)::t1::(group_intros [] others)) + +let rec optim2 = function + | PbpThens (tl1,tl2) -> PbpThens (group_intros [] tl1, List.map optim2 tl2) + | PbpThen tl -> PbpThen (group_intros [] tl) + +(* let rec merge_ast_in_command com1 com2 = let args1 = (match com1 with @@ -635,7 +826,8 @@ let rec merge_ast_in_command com1 com2 = Node(_, "APPLIST", hyp::args) -> args | _ -> failwith "unexpected com2 in merge_ast_in_command") in Node(zz, "COMMAND", [Node(zz, "APPLIST", args1@args2)]);; - +*) +(* let cleanup_clears empty_allowed names str_list other = let rec clean_aux = function [] -> [] @@ -653,42 +845,44 @@ let cleanup_clears empty_allowed names str_list other = else [Node(zz,"Idtac",[])] | _ -> other) | l -> Node(zz, "TRY", [Node(zz, "Clear", [Node(zz,"CLAUSE", l)])])::other;; - +*) +let rec cleanup_clears str_list = function + [] -> [] + | x::tail -> + if List.mem x str_list then cleanup_clears str_list tail + else x::(cleanup_clears str_list tail);; (* This function takes care of compacting instanciations of universal quantifications. *) + +let rec optim3_aux str_list = function + (PbpGeneralize (h,l1))::(PbpIntros [s])::(PbpGeneralize (h',l2))::others + when s=h' -> + optim3_aux (s::str_list) (PbpGeneralize (h,l1@l2)::others) + | (PbpTryClear names)::other -> + (match cleanup_clears str_list names with + [] -> other + | l -> (PbpTryClear l)::other) + | a::l -> a::(optim3_aux str_list l) + | [] -> [];; + let rec optim3 str_list = function - Node(a,"TACTICLIST", b) -> - let rec optim3_aux empty_allowed str_list = function - ((Node(a, "Generalize", [Node(_, "COMMAND", [com1])])) as t1):: - intro:: - (Node(b, "Generalize", [Node(_, "COMMAND", [com2])]) as t2)::others -> - (match get_name_from_intro intro with - None -> t1::intro::(optim3_aux true str_list (t2::others)) - | Some s -> optim3_aux true (s::str_list) - (Node(a, "Generalize", - [merge_ast_in_command com1 com2])::others)) - |( Node(zz, "TRY", [Node(a,"Clear", [Node(_,"CLAUSE", names)])]))::other -> - cleanup_clears empty_allowed names str_list other - | [Node(a,"TACLIST",branches)] -> - [Node(a,"TACLIST",List.map (optim3 str_list) branches)] - | a::l -> a::(optim3_aux true str_list l) - | [] -> if empty_allowed then - [] - else failwith "strange value in optim3_aux" in - Node(a, "TACTICLIST", optim3_aux false str_list b) - | t -> t;; + PbpThens (tl1, tl2) -> + PbpThens (optim3_aux str_list tl1, List.map (optim3 str_list) tl2) + | PbpThen tl -> PbpThen (optim3_aux str_list tl) -let optim x = optim3 [] (optim2 (optim1 x));; +let optim x = make_pbp_tactic (optim3 [] (optim2 x));; +(* TODO add_tactic "Traced_Try" traced_try_entry;; +*) let rec tactic_args_to_ints = function [] -> [] | (Integer n)::l -> n::(tactic_args_to_ints l) | _ -> failwith "expecting only numbers";; - +(* let pbp_tac display_function = function (Identifier a)::l -> (function g -> @@ -720,3 +914,34 @@ let pbp_tac display_function = function | _ -> failwith "expecting other arguments";; +*) +let pbp_tac display_function idopt nl = + match idopt with + | Some str -> + (function g -> + let (ou,tstr) = (get_hyp_by_name g str) in + let exp_ast = + pbpt default_ast + (match ou with + "hyp" ->(pf_ids_of_hyps g) + |_ -> (str::(pf_ids_of_hyps g))) + [] + false + (Some str) + (kind_of_term tstr) + nl in + (display_function (optim exp_ast); tclIDTAC g)) + | None -> + if nl <> [] then + (function g -> + let exp_ast = + (pbpt default_ast (pf_ids_of_hyps g) [] false + None (kind_of_term (pf_concl g)) nl) in + (display_function (optim exp_ast); tclIDTAC g)) + else + (function g -> + (display_function + (make_pbp_tactic (default_ast None (pf_concl g) [])); + tclIDTAC g));; + + diff --git a/contrib/interface/pbp.mli b/contrib/interface/pbp.mli index 57794ec22..43ec1274d 100644 --- a/contrib/interface/pbp.mli +++ b/contrib/interface/pbp.mli @@ -1,4 +1,4 @@ -val pbp_tac : (Ctast.t -> 'a) -> - Proof_type.tactic_arg list -> +val pbp_tac : (Tacexpr.raw_tactic_expr -> 'a) -> + Names.identifier option -> int list -> Proof_type.goal Tacmach.sigma -> Proof_type.goal list Proof_type.sigma * Proof_type.validation;; diff --git a/contrib/interface/showproof.ml b/contrib/interface/showproof.ml index 9a5f803d3..ae876e129 100644 --- a/contrib/interface/showproof.ml +++ b/contrib/interface/showproof.ml @@ -27,6 +27,9 @@ open Proof_trees open Sign open Pp open Printer +open Rawterm +open Tacexpr +open Genarg (*****************************************************************************) (* Arbre de preuve maison: @@ -40,7 +43,7 @@ type nhyp = {hyp_name : identifier; hyp_full_type: Term.constr} ;; -type ntactic = Coqast.t list +type ntactic = tactic_expr ;; type nproof = @@ -154,14 +157,15 @@ let seq_to_lnhyp sign sign' cl = let rule_is_complex r = match r with - Tactic ("Interp",_) -> true - | Tactic ("Auto", _) -> true - | Tactic ("Symmetry", _) -> true - |_ -> false + Tactic (TacArg (Tacexp t),_) -> true + | Tactic (TacAtom (_,TacAuto _), _) -> true + | Tactic (TacAtom (_,TacSymmetry), _) -> true + |_ -> false ;; let ast_of_constr = Termast.ast_of_constr true (Global.env()) ;; +(* let rule_to_ntactic r = let rast = (match r with @@ -178,6 +182,20 @@ let rule_to_ntactic r = else [rast ] ;; +*) +let rule_to_ntactic r = + let rt = + (match r with + Tactic (t,_) -> t + | Prim (Refine h) -> TacAtom (Ast.dummy_loc,TacExact h) + | _ -> TacAtom (Ast.dummy_loc, TacIntroPattern [])) in + if rule_is_complex r + then (match rt with + TacArg (Tacexp _) as t -> t + | _ -> assert false) + + else rt +;; let term_of_command x = @@ -238,14 +256,12 @@ let to_nproof sigma osign pf = | Some(r,spfl) -> if rule_is_complex r then ( - let p1=(match pf.subproof with - Some x -> to_nproof_rec sigma sign x - |_-> assert false) in + let p1= to_nproof_rec sigma sign (subproof_of_proof pf) in let ntree= fill_unproved p1 (List.map (fun x -> (to_nproof_rec sigma sign x).t_proof) spfl) in (match r with - Tactic ("Auto",_) -> + Tactic (TacAtom (_, TacAuto _),_) -> if spfl=[] then {t_info="to_prove"; @@ -253,7 +269,7 @@ let to_nproof sigma osign pf = t_concl=concl ntree; t_full_concl=ntree.t_goal.t_full_concl; t_full_env=ntree.t_goal.t_full_env}; - t_proof= Proof ([Node ((0,0), "InfoAuto",[])], [ntree])} + t_proof= Proof (TacAtom (Ast.dummy_loc,TacExtend ("InfoAuto",[])), [ntree])} else ntree | _ -> ntree)) else @@ -399,11 +415,15 @@ let enumerate f ln = let constr_of_ast = Astterm.interp_constr Evd.empty (Global.env());; + +(* let sp_tac tac = try spt (constr_of_ast (term_of_command tac)) with _ -> (* let Node(_,t,_) = tac in *) spe (* sps ("error in sp_tac " ^ t) *) ;; +*) +let sp_tac tac = failwith "TODO" let soit_A_une_proposition nh ln t= match !natural_language with French -> @@ -942,6 +962,7 @@ let natural_lhyp lh hi = Analyse des tactiques. *) +(* let name_tactic tac = match tac with (Node(_,"Interp", @@ -950,7 +971,14 @@ let name_tactic tac = |(Node(_,t,_))::_ -> t | _ -> assert false ;; +*) +let name_tactic = function + | TacIntroPattern _ -> "Intro" + | TacAssumption -> "Assumption" + | _ -> failwith "TODO" +;; +(* let arg1_tactic tac = match tac with (Node(_,"Interp", @@ -960,6 +988,9 @@ let arg1_tactic tac = | x::_ -> x | _ -> assert false ;; +*) + +let arg1_tactic tac = failwith "TODO" let arg2_tactic tac = match tac with @@ -970,6 +1001,7 @@ let arg2_tactic tac = | _ -> assert false ;; +(* type nat_tactic = Split of (Coqast.t list) | Generalize of (Coqast.t list) @@ -992,7 +1024,7 @@ let analyse_tac tac = | [Node (_, x,la)] -> Other (x,la) | _ -> assert false ;; - +*) @@ -1103,6 +1135,17 @@ let terms_of_equality e = let eq_term = eq_constr;; +let is_equality_tac = function + | TacAtom (_, + (TacExtend + (("ERewriteLR"|"ERewriteRL"|"ERewriteLRocc"|"ERewriteRLocc" + |"ERewriteParallel"|"ERewriteNormal" + |"RewriteLR"|"RewriteRL"|"Replace"),_) + | TacReduce _ + | TacSymmetry | TacReflexivity + | TacExact _ | TacIntroPattern _ | TacIntroMove _ | TacAuto _)) -> true + | _ -> false + let equalities_ntree ig ntree = let rec equalities_ntree ig ntree = if not (is_equality (concl ntree)) @@ -1111,12 +1154,7 @@ let equalities_ntree ig ntree = match (proof ntree) with Notproved -> [(ig,ntree)] | Proof (tac,ltree) -> - if List.mem (name_tactic tac) - ["ERewriteLR";"ERewriteRL"; "ERewriteLRocc";"ERewriteRLocc"; - "ERewriteParallel";"ERewriteNormal"; "Reduce"; - "RewriteLR";"RewriteRL"; - "Replace";"Symmetry";"Reflexivity"; - "Exact";"Intros";"Intro";"Auto"] + if is_equality_tac tac then (match ltree with [] -> [(ig,ntree)] | t::_ -> let res=(equalities_ntree ig t) in @@ -1176,7 +1214,7 @@ let rec natural_ntree ig ntree = let lemma = match (proof ntree) with Proof (tac,ltree) -> (try (sph [spt (dbize (gLOB ge) - (term_of_command (arg1_tactic tac))); + (term_of_command (arg1_tactic tac)));(* TODO *) (match ltree with [] ->spe | [_] -> spe @@ -1224,55 +1262,62 @@ let rec natural_ntree ig ntree = sph [spi; sps (intro_not_proved_goal gs); spb; tag_toprove g ] ] - | Proof (tac,ltree) -> - (let tactic= name_tactic tac in - let ntext = - (match tactic with + + | Proof (TacId,ltree) -> natural_ntree ig (List.hd ltree) + | Proof (TacAtom (_,tac),ltree) -> + (let ntext = + match tac with (* Pas besoin de l'argument ventuel de la tactique *) - "Intros" -> natural_intros ig lh g gs ltree - | "Intro" -> natural_intros ig lh g gs ltree - | "Idtac" -> natural_ntree ig (List.hd ltree) - | "Fix" -> natural_fix ig lh g gs (arg2_tactic tac) ltree - | "Split" -> natural_split ig lh g gs ge tac ltree - | "Generalize" -> natural_generalize ig lh g gs ge tac ltree - | "Right" -> natural_right ig lh g gs ltree - | "Left" -> natural_left ig lh g gs ltree - | (* "Simpl" *)"Reduce" -> natural_reduce ig lh g gs ge tac ltree - | "InfoAuto" -> natural_infoauto ig lh g gs ltree - | "Auto" -> natural_auto ig lh g gs ltree - | "EAuto" -> natural_auto ig lh g gs ltree - | "Trivial" -> natural_trivial ig lh g gs ltree - | "Assumption" -> natural_trivial ig lh g gs ltree - | "Clear" -> natural_clear ig lh g gs ltree + TacIntroPattern _ -> natural_intros ig lh g gs ltree + | TacIntroMove _ -> natural_intros ig lh g gs ltree + | TacFix (_,n) -> natural_fix ig lh g gs n ltree + | TacSplit NoBindings -> natural_split ig lh g gs ge [] ltree + | TacSplit(ImplicitBindings l) -> natural_split ig lh g gs ge l ltree + | TacGeneralize l -> natural_generalize ig lh g gs ge l ltree + | TacRight _ -> natural_right ig lh g gs ltree + | TacLeft _ -> natural_left ig lh g gs ltree + | (* "Simpl" *)TacReduce (r,cl) -> + natural_reduce ig lh g gs ge r cl ltree + | TacExtend ("InfoAuto",[]) -> natural_infoauto ig lh g gs ltree + | TacAuto _ -> natural_auto ig lh g gs ltree + | TacExtend ("EAuto",_) -> natural_auto ig lh g gs ltree + | TacTrivial _ -> natural_trivial ig lh g gs ltree + | TacAssumption -> natural_trivial ig lh g gs ltree + | TacClear _ -> natural_clear ig lh g gs ltree (* Besoin de l'argument de la tactique *) - | "Induction" -> - let arg1=(term_of_command (arg1_tactic tac)) in - natural_induction ig lh g gs ge arg1 ltree false - | "InductionIntro" -> - let arg1=(term_of_command (arg1_tactic tac)) in - natural_induction ig lh g gs ge arg1 ltree true - | _ -> - let arg1=(term_of_command (arg1_tactic tac)) in - let arg1=dbize (gLOB ge) arg1 in - (match tactic with - "Apply" -> natural_apply ig lh g gs arg1 ltree - | "Exact" -> natural_exact ig lh g gs arg1 ltree - | "Cut" -> natural_cut ig lh g gs arg1 ltree - | "CutIntro" -> natural_cutintro ig lh g gs arg1 ltree - | "Case" -> natural_case ig lh g gs ge arg1 ltree false - | "CaseIntro" -> natural_case ig lh g gs ge arg1 ltree true - | "Elim" -> natural_elim ig lh g gs ge arg1 ltree false - | "ElimIntro" -> natural_elim ig lh g gs ge arg1 ltree true - | "RewriteRL" -> natural_rewrite ig lh g gs arg1 ltree - | "RewriteLR" -> natural_rewrite ig lh g gs arg1 ltree - | "ERewriteRL" -> natural_rewrite ig lh g gs arg1 ltree - | "ERewriteLR" -> natural_rewrite ig lh g gs arg1 ltree - |_ -> natural_generic ig lh g gs (sps tactic) (prl sp_tac tac) ltree) - ) + | TacOldInduction (NamedHyp id) -> + natural_induction ig lh g gs ge id ltree false + | TacExtend ("InductionIntro",[a]) -> + let id=(out_gen wit_ident a) in + natural_induction ig lh g gs ge id ltree true + | TacApply (c,_) -> natural_apply ig lh g gs c ltree + | TacExact c -> natural_exact ig lh g gs c ltree + | TacCut c -> natural_cut ig lh g gs c ltree + | TacExtend ("CutIntro",[a]) -> + let c = out_gen wit_constr a in + natural_cutintro ig lh g gs a ltree + | TacCase (c,_) -> natural_case ig lh g gs ge c ltree false + | TacExtend ("CaseIntro",[a]) -> + let c = out_gen wit_constr a in + natural_case ig lh g gs ge c ltree true + | TacElim ((c,_),_) -> natural_elim ig lh g gs ge c ltree false + | TacExtend ("ElimIntro",[a]) -> + let c = out_gen wit_constr a in + natural_elim ig lh g gs ge c ltree true + | TacExtend ("Rewrite",[_;a]) -> + let (c,_) = out_gen wit_constr_with_bindings a in + natural_rewrite ig lh g gs c ltree + | TacExtend ("ERewriteRL",[a]) -> + let c = out_gen wit_constr a in (* TODO *) + natural_rewrite ig lh g gs c ltree + | TacExtend ("ERewriteLR",[a]) -> + let c = out_gen wit_constr a in (* TODO *) + natural_rewrite ig lh g gs c ltree + |_ -> natural_generic ig lh g gs (sps (name_tactic tac)) (prl sp_tac [tac]) ltree in ntext (* spwithtac ntext tactic*) ) - + | Proof _ -> failwith "Don't know what to do with that" in if info<>"not_proved" then spshrink info ntext @@ -1580,12 +1625,9 @@ and natural_elim ig lh g gs ge arg1 ltree with_intros= (* InductionIntro n *) -and natural_induction ig lh g gs ge arg1 ltree with_intros= +and natural_induction ig lh g gs ge arg2 ltree with_intros= let env = (gLOB (g_env (List.hd ltree))) in - let arg1=dbize env arg1 in - let arg2 = match (kind_of_term arg1) with - Var(arg2) -> arg2 - | _ -> assert false in + let arg1= mkVar arg2 in let targ1 = prod_head (type_of env Evd.empty arg1) in let IndType (indf,targ) = find_rectype env Evd.empty targ1 in let ncti= Array.length(get_constructors env indf) in @@ -1637,8 +1679,7 @@ and natural_induction ig lh g gs ge arg1 ltree with_intros= (***********************************************************************) (* Points fixes *) -and natural_fix ig lh g gs arg ltree = - let narg = match arg with Num(_,narg) -> narg |_ -> assert false in +and natural_fix ig lh g gs narg ltree = let {t_info=info; t_goal={newhyp=lh1;t_concl=g1;t_full_concl=gf1; t_full_env=ge1};t_proof=p1}=(List.hd ltree) in @@ -1656,10 +1697,7 @@ and natural_fix ig lh g gs arg ltree = ltree) ] | _ -> assert false -and natural_reduce ig lh g gs ge tac ltree = - let (mode,la) = match analyse_tac tac with - Reduce (mode,la) -> (mode,la) - |_ -> assert false in +and natural_reduce ig lh g gs ge mode la ltree = match la with [] -> spv @@ -1669,7 +1707,7 @@ and natural_reduce ig lh g gs ge tac ltree = {ihsg=All_subgoals_hyp;isgintro="simpl"}) ltree) ] - | [Nvar (_,hyp)] -> + | [hyp] -> spv [ (natural_lhyp lh ig.ihsg); (show_goal2 lh ig g gs ""); @@ -1678,14 +1716,11 @@ and natural_reduce ig lh g gs ge tac ltree = ltree) ] | _ -> assert false -and natural_split ig lh g gs ge tac ltree = - let la = match analyse_tac tac with - Split la -> la - |_ -> assert false in +and natural_split ig lh g gs ge la ltree = match la with [arg] -> let env= (gLOB ge) in - let arg1= dbize env arg in + let arg1= (*dbize env*) arg in spv [ (natural_lhyp lh ig.ihsg); (show_goal2 lh ig g gs ""); @@ -1702,14 +1737,13 @@ and natural_split ig lh g gs ge tac ltree = ltree) ] | _ -> assert false -and natural_generalize ig lh g gs ge tac ltree = - match analyse_tac tac with - Generalize la -> - (match la with +and natural_generalize ig lh g gs ge la ltree = + match la with [arg] -> let env= (gLOB ge) in - let arg1= dbize env arg in - let type_arg=type_of_ast ge arg in + let arg1= (*dbize env*) arg in + let type_arg=type_of (Global.env()) Evd.empty arg in +(* let type_arg=type_of_ast ge arg in*) spv [ (natural_lhyp lh ig.ihsg); (show_goal2 lh ig g gs ""); @@ -1718,8 +1752,7 @@ and natural_generalize ig lh g gs ge tac ltree = {ihsg=All_subgoals_hyp;isgintro=""}) ltree) ] - | _ -> assert false) - | _ -> assert false + | _ -> assert false and natural_right ig lh g gs ltree = spv [ (natural_lhyp lh ig.ihsg); diff --git a/contrib/interface/xlate.ml b/contrib/interface/xlate.ml index 7f38ca0d2..2e49ec4e3 100644 --- a/contrib/interface/xlate.ml +++ b/contrib/interface/xlate.ml @@ -8,6 +8,10 @@ open Ast;; open Names;; open Ctast;; open Ascent;; +open Genarg;; +open Rawterm;; +open Tacexpr;; +open Vernacexpr;; let in_coq_ref = ref false;; @@ -77,6 +81,10 @@ let ctf_STRING_OPT_NONE = CT_coerce_NONE_to_STRING_OPT CT_none;; let ctf_STRING_OPT_SOME s = CT_coerce_STRING_to_STRING_OPT s;; +let ctf_STRING_OPT = function + | None -> ctf_STRING_OPT_NONE + | Some s -> ctf_STRING_OPT_SOME s + let ctv_ID_OPT_NONE = CT_coerce_NONE_to_ID_OPT CT_none;; let ctf_ID_OPT_SOME s = CT_coerce_ID_to_ID_OPT s;; @@ -102,7 +110,12 @@ let castc x = CT_coerce_TYPED_FORMULA_to_FORMULA x;; let varc x = CT_coerce_ID_to_FORMULA x;; +let xlate_ident id = CT_ident (string_of_id id) + +(* let ident_tac s = CT_user_tac (CT_ident s, CT_targ_list []);; +*) +let ident_tac s = CT_user_tac (xlate_ident s, CT_targ_list []);; let ident_vernac s = CT_user_vernac (CT_ident s, CT_varg_list []);; @@ -141,6 +154,7 @@ type iVARG = Varg_binder of ct_BINDER | Varg_tactic_arg of iTARG | Varg_varglist of iVARG list;; +(* let coerce_iTARG_to_TARG = function | Targ_intropatt x -> xlate_error "coerce_iTARG_to_TARG (3)" @@ -160,7 +174,9 @@ let coerce_iTARG_to_TARG = | Targ_cofixtac x -> CT_coerce_COFIXTAC_to_TARG x | Targ_tacexp x -> CT_coerce_TACTIC_COM_to_TARG x | Targ_redexp x -> xlate_error "coerce_iTarg_to_TARG(2)";; +*) +(* let rec coerce_iVARG_to_VARG = function | Varg_binder x -> CT_coerce_BINDER_to_VARG x @@ -189,6 +205,7 @@ let rec coerce_iVARG_to_VARG = CT_coerce_VARG_LIST_to_VARG (CT_varg_list (List.map coerce_iVARG_to_VARG x)) | _ -> xlate_error "coerce_iVARG_to_VARG: leftover case";; +*) let coerce_iVARG_to_FORMULA = function @@ -219,9 +236,14 @@ let xlate_id = | s -> CT_ident s) | _ -> xlate_error "xlate_id: not an identifier";; +(* let xlate_id_unit = function Node(_, "VOID", []) -> CT_unit | x -> CT_coerce_ID_to_ID_UNIT (xlate_id x);; +*) +let xlate_id_unit = function + None -> CT_unit + | Some x -> CT_coerce_ID_to_ID_UNIT (xlate_ident x);; let xlate_id_opt = function @@ -231,11 +253,20 @@ let xlate_id_opt = | s -> ctf_ID_OPT_SOME (CT_ident s)) | _ -> xlate_error "xlate_id: not an identifier";; +let xlate_ident_opt = + function + | None -> ctv_ID_OPT_NONE + | Some id -> ctf_ID_OPT_SOME (xlate_ident id) + let xlate_int = function | Num (_, n) -> CT_int n | _ -> xlate_error "xlate_int: not an int";; +let xlate_int_opt = function + | Some n -> CT_coerce_INT_to_INT_OPT (CT_int n) + | None -> CT_coerce_NONE_to_INT_OPT CT_none + let xlate_string = function | Str (_, s) -> CT_string s @@ -313,6 +344,28 @@ let qualid_to_ct_ID = | Node(_, "QUALIDMETA",[Num(_,n)]) -> Some(CT_metac (CT_int n)) | _ -> None;; +let tac_qualid_to_ct_ID qid = CT_ident (Nametab.string_of_qualid qid) + +let loc_qualid_to_ct_ID (_,qid) = CT_ident (Nametab.string_of_qualid qid) + +let qualid_or_meta_to_ct_ID = function + | AN (_,qid) -> tac_qualid_to_ct_ID qid + | MetaNum (_,n) -> CT_metac (CT_int n) + +let ident_or_meta_to_ct_ID = function + | AN (_,id) -> xlate_ident id + | MetaNum (_,n) -> CT_metac (CT_int n) + +let xlate_qualid_list l = CT_id_list (List.map loc_qualid_to_ct_ID l) + +let reference_to_ct_ID = function + | RIdent (_,id) -> CT_ident (Names.string_of_id id) + | RQualid (_,qid) -> CT_ident (Nametab.string_of_qualid qid) + +let xlate_class = function + | FunClass -> CT_ident "FUNCLASS" + | SortClass -> CT_ident "SORTCLASS" + | RefClass qid -> loc_qualid_to_ct_ID qid let special_case_qualid cont_function astnode = match qualid_to_ct_ID astnode with @@ -744,6 +797,13 @@ let xlate_special_cases cont_function arg = | [] -> None in xlate_rec xlate_formula_special_cases;; +let xlate_sort = + function + | Coqast.Node (_, "SET", []) -> CT_sortc "Set" + | Coqast.Node (_, "PROP", []) -> CT_sortc "Prop" + | Coqast.Node (_, "TYPE", []) -> CT_sortc "Type" + | _ -> xlate_error "xlate_sort";; + let xlate_formula a = !set_flags (); let rec ctrec = @@ -766,10 +826,28 @@ let xlate_formula a = | _ -> xlate_error "xlate_formula" in strip_Rform (ctrec a);; +(* let xlate_formula_opt = function | Node (_, "None", []) -> ctv_FORMULA_OPT_NONE | e -> CT_coerce_FORMULA_to_FORMULA_OPT (xlate_formula e);; +*) +let xlate_formula_opt = + function + | None -> ctv_FORMULA_OPT_NONE + | Some e -> CT_coerce_FORMULA_to_FORMULA_OPT (xlate_formula e);; + +let xlate_constr c = xlate_formula (Ctast.ast_to_ct c) + +let xlate_constr_opt c = xlate_formula_opt (option_app Ctast.ast_to_ct c) + +let xlate_hyp_location = + function + | InHyp (AI (_,id)) -> xlate_ident id + | InHyp (MetaId _) -> xlate_error "MetaId should occur only in quotations" + | InHypType id -> xlate_error "TODO: in (Type of id)" + +let xlate_clause l = CT_id_list (List.map xlate_hyp_location l) (** Tactics *) @@ -826,7 +904,7 @@ let make_ID_from_FORMULA = | _ -> xlate_error "make_ID_from_FORMULA: non-formula argument";; let make_ID_from_iTARG_FORMULA x = make_ID_from_FORMULA (strip_targ_command x);; - +(* let filter_binding_or_command_list bl = match bl with | (Targ_binding_com cmd) :: bl' -> @@ -837,6 +915,23 @@ let filter_binding_or_command_list bl = (CT_binding_list (List.map strip_targ_binding bl)) | [] -> CT_coerce_FORMULA_LIST_to_SPEC_LIST (CT_formula_list []) | _ -> xlate_error "filter_binding_or_command_list";; +*) +let xlate_quantified_hypothesis = function + | AnonHyp n -> CT_coerce_INT_to_ID_OR_INT (CT_int n) + | NamedHyp id -> CT_coerce_ID_to_ID_OR_INT (xlate_ident id) + +let xlate_explicit_binding (h,c) = + CT_binding (xlate_quantified_hypothesis h, xlate_constr c) + +let xlate_bindings = function + | ImplicitBindings l -> + CT_coerce_FORMULA_LIST_to_SPEC_LIST + (CT_formula_list (List.map xlate_constr l)) + | ExplicitBindings l -> + CT_coerce_BINDING_LIST_to_SPEC_LIST + (CT_binding_list (List.map xlate_explicit_binding l)) + | NoBindings -> + CT_coerce_FORMULA_LIST_to_SPEC_LIST (CT_formula_list []) let strip_targ_spec_list = function @@ -849,6 +944,7 @@ let strip_targ_intropatt = | _ -> xlate_error "strip_targ_intropatt";; +(* let rec get_flag_rec = function | n1 :: tail -> @@ -873,7 +969,23 @@ let rec get_flag_rec = | Node (_, a, _) -> error ("get_flag_rec : unexpected flag " ^ a) | _ -> error "get_flag_rec : unexpected flag") | [] -> [], CT_unf [];; - +*) +let get_flag r = + let conv_flags, red_ids = + if r.rDelta then + [CT_delta], CT_unfbut (List.map qualid_or_meta_to_ct_ID r.rConst) + else + (if r.rConst = [] + then (* probably useless: just for compatibility *) [] + else [CT_delta]), + CT_unf (List.map qualid_or_meta_to_ct_ID r.rConst) in + let conv_flags = if r.rBeta then CT_beta::conv_flags else conv_flags in + let conv_flags = if r.rIota then CT_iota::conv_flags else conv_flags in + let conv_flags = if r.rZeta then CT_zeta::conv_flags else conv_flags in + (* Rem: EVAR flag obsolte *) + conv_flags, red_ids + +(* let rec xlate_intro_pattern = function | Node(_,"CONJPATTERN", l) -> @@ -886,10 +998,21 @@ let rec xlate_intro_pattern = CT_coerce_ID_to_INTRO_PATT(CT_ident c) | Node(_, a, _) -> failwith ("xlate_intro_pattern on node " ^ a) | _ -> failwith "xlate_intro_pattern";; +*) +let rec xlate_intro_pattern = + function + | IntroOrAndPattern [l] -> + CT_conj_pattern(CT_intro_patt_list (List.map xlate_intro_pattern l)) + | IntroOrAndPattern ll -> + let insert_conj l = CT_conj_pattern (CT_intro_patt_list + (List.map xlate_intro_pattern l)) + in CT_disj_pattern(CT_intro_patt_list (List.map insert_conj ll)) + | IntroWildcard -> xlate_error "TODO: '_' intro pattern" + | IntroIdentifier c -> CT_coerce_ID_to_INTRO_PATT(xlate_ident c) let compute_INV_TYPE_from_string = function "InversionClear" -> CT_inv_clear - | "HalfInversion" -> CT_inv_simple + | "SimpleInversion" -> CT_inv_simple | "Inversion" -> CT_inv_regular | _ -> failwith "unexpected Inversion type";; @@ -909,6 +1032,7 @@ let tactic_special_case cont_function cvt_arg = function | _ -> assert false;; let xlate_context_pattern = function +(* Node(_,"TERM", [Node(_, "COMMAND", [v])]) -> CT_coerce_FORMULA_to_CONTEXT_PATTERN (xlate_formula v) | Node(_,"SUBTERM", [Node(_,"COMMAND",[v])]) -> @@ -916,12 +1040,23 @@ let xlate_context_pattern = function | Node(_,"SUBTERM", [Nvar(_, s); Node(_, "COMMAND", [v])]) -> CT_context(ctf_ID_OPT_SOME (CT_ident s), xlate_formula v) | _ -> assert false;; +*) + | Term v -> + CT_coerce_FORMULA_to_CONTEXT_PATTERN (xlate_constr v) + | Subterm (idopt, v) -> + CT_context(xlate_ident_opt idopt, xlate_constr v) + let xlate_match_context_hyps = function +(* [b] -> CT_premise_pattern(ctv_ID_OPT_NONE, xlate_context_pattern b) | [Nvar(_,s);b] -> CT_premise_pattern(ctf_ID_OPT_SOME (CT_ident s), xlate_context_pattern b) | _ -> assert false;; +*) + | NoHypId b -> CT_premise_pattern(ctv_ID_OPT_NONE, xlate_context_pattern b) + | Hyp ((_,id),b) -> CT_premise_pattern(ctf_ID_OPT_SOME (xlate_ident id), + xlate_context_pattern b) let xlate_largs_to_id_unit largs = @@ -929,6 +1064,8 @@ let xlate_largs_to_id_unit largs = fst::rest -> fst, rest | _ -> assert false;; +(* Obsolete, partially replaced by xlate_tacarg and partially dispatched on + throughout the code for each tactic entry let rec cvt_arg = function | Nvar (_, id) -> Targ_ident (CT_ident id) @@ -991,8 +1128,75 @@ let rec cvt_arg = (string_of_int l1) ^ " " ^ (string_of_int l2)) | _ -> failwith "cvt_arg" +*) +let rec (xlate_tacarg:raw_tactic_arg -> ct_TACTIC_ARG) = + function + | TacVoid -> + CT_void + | Tacexp t -> + CT_coerce_TACTIC_COM_to_TACTIC_ARG(xlate_tactic t) + | Integer n -> + CT_coerce_ID_OR_INT_to_TACTIC_ARG + (CT_coerce_INT_to_ID_OR_INT (CT_int n)) + | Reference r -> + CT_coerce_ID_OR_INT_to_TACTIC_ARG + (CT_coerce_ID_to_ID_OR_INT (reference_to_ct_ID r)) + | TacDynamic _ -> + failwith "Dynamics not treated in xlate_ast" + | ConstrMayEval (ConstrTerm c) -> + CT_coerce_FORMULA_to_TACTIC_ARG (xlate_constr c) + | ConstrMayEval _ -> + xlate_error "TODO: Eval/Inst as tactic argument" + | MetaIdArg _ -> + xlate_error "MetaIdArg should only be used in quotations" + | MetaNumArg (_,n) -> + CT_coerce_FORMULA_to_TACTIC_ARG + (CT_coerce_ID_to_FORMULA(CT_metac (CT_int n))) + | t -> + CT_coerce_TACTIC_COM_to_TACTIC_ARG(xlate_call_or_tacarg t) + +and (xlate_call_or_tacarg:raw_tactic_arg -> ct_TACTIC_COM) = + function + (* Moved from xlate_tactic *) + | TacCall (_,Reference r, a::l) -> + CT_simple_user_tac + (reference_to_ct_ID r, + CT_tactic_arg_list(xlate_tacarg a,List.map xlate_tacarg l)) + | TacCall (_,_,_) -> xlate_error "" + | Reference (RIdent (_,s)) -> ident_tac s + | t -> xlate_error "TODO: result other than tactic or constr" + and xlate_red_tactic = function + | Red true -> xlate_error "" + | Red false -> CT_red + | Hnf -> CT_hnf + | Simpl -> CT_simpl + | Cbv flag_list -> + let conv_flags, red_ids = get_flag flag_list in + CT_cbv (CT_conversion_flag_list conv_flags, red_ids) + | Lazy flag_list -> + let conv_flags, red_ids = get_flag flag_list in + CT_cbv (CT_conversion_flag_list conv_flags, red_ids) + | Unfold unf_list -> + let ct_unf_list = List.map (fun (nums,qid) -> + CT_unfold_occ (CT_int_list (List.map (fun x -> CT_int x) nums), + qualid_or_meta_to_ct_ID qid)) unf_list in + (match ct_unf_list with + | first :: others -> CT_unfold (CT_unfold_ne_list (first, others)) + | [] -> error "there should be at least one thing to unfold") + | Fold formula_list -> + CT_fold(CT_formula_list(List.map xlate_constr formula_list)) + | Pattern l -> + let pat_list = List.map (fun (nums,c) -> + CT_pattern_occ + (CT_int_list (List.map (fun x -> CT_int x) nums), + xlate_constr c)) l in + (match pat_list with + | first :: others -> CT_pattern (CT_pattern_ne_list (first, others)) + | [] -> error "Expecting at least one pattern in a Pattern command") + | ExtraRedExpr _ -> xlate_error "TODO: ExtraRedExpr" +(* | Node (loc, s, []) -> (match s with | "Red" -> CT_red @@ -1038,14 +1242,20 @@ and xlate_red_tactic = formula_list)) | Node (_, a, _) -> error ("xlate_red_tactic: unexpected argument " ^ a) | _ -> error "xlate_red_tactic : unexpected argument" +*) and xlate_tactic = function - | Node (_, s, l) -> +(* | Node (_, s, l) -> (match s, l with +*) +(* | "FUN", [Node(_, "FUNVAR", largs); t] -> - let fst, rest = xlate_largs_to_id_unit largs in - CT_tactic_fun - (CT_id_unit_list(fst, rest), xlate_tactic t) +*) + | TacFun (largs, t) -> + let fst, rest = xlate_largs_to_id_unit largs in + CT_tactic_fun (CT_id_unit_list(fst, rest), xlate_tactic t) + | TacFunRec _ -> xlate_error "Merged with Tactic Definition" +(* | "TACTICLIST", (t :: tl) -> (match List.map xlate_tactic (t::tl) with | [] -> xlate_error "xlate_tactic: internal xlate_error" @@ -1054,17 +1264,37 @@ and xlate_tactic = | xt :: xtl -> CT_then (xt, xtl)) | "TACTICLIST", _ -> xlate_error "xlate_tactic: malformed tactic-expression TACTICLIST" +*) + | TacThen (t1, t2) -> + (match xlate_tactic t1 with + | CT_then(t,tl) -> CT_then (t, tl@[xlate_tactic t2]) + | xt1 -> CT_then (xt1, [xlate_tactic t2])) +(* | "TACLIST", (t :: tl) -> (match List.map xlate_tactic (t::tl) with | [] -> xlate_error "xlate_tactic: internal xlate_error" | xt :: [] -> xt | xt :: xtl -> CT_parallel (xt, xtl)) +*) + | TacThens (t, tl) -> CT_parallel (xlate_tactic t, List.map xlate_tactic tl) +(* | "FIRST", (a::l) -> +*) + | TacFirst [] -> xlate_error "" + | TacFirst (a::l) -> CT_first(xlate_tactic a,List.map xlate_tactic l) +(* | "TCLSOLVE", (a::l) -> +*) + | TacSolve [] -> xlate_error "" + | TacSolve (a::l) -> CT_tacsolve(xlate_tactic a, List.map xlate_tactic l) +(* | "DO", ((Num (_, n)) :: (t :: [])) -> CT_do (CT_int n, xlate_tactic t) | "DO", _ -> xlate_error "xlate_tactic: malformed tactic-expression DO" +*) + | TacDo (n, t) -> CT_do (CT_int n, xlate_tactic t) +(* | "TRY", (t :: []) -> CT_try (xlate_tactic t) | "TRY", _ -> xlate_error "xlate_tactic: malformed tactic-expression TRY" | "REPEAT", (t :: []) -> CT_repeat (xlate_tactic t) @@ -1074,12 +1304,27 @@ and xlate_tactic = | "INFO", (t :: []) -> CT_info (xlate_tactic t) | "REPEAT", _ -> xlate_error "xlate_tactic: malformed tactic-expression REPEAT" +*) + | TacTry t -> CT_try (xlate_tactic t) + | TacRepeat t -> CT_repeat (xlate_tactic t) + | TacAbstract (t, None) -> CT_abstract(ctv_ID_OPT_NONE, (xlate_tactic t)) + | TacAbstract (t, Some id) -> + CT_abstract(ctf_ID_OPT_SOME (xlate_ident id), (xlate_tactic t)) + | TacInfo t -> CT_info (xlate_tactic t) + | TacProgress t -> xlate_error "TODO: Progress t" +(* | "ORELSE", (t1 :: (t2 :: [])) -> CT_orelse (xlate_tactic t1, xlate_tactic t2) | "ORELSE", _ -> xlate_error "xlate_tactic: malformed tactic-expression ORELSE" +*) + | TacOrelse (t1, t2) -> CT_orelse (xlate_tactic t1, xlate_tactic t2) + +(* | ((s, l) as it) when (is_tactic_special_case s) -> tactic_special_case xlate_tactic cvt_arg it +*) +(* moved to xlate_call_or_tacarg | "APP", (Nvar(_,s))::l -> let args = List.map (function @@ -1093,6 +1338,8 @@ and xlate_tactic = fst::args2 -> fst, args2 | _ -> assert false in CT_simple_user_tac(CT_ident s, CT_tactic_arg_list(fst, args2)) +*) +(* | "MATCH", exp::rules -> CT_match_tac(mk_let_value exp, match List.map @@ -1119,9 +1366,33 @@ and xlate_tactic = | [] -> assert false | fst::others -> CT_match_tac_rules(fst, others)) +*) + | TacMatch (exp, rules) -> + CT_match_tac(CT_coerce_DEF_BODY_to_LET_VALUE(formula_to_def_body exp), + match List.map + (function + | Pat ([],p,tac) -> + CT_match_tac_rule(xlate_context_pattern p, + mk_let_value tac) + | Pat (_,p,tac) -> xlate_error"No hyps in pure Match" + | All tac -> + CT_match_tac_rule + (CT_coerce_FORMULA_to_CONTEXT_PATTERN + CT_existvarc, + mk_let_value tac)) rules with + | [] -> assert false + | fst::others -> + CT_match_tac_rules(fst, others)) + +(* | "MATCHCONTEXT", rule1::rules -> +*) + | TacMatchContext (_,[]) -> failwith "" + | TacMatchContext (lr,rule1::rules) -> + (* TODO : traiter la direction "lr" *) CT_match_context(xlate_context_rule rule1, List.map xlate_context_rule rules) +(* | "LET", [Node(_, "LETDECL",l); t] -> let cvt_clause = @@ -1136,30 +1407,87 @@ and xlate_tactic = (xlate_tactic v)) | Node(_, s, _) -> failwith ("cvt_clause : unexpected " ^ s) | _ -> assert false in +*) + | TacLetIn (l, t) -> + let cvt_clause = + function + ((_,s),None,ConstrMayEval v) -> + CT_let_clause(xlate_ident s, + CT_coerce_DEF_BODY_to_LET_VALUE + (formula_to_def_body v)) + | ((_,s),None,Tacexp t) -> + CT_let_clause(xlate_ident s, + CT_coerce_TACTIC_COM_to_LET_VALUE + (xlate_tactic t)) + | ((_,s),None,t) -> + CT_let_clause(xlate_ident s, + CT_coerce_TACTIC_COM_to_LET_VALUE + (xlate_call_or_tacarg t)) + | ((_,s),Some c,v) -> xlate_error "TODO: Let id : c := t In t'" in let cl_l = List.map cvt_clause l in (match cl_l with | [] -> assert false | fst::others -> CT_lettac (CT_let_clauses(fst, others), mk_let_value t)) + | TacLetCut _ -> xlate_error "Unclear future of syntax Let x := t" + | TacLetRecIn _ -> xlate_error "TODO: Rec x = t In" + +(* | s, l -> xlate_tac (s, List.map cvt_arg l)) +*) + | TacAtom (_, t) -> xlate_tac t +(* was in xlate_tac *) + | TacFail 0 -> CT_fail + | TacFail n -> xlate_error "TODO: Fail n" + | TacId -> CT_idtac +(* moved to xlate_call_or_tacarg | Nvar(_, s) -> ident_tac s +*) +(* | the_node -> xlate_error ("xlate_tactic at " ^ (string_of_node_loc the_node) ) +*) + | TacArg a -> xlate_call_or_tacarg a and xlate_tac = - function + function +(* | "Absurd", ((Targ_command c) :: []) -> CT_absurd c | "Change", [Targ_command f; Targ_id_list b] -> CT_change(f,b) | "Contradiction", [] -> CT_contradiction +*) + | TacExtend ("Absurd",[c]) -> + CT_absurd (xlate_constr (out_gen rawwit_constr c)) + | TacChange (f, b) -> CT_change (xlate_constr f, xlate_clause b) + | TacExtend ("Contradiction",[]) -> CT_contradiction +(* | "DoubleInd", ((Targ_int n1) :: ((Targ_int n2) :: [])) -> - CT_tac_double (n1, n2) + CT_tac_double (n1, n2) +*) + | TacDoubleInduction (AnonHyp n1, AnonHyp n2) -> + CT_tac_double (CT_int n1, CT_int n2) + | TacDoubleInduction _ -> xlate_error "TODO: Double Induction id1 id2" +(* | "Discr", [] -> CT_discriminate_eq ctv_ID_OPT_NONE | "DiscrHyp", ((Targ_ident id) :: []) -> CT_discriminate_eq (ctf_ID_OPT_SOME id) | "DEqConcl", [] -> CT_simplify_eq ctv_ID_OPT_NONE | "DEqHyp", ((Targ_ident id) :: []) -> CT_simplify_eq (ctf_ID_OPT_SOME id) +*) + | TacExtend ("Discriminate", [idopt]) -> + CT_discriminate_eq + (xlate_ident_opt (out_gen (wit_opt rawwit_ident) idopt)) + | TacExtend ("DEq", [idopt]) -> + CT_simplify_eq + (xlate_ident_opt (out_gen (wit_opt rawwit_ident) idopt)) +(* | "Inj", [] -> CT_injection_eq ctv_ID_OPT_NONE | "InjHyp", ((Targ_ident id) :: []) -> CT_injection_eq (ctf_ID_OPT_SOME id) +*) + | TacExtend ("Injection", [idopt]) -> + CT_injection_eq + (xlate_ident_opt (out_gen (wit_opt rawwit_ident) idopt)) +(* | "Fix", ((Targ_int n) :: []) -> CT_fixtactic (ctv_ID_OPT_NONE, n, CT_fix_tac_list []) | "Fix", ((Targ_ident id) :: ((Targ_int n) :: fixtac_list)) -> @@ -1171,40 +1499,126 @@ and xlate_tac = CT_cofixtactic (CT_coerce_ID_to_ID_OPT id, CT_cofix_tac_list (List.map strip_targ_cofixtac cofixtac_list)) +*) + | TacFix (idopt, n) -> + CT_fixtactic (xlate_ident_opt idopt, CT_int n, CT_fix_tac_list []) + | TacMutualFix (id, n, fixtac_list) -> + let f (id,n,c) = CT_fixtac (xlate_ident id, CT_int n, xlate_constr c) in + CT_fixtactic + (ctf_ID_OPT_SOME (xlate_ident id), CT_int n, + CT_fix_tac_list (List.map f fixtac_list)) + | TacCofix idopt -> + CT_cofixtactic (xlate_ident_opt idopt, CT_cofix_tac_list []) + | TacMutualCofix (id, cofixtac_list) -> + let f (id,c) = CT_cofixtac (xlate_ident id, xlate_constr c) in + CT_cofixtactic + (CT_coerce_ID_to_ID_OPT (xlate_ident id), + CT_cofix_tac_list (List.map f cofixtac_list)) +(* | "IntrosUntil", ((Targ_ident id) :: []) -> CT_intros_until id +*) + | TacIntrosUntil (NamedHyp id) -> CT_intros_until (xlate_ident id) + | TacIntrosUntil (AnonHyp n) -> xlate_error "TODO: Intros until n" +(* | "IntroMove", [Targ_ident id1;Targ_ident id2] -> - CT_intro_after(CT_coerce_ID_to_ID_OPT id1, id2) +*) + | TacIntroMove (Some id1, Some (_,id2)) -> + CT_intro_after(CT_coerce_ID_to_ID_OPT (xlate_ident id1),xlate_ident id2) +(* | "IntroMove", [Targ_ident id2] -> - CT_intro_after(CT_coerce_NONE_to_ID_OPT CT_none, id2) +*) + | TacIntroMove (None, Some (_,id2)) -> + CT_intro_after(CT_coerce_NONE_to_ID_OPT CT_none, xlate_ident id2) +(* | "MoveDep", [Targ_ident id1;Targ_ident id2] -> - CT_move_after(id1, id2) +*) + | TacMove (true, (_,id1), (_,id2)) -> + CT_move_after(xlate_ident id1, xlate_ident id2) + | TacMove (false, id1, id2) -> xlate_error "Non dep Move is only internal" +(* | "Intros", [] -> CT_intros (CT_intro_patt_list []) | "Intros", [patt_list] -> CT_intros (strip_targ_intropatt patt_list) +*) + | TacIntroPattern [] -> CT_intros (CT_intro_patt_list []) + | TacIntroPattern patt_list -> + CT_intros (CT_intro_patt_list (List.map xlate_intro_pattern patt_list)) +(* | "Intro", [Targ_ident (CT_ident id)] -> - CT_intros (CT_intro_patt_list [CT_coerce_ID_to_INTRO_PATT(CT_ident id)]) +*) + | TacIntroMove (Some id, None) -> + CT_intros (CT_intro_patt_list[CT_coerce_ID_to_INTRO_PATT(xlate_ident id)]) + | TacIntroMove (None, None) -> xlate_error "TODO: Intro" +(* | "Left", (bindl :: []) -> CT_left (strip_targ_spec_list bindl) | "Right", (bindl :: []) -> CT_right (strip_targ_spec_list bindl) | "Split", (bindl :: []) -> CT_split (strip_targ_spec_list bindl) +*) + | TacLeft bindl -> CT_left (xlate_bindings bindl) + | TacRight bindl -> CT_right (xlate_bindings bindl) + | TacSplit bindl -> CT_split (xlate_bindings bindl) +(* | "Replace", ((Targ_command c1) :: ((Targ_command c2) :: [])) -> CT_replace_with (c1, c2) +*) + | TacExtend ("Replace", [c1; c2]) -> + let c1 = xlate_constr (out_gen rawwit_constr c1) in + let c2 = xlate_constr (out_gen rawwit_constr c2) in + CT_replace_with (c1, c2) | (*Changes to Equality.v some more rewrite possibilities *) +(* "RewriteLR", [(Targ_command c); bindl] -> CT_rewrite_lr (c, strip_targ_spec_list bindl, ctv_ID_OPT_NONE) - | "RewriteLRin", ((Targ_ident id) :: ((Targ_command c) :: (bindl::[]))) -> - CT_rewrite_lr (c, strip_targ_spec_list bindl, ctf_ID_OPT_SOME id) | "RewriteRL", [Targ_command c; bindl] -> CT_rewrite_rl (c, strip_targ_spec_list bindl, ctv_ID_OPT_NONE) +*) + TacExtend ("Rewrite", [b; cbindl]) -> + let b = out_gen Extraargs.rawwit_orient b in + let (c,bindl) = out_gen rawwit_constr_with_bindings cbindl in + let c = xlate_constr c and bindl = xlate_bindings bindl in + if b then CT_rewrite_lr (c, bindl, ctv_ID_OPT_NONE) + else CT_rewrite_rl (c, bindl, ctv_ID_OPT_NONE) +(* + | "RewriteLRin", ((Targ_ident id) :: ((Targ_command c) :: (bindl::[]))) -> + CT_rewrite_lr (c, strip_targ_spec_list bindl, ctf_ID_OPT_SOME id) | "RewriteRLin", [Targ_ident id; Targ_command c; bindl] -> CT_rewrite_rl (c, strip_targ_spec_list bindl, ctf_ID_OPT_SOME id) +*) + | TacExtend ("RewriteIn", [b; cbindl; id]) -> + let b = out_gen Extraargs.rawwit_orient b in + let (c,bindl) = out_gen rawwit_constr_with_bindings cbindl in + let c = xlate_constr c and bindl = xlate_bindings bindl in + let id = ctf_ID_OPT_SOME (xlate_ident (out_gen rawwit_ident id)) in + if b then CT_rewrite_lr (c, bindl, id) + else CT_rewrite_rl (c, bindl, id) +(* | "CondRewriteLR", [Targ_tacexp t; Targ_command c; bindl] -> CT_condrewrite_lr (t, c, strip_targ_spec_list bindl, ctv_ID_OPT_NONE) | "CondRewriteRL", [Targ_tacexp t; Targ_command c; bindl] -> CT_condrewrite_rl (t, c, strip_targ_spec_list bindl, ctv_ID_OPT_NONE) +*) + | TacExtend ("ConditionalRewrite", [t; b; cbindl]) -> + let t = out_gen rawwit_tactic t in + let b = out_gen Extraargs.rawwit_orient b in + let (c,bindl) = out_gen rawwit_constr_with_bindings cbindl in + let c = xlate_constr c and bindl = xlate_bindings bindl in + if b then CT_condrewrite_lr (xlate_tactic t, c, bindl, ctv_ID_OPT_NONE) + else CT_condrewrite_rl (xlate_tactic t, c, bindl, ctv_ID_OPT_NONE) +(* | "CondRewriteLRin", [Targ_tacexp t; Targ_ident id; Targ_command c; bindl] -> CT_condrewrite_lr (t, c, strip_targ_spec_list bindl, ctf_ID_OPT_SOME id) | "CondRewriteRLin", [Targ_tacexp t; Targ_ident id; Targ_command c; bindl] -> CT_condrewrite_rl (t, c, strip_targ_spec_list bindl, ctf_ID_OPT_SOME id) +*) + | TacExtend ("ConditionalRewriteIn", [t; b; cbindl; id]) -> + let t = out_gen rawwit_tactic t in + let b = out_gen Extraargs.rawwit_orient b in + let (c,bindl) = out_gen rawwit_constr_with_bindings cbindl in + let c = xlate_constr c and bindl = xlate_bindings bindl in + let id = ctf_ID_OPT_SOME (xlate_ident (out_gen rawwit_ident id)) in + if b then CT_condrewrite_lr (xlate_tactic t, c, bindl, id) + else CT_condrewrite_rl (xlate_tactic t, c, bindl, id) +(* | "SubstConcl_LR", ((Targ_command c) :: []) -> CT_cutrewrite_lr (c, ctv_ID_OPT_NONE) | "SubstHyp_LR", ((Targ_command c) :: ((Targ_ident id) :: [])) -> @@ -1215,16 +1629,53 @@ and xlate_tac = CT_cutrewrite_rl (c, ctf_ID_OPT_SOME id) | "SubstHypInConcl_LR", ((Targ_ident id) :: []) -> CT_deprewrite_lr id | "SubstHypInConcl_RL", ((Targ_ident id) :: []) -> CT_deprewrite_rl id +*) + | TacExtend ("DependentRewrite", [b; id_or_constr]) -> + let b = out_gen Extraargs.rawwit_orient b in + (match genarg_tag id_or_constr with + | IdentArgType -> (*Dependent Rewrite/SubstHypInConcl*) + let id = xlate_ident (out_gen rawwit_ident id_or_constr) in + if b then CT_deprewrite_lr id else CT_deprewrite_rl id + | ConstrArgType -> (*CutRewrite/SubstConcl*) + let c = xlate_constr (out_gen rawwit_constr id_or_constr) in + if b then CT_cutrewrite_lr (c, ctv_ID_OPT_NONE) + else CT_cutrewrite_rl (c, ctv_ID_OPT_NONE) + | _ -> xlate_error "") + | TacExtend ("DependentRewrite", [b; c; id]) -> (*CutRewrite in/SubstHyp*) + let b = out_gen Extraargs.rawwit_orient b in + let c = xlate_constr (out_gen rawwit_constr c) in + let id = xlate_ident (out_gen rawwit_ident id) in + if b then CT_cutrewrite_lr (c, ctf_ID_OPT_SOME id) + else CT_cutrewrite_lr (c, ctf_ID_OPT_SOME id) +(* | "Reflexivity", [] -> CT_reflexivity | "Symmetry", [] -> CT_symmetry | "Transitivity", ((Targ_command c) :: []) -> CT_transitivity c +*) + | TacReflexivity -> CT_reflexivity + | TacSymmetry -> CT_symmetry + | TacTransitivity c -> CT_transitivity (xlate_constr c) +(* | "Assumption", [] -> CT_assumption +*) + | TacAssumption -> CT_assumption +(* Moved to xlate_tactic | "FAIL", [] -> CT_fail | "IDTAC", [] -> CT_idtac +*) +(* | "Exact", ((Targ_command c) :: []) -> CT_exact c +*) + | TacExact c -> CT_exact (xlate_constr c) +(* | "DHyp", [Targ_ident id] -> CT_dhyp id | "CDHyp", [Targ_ident id] -> CT_cdhyp id | "DConcl", [] -> CT_dconcl +*) + | TacDestructHyp (true, (_,id)) -> CT_cdhyp (xlate_ident id) + | TacDestructHyp (false, (_,id)) -> CT_dhyp (xlate_ident id) + | TacDestructConcl -> CT_dconcl +(* | "SuperAuto", [a1;a2;a3;a4] -> CT_superauto( (match a1 with @@ -1242,8 +1693,19 @@ and xlate_tac = | _ -> (CT_coerce_NONE_to_USINGTDB CT_none))) +*) + | TacSuperAuto (nopt,l,a3,a4) -> + CT_superauto( + xlate_int_opt nopt, + xlate_qualid_list l, + (if a3 then CT_destructing else CT_coerce_NONE_to_DESTRUCTING CT_none), + (if a4 then CT_usingtdb else CT_coerce_NONE_to_USINGTDB CT_none)) +(* | "AutoTDB", [Targ_int n] -> CT_autotdb (CT_coerce_INT_to_INT_OPT n) | "AutoTDB", [] -> CT_autotdb (CT_coerce_NONE_to_INT_OPT CT_none) +*) + | TacAutoTDB nopt -> CT_autotdb (xlate_int_opt nopt) +(* | "Auto", ((Targ_int n) :: []) -> CT_auto (CT_coerce_INT_to_INT_OPT n) | "Auto", ((Targ_string (CT_string "*"))::[]) -> CT_auto_with((CT_coerce_NONE_to_INT_OPT CT_none), CT_star) @@ -1264,6 +1726,14 @@ and xlate_tac = idl))) | "Auto", ((Targ_int n) :: ((Targ_string (CT_string "*")) :: [])) -> CT_auto_with ((CT_coerce_INT_to_INT_OPT n), CT_star) +*) + | TacAuto (nopt, Some []) -> CT_auto (xlate_int_opt nopt) + | TacAuto (nopt, None) -> CT_auto_with (xlate_int_opt nopt, CT_star) + | TacAuto (nopt, Some (id1::idl)) -> + CT_auto_with(xlate_int_opt nopt, + CT_coerce_ID_NE_LIST_to_ID_NE_LIST_OR_STAR( + CT_id_ne_list(CT_ident id1, List.map (fun x -> CT_ident x) idl))) +(* | "EAuto", ((Targ_int n) :: []) -> CT_eauto (CT_coerce_INT_to_INT_OPT n) | "EAuto", [] -> CT_eauto (CT_coerce_NONE_to_INT_OPT CT_none) | "EAuto", ((Targ_int n) :: ((Targ_ident id1) :: idl)) -> @@ -1284,12 +1754,42 @@ and xlate_tac = CT_eauto_with ((CT_coerce_INT_to_INT_OPT n), CT_star) | "EAuto", ((Targ_string (CT_string "*"))::[]) -> CT_eauto_with((CT_coerce_NONE_to_INT_OPT CT_none), CT_star) +*) + | TacExtend ("EAuto", [nopt; popt; idl]) -> + let control = + match out_gen (wit_opt rawwit_int_or_var) nopt with + | Some breadth -> Some (true, breadth) + | None -> + match out_gen (wit_opt rawwit_int_or_var) popt with + | Some depth -> Some (false, depth) + | None -> None in + let idl = out_gen (wit_opt (wit_list0 rawwit_string)) idl in + xlate_error "TODO: EAuto n p" + (* Something like: + match idl with + | None -> CT_eauto_with (..., CT_star) + | Some [] -> CT_eauto ... + | Some (id::l) -> CT_eauto_with (..., ...) + *) +(* | "Prolog", ((Targ_int n) :: idlist) -> (*according to coqdev the code is right, they want formula *) CT_prolog (CT_formula_list (List.map strip_targ_command idlist), n) - | (**) +*) + | TacExtend ("Prolog", [cl; n]) -> + let cl = List.map xlate_constr (out_gen (wit_list0 rawwit_constr) cl) in + (match out_gen wit_int_or_var n with + | ArgVar _ -> xlate_error "" + | ArgArg n -> CT_prolog (CT_formula_list cl, CT_int n)) +(* "EApplyWithBindings", ((Targ_command c) :: (bindl :: [])) -> CT_eapply (c, strip_targ_spec_list bindl) +*) + | TacExtend ("EApply", [cbindl]) -> + let (c,bindl) = out_gen rawwit_constr_with_bindings cbindl in + let c = xlate_constr c and bindl = xlate_bindings bindl in + CT_eapply (c, bindl) +(* | "Trivial", [] -> CT_trivial | "Trivial", ((Targ_string (CT_string "*"))::[]) -> CT_trivial_with(CT_star) @@ -1299,25 +1799,61 @@ and xlate_tac = List.map (function Targ_ident x -> x | _ -> xlate_error "Trivial expects identifiers") idl)))) +*) + | TacTrivial (Some []) -> CT_trivial + | TacTrivial None -> CT_trivial_with(CT_star) + | TacTrivial (Some (id1::idl)) -> + CT_trivial_with(CT_coerce_ID_NE_LIST_to_ID_NE_LIST_OR_STAR( + (CT_id_ne_list(CT_ident id1,List.map (fun x -> CT_ident x) idl)))) +(* | "Reduce", ((Targ_redexp id) :: ((Targ_id_list l) :: [])) -> CT_reduce (id, l) +*) + | TacReduce (red, l) -> + CT_reduce (xlate_red_tactic red, xlate_clause l) +(* | "Apply", ((Targ_command c) :: (bindl :: [])) -> CT_apply (c, strip_targ_spec_list bindl) +*) + | TacApply (c,bindl) -> + CT_apply (xlate_constr c, xlate_bindings bindl) +(* | "Constructor", ((Targ_int n) :: (bindl :: [])) -> CT_constructor (n, strip_targ_spec_list bindl) +*) + | TacConstructor (n_or_meta, bindl) -> + let n = match n_or_meta with AI n -> n | MetaId _ -> xlate_error "" + in CT_constructor (CT_int n, xlate_bindings bindl) +(* | "Specialize", ((Targ_int n) :: ((Targ_command c) :: ((Targ_spec_list sl) :: []))) -> CT_specialize (CT_coerce_INT_to_INT_OPT n, c, sl) | "Specialize", ((Targ_command c) :: ((Targ_spec_list sl) :: [])) -> CT_specialize (CT_coerce_NONE_to_INT_OPT CT_none, c, sl) +*) + | TacSpecialize (nopt, (c,sl)) -> + CT_specialize (xlate_int_opt nopt, xlate_constr c, xlate_bindings sl) +(* | "Generalize", (first :: cl) -> CT_generalize (CT_formula_ne_list (strip_targ_command first, List.map strip_targ_command cl)) | "GeneralizeDep", [Targ_command c] -> CT_generalize_dependent c +*) + | TacGeneralize [] -> xlate_error "" + | TacGeneralize (first :: cl) -> + CT_generalize + (CT_formula_ne_list (xlate_constr first, List.map xlate_constr cl)) + | TacGeneralizeDep c -> + CT_generalize_dependent (xlate_constr c) +(* | "ElimType", ((Targ_command c) :: []) -> CT_elim_type c | "CaseType", ((Targ_command c) :: []) -> CT_case_type c +*) + | TacElimType c -> CT_elim_type (xlate_constr c) + | TacCaseType c -> CT_case_type (xlate_constr c) +(* | "Elim", ((Targ_command c1) :: ((Targ_spec_list sl) :: [])) -> CT_elim (c1, sl, CT_coerce_NONE_to_USING CT_none) | "Elim", @@ -1327,28 +1863,67 @@ and xlate_tac = CT_elim (c1, sl, CT_using (c2, sl2)) | "Case", ((Targ_command c1) :: ((Targ_spec_list sl) :: [])) -> CT_casetac (c1, sl) +*) + | TacElim ((c1,sl), None) -> + CT_elim (xlate_constr c1, xlate_bindings sl, + CT_coerce_NONE_to_USING CT_none) + | TacElim ((c1,sl), Some (c2,sl2)) -> + CT_elim (xlate_constr c1, xlate_bindings sl, + CT_using (xlate_constr c2, xlate_bindings sl2)) + | TacCase (c1,sl) -> + CT_casetac (xlate_constr c1, xlate_bindings sl) +(* | "Induction", ((Targ_ident id) :: []) -> CT_induction (CT_coerce_ID_to_ID_OR_INT id) | "Induction", ((Targ_int n) :: []) -> CT_induction (CT_coerce_INT_to_ID_OR_INT n) +*) + | TacOldInduction h -> CT_induction (xlate_quantified_hypothesis h) +(* | "Destruct", ((Targ_ident id) :: []) -> CT_destruct (CT_coerce_ID_to_ID_OR_INT id) | "Destruct", ((Targ_int n) :: []) -> CT_destruct (CT_coerce_INT_to_ID_OR_INT n) +*) + | TacOldDestruct h -> CT_destruct (xlate_quantified_hypothesis h) +(* | "Cut", ((Targ_command c) :: []) -> CT_cut c +*) + | TacCut c -> CT_cut (xlate_constr c) +(* | "CutAndApply", ((Targ_command c) :: []) -> CT_use c +*) + | TacLApply c -> CT_use (xlate_constr c) +(* | "DecomposeThese", ((Targ_id_list l) :: ((Targ_command c) :: [])) -> (match l with CT_id_list (id :: l') -> CT_decompose_list( CT_id_ne_list(id,l'),c) | _ -> xlate_error "DecomposeThese : empty list of identifiers?") +*) + | TacDecompose ([],c) -> + xlate_error "Decompose : empty list of identifiers?" + | TacDecompose (id::l,c) -> + let id' = qualid_or_meta_to_ct_ID id in + let l' = List.map qualid_or_meta_to_ct_ID l in + CT_decompose_list(CT_id_ne_list(id',l'),xlate_constr c) + | TacDecomposeAnd c -> xlate_error "TODO: Decompose Record" + | TacDecomposeOr c -> xlate_error "TODO: Decompose Sum" +(* | "Clear", [id_list] -> (match id_list with Targ_id_list(CT_id_list(id::idl)) -> CT_clear (CT_id_ne_list (id, idl)) | _ -> xlate_error "Clear expects a non empty list of identifiers") +*) + | TacClear [] -> + xlate_error "Clear expects a non empty list of identifiers" + | TacClear (id::idl) -> + let idl' = List.map ident_or_meta_to_ct_ID idl in + CT_clear (CT_id_ne_list (ident_or_meta_to_ct_ID id, idl')) | (*For translating tactics/Inv.v *) +(* "Inv", [Targ_string (CT_string s); Targ_ident id] -> CT_inversion (compute_INV_TYPE_from_string s, id, CT_id_list []) | "InvIn", ((Targ_ident (CT_ident s))::((Targ_ident id) :: idlist)) -> @@ -1362,15 +1937,107 @@ and xlate_tac = ((Targ_ident id) :: ((Targ_command c) :: []))) -> CT_depinversion (compute_INV_TYPE_from_string s, id, CT_coerce_FORMULA_to_FORMULA_OPT c) +*) + TacExtend ("SimpleInversion"|"Inversion"|"InversionClear" as s, [id]) -> + let id = xlate_ident (out_gen rawwit_ident id) in + CT_inversion (compute_INV_TYPE_from_string s, id, CT_id_list []) + | TacExtend ("SimpleInversion"|"Inversion"|"InversionClear" as s, + [id;copt_or_idl]) -> + let id = xlate_ident (out_gen rawwit_ident id) in + (match genarg_tag copt_or_idl with + | List1ArgType IdentArgType -> (* InvIn *) + let idl = out_gen (wit_list1 rawwit_ident) copt_or_idl in + CT_inversion (compute_INV_TYPE_from_string s, id, + CT_id_list (List.map xlate_ident idl)) + | OptArgType ConstrArgType -> (* DInv *) + let copt = out_gen (wit_opt rawwit_constr) copt_or_idl in + CT_depinversion + (compute_INV_TYPE_from_string s, id, xlate_constr_opt copt) + | _ -> xlate_error "") +(* | "UseInversionLemma", ((Targ_ident id) :: ((Targ_command c) :: [])) -> CT_use_inversion (id, c, CT_id_list []) | "UseInversionLemmaIn", ((Targ_ident id) :: ((Targ_command c) :: idlist)) -> CT_use_inversion (id, c, CT_id_list (List.map strip_targ_ident idlist)) +*) + | TacExtend ("InversionUsing", [id; c]) -> + let id = xlate_ident (out_gen rawwit_ident id) in + let c = out_gen rawwit_constr c in + CT_use_inversion (id, xlate_constr c, CT_id_list []) + | TacExtend ("InversionUsing", [id; c; idlist]) -> + let id = xlate_ident (out_gen rawwit_ident id) in + let c = out_gen rawwit_constr c in + let idlist = out_gen (wit_list1 rawwit_ident) idlist in + CT_use_inversion (id, xlate_constr c, + CT_id_list (List.map xlate_ident idlist)) +(* | "Omega", [] -> CT_omega +*) + | TacExtend ("Omega", []) -> CT_omega +(* | "APP", (Targ_ident id)::l -> CT_user_tac (id, CT_targ_list (List.map coerce_iTARG_to_TARG l)) +*) + | TacRename (_, _) -> xlate_error "TODO: Rename id into id'" + | TacClearBody _ -> xlate_error "TODO: Clear Body H" + | TacDAuto (_, _) -> xlate_error "TODO: DAuto" + | TacNewDestruct _ -> xlate_error "TODO: NewDestruct" + | TacNewInduction _ -> xlate_error "TODO: NewInduction" + | TacInstantiate (_, _) -> xlate_error "TODO: Instantiate" + | TacLetTac (_, _, _) -> xlate_error "TODO: LetTac" + | TacForward (_, _, _) -> xlate_error "TODO: Assert/Pose id:=c" + | TacTrueCut (_, _) -> xlate_error "TODO: Assert id:t" + | TacAnyConstructor tacopt -> xlate_error "TODO: Constructor tac" +(* | s, l -> CT_user_tac (CT_ident s, CT_targ_list (List.map coerce_iTARG_to_TARG l)) +*) + | TacExtend (id, l) -> + CT_user_tac (CT_ident id, CT_targ_list (List.map coerce_genarg_to_TARG l)) + | TacAlias (_, _, _) -> xlate_error "TODO: aliases" + +and coerce_genarg_to_TARG x = + match Genarg.genarg_tag x with + (* Basic types *) + | BoolArgType -> xlate_error "TODO: generic boolean argument" + | IntArgType -> + let n = out_gen rawwit_int x in + CT_coerce_ID_OR_INT_to_TARG (CT_coerce_INT_to_ID_OR_INT (CT_int n)) + | IntOrVarArgType -> + let x = match out_gen rawwit_int_or_var x with + | ArgArg n -> CT_coerce_INT_to_ID_OR_INT (CT_int n) + | ArgVar (_,id) -> CT_coerce_ID_to_ID_OR_INT (xlate_ident id) in + CT_coerce_ID_OR_INT_to_TARG x + | StringArgType -> + let s = CT_string (out_gen rawwit_string x) in + CT_coerce_ID_OR_STRING_to_TARG (CT_coerce_STRING_to_ID_OR_STRING s) + | PreIdentArgType -> + let id = CT_ident (out_gen rawwit_pre_ident x) in + CT_coerce_ID_OR_INT_to_TARG (CT_coerce_ID_to_ID_OR_INT id) + | IdentArgType -> + let id = xlate_ident (out_gen rawwit_ident x) in + CT_coerce_ID_OR_INT_to_TARG (CT_coerce_ID_to_ID_OR_INT id) + | QualidArgType -> + let id = tac_qualid_to_ct_ID (snd (out_gen rawwit_qualid x)) in + CT_coerce_ID_OR_INT_to_TARG (CT_coerce_ID_to_ID_OR_INT id) + (* Specific types *) + | ConstrArgType -> + CT_coerce_FORMULA_to_TARG (xlate_constr (out_gen rawwit_constr x)) + | ConstrMayEvalArgType -> xlate_error"TODO: generic constr-may-eval argument" + | QuantHypArgType ->xlate_error"TODO: generic quantified hypothesis argument" + | TacticArgType -> + let t = xlate_tactic (out_gen rawwit_tactic x) in + CT_coerce_TACTIC_COM_to_TARG t + | CastedOpenConstrArgType -> xlate_error "TODO: open constr" + | ConstrWithBindingsArgType -> xlate_error "TODO: constr with bindings" + | RedExprArgType -> xlate_error "TODO: red expr as generic argument" + | List0ArgType l -> xlate_error "TODO: lists of generic arguments" + | List1ArgType l -> xlate_error "TODO: non empty lists of generic arguments" + | OptArgType x -> xlate_error "TODO: optional generic arguments" + | PairArgType (u,v) -> xlate_error "TODO: pairs of generic arguments" + | ExtraArgType s -> xlate_error "Cannot treat extra generic arguments" + +(* and (xlate_context_rule: Ctast.t -> ct_CONTEXT_RULE) = function | Node(_, "MATCHCONTEXTRULE", parts) -> @@ -1384,6 +2051,16 @@ and (xlate_context_rule: Ctast.t -> ct_CONTEXT_RULE) = let hyps, cpat, tactic = xlate_ctxt_rule_aux parts in CT_context_rule(CT_context_hyp_list hyps, cpat, tactic) | _ -> assert false +*) +and xlate_context_rule = + function + | Pat (hyps, concl_pat, tactic) -> + CT_context_rule( + CT_context_hyp_list (List.map xlate_match_context_hyps hyps), + xlate_context_pattern concl_pat, xlate_tactic tactic) + | All te -> + xlate_error "TODO: wildcard match_context_rule" +(* and (formula_to_def_body : Ctast.t -> ct_DEF_BODY) = function | Node(_, "EVAL", [f;Node(_, "REDEXP", [tac])]) -> @@ -1404,11 +2081,29 @@ and (formula_to_def_body : Ctast.t -> ct_DEF_BODY) = | _ -> failwith("error raised inside formula_to_def_body (3) " ^ s)) +*) +and formula_to_def_body = + function + | ConstrEval (red, f) -> + CT_coerce_EVAL_CMD_to_DEF_BODY( + CT_eval(CT_coerce_NONE_to_INT_OPT CT_none, + xlate_red_tactic red, xlate_constr f)) + | ConstrContext _ -> xlate_error "TODO: Inst" + | ConstrTypeOf _ -> xlate_error "TODO: Check" + | ConstrTerm c -> ct_coerce_FORMULA_to_DEF_BODY(xlate_constr c) + +(* and mk_let_value = function Node(_, "COMMAND", [v]) -> CT_coerce_DEF_BODY_to_LET_VALUE(formula_to_def_body v) | v -> CT_coerce_TACTIC_COM_to_LET_VALUE(xlate_tactic v);; +*) +and mk_let_value = function + TacArg (ConstrMayEval v) -> + CT_coerce_DEF_BODY_to_LET_VALUE(formula_to_def_body v) + | v -> CT_coerce_TACTIC_COM_to_LET_VALUE(xlate_tactic v);; +(* let strip_varg_int = function | Varg_int n -> n @@ -1428,14 +2123,75 @@ let strip_varg_binder = function | Varg_binder n -> n | _ -> xlate_error "strip vernac: non-binder argument";; - +*) + +let coerce_genarg_to_VARG x = + match Genarg.genarg_tag x with + (* Basic types *) + | BoolArgType -> xlate_error "TODO: generic boolean argument" + | IntArgType -> + let n = out_gen rawwit_int x in + CT_coerce_INT_OPT_to_VARG (CT_coerce_INT_to_INT_OPT (CT_int n)) + | IntOrVarArgType -> + (match out_gen rawwit_int_or_var x with + | ArgArg n -> + CT_coerce_INT_OPT_to_VARG (CT_coerce_INT_to_INT_OPT (CT_int n)) + | ArgVar (_,id) -> + CT_coerce_ID_OPT_OR_ALL_to_VARG + (CT_coerce_ID_OPT_to_ID_OPT_OR_ALL + (CT_coerce_ID_to_ID_OPT (xlate_ident id)))) + | StringArgType -> + let s = CT_string (out_gen rawwit_string x) in + CT_coerce_STRING_OPT_to_VARG (CT_coerce_STRING_to_STRING_OPT s) + | PreIdentArgType -> + let id = CT_ident (out_gen rawwit_pre_ident x) in + CT_coerce_ID_OPT_OR_ALL_to_VARG + (CT_coerce_ID_OPT_to_ID_OPT_OR_ALL + (CT_coerce_ID_to_ID_OPT id)) + | IdentArgType -> + let id = xlate_ident (out_gen rawwit_ident x) in + CT_coerce_ID_OPT_OR_ALL_to_VARG + (CT_coerce_ID_OPT_to_ID_OPT_OR_ALL + (CT_coerce_ID_to_ID_OPT id)) + | QualidArgType -> + let id = tac_qualid_to_ct_ID (snd (out_gen rawwit_qualid x)) in + CT_coerce_ID_OPT_OR_ALL_to_VARG + (CT_coerce_ID_OPT_to_ID_OPT_OR_ALL + (CT_coerce_ID_to_ID_OPT id)) + (* Specific types *) + | ConstrArgType -> + CT_coerce_FORMULA_OPT_to_VARG + (CT_coerce_FORMULA_to_FORMULA_OPT (xlate_constr (out_gen rawwit_constr x))) + | ConstrMayEvalArgType -> xlate_error"TODO: generic constr-may-eval argument" + | QuantHypArgType ->xlate_error"TODO: generic quantified hypothesis argument" + | TacticArgType -> + let t = xlate_tactic (out_gen rawwit_tactic x) in + CT_coerce_TACTIC_OPT_to_VARG (CT_coerce_TACTIC_COM_to_TACTIC_OPT t) + | CastedOpenConstrArgType -> xlate_error "TODO: open constr" + | ConstrWithBindingsArgType -> xlate_error "TODO: constr with bindings" + | RedExprArgType -> xlate_error "TODO: red expr as generic argument" + | List0ArgType l -> xlate_error "TODO: lists of generic arguments" + | List1ArgType l -> xlate_error "TODO: non empty lists of generic arguments" + | OptArgType x -> xlate_error "TODO: optional generic arguments" + | PairArgType (u,v) -> xlate_error "TODO: pairs of generic arguments" + | ExtraArgType s -> xlate_error "Cannot treat extra generic arguments" + +(* let xlate_thm x = CT_thm (match x with | "THEOREM" -> "Theorem" | "REMARK" -> "Remark" | "LEMMA" -> "Lemma" | "FACT" -> "Fact" | _ -> xlate_error "xlate_thm");; +*) +let xlate_thm x = CT_thm (match x with + | Theorem -> "Theorem" + | Remark -> "Remark" + | Lemma -> "Lemma" + | Fact -> "Fact" + | Decl -> "Decl") +(* let xlate_defn x = CT_defn (match x with | "DEFINITION" -> "Definition" | "LOCAL" -> "Local" @@ -1448,12 +2204,26 @@ let xlate_defn x = CT_defn (match x with | "COERCION" -> "Coercion" | "LCOERCION" -> "LCOERCION" | _ -> xlate_error "xlate_defn");; +*) +let xlate_defn x = CT_defn (match x with + | LocalDefinition -> "Local" + | Definition -> "Definition") +(* let xlate_defn_or_thm s = try CT_coerce_THM_to_DEFN_OR_THM (xlate_thm s) with | _ -> CT_coerce_DEFN_to_DEFN_OR_THM (xlate_defn s);; +*) +let xlate_defn_or_thm = + function + (* Unable to decide if a fact in one section or at toplevel, a remark + at toplevel or a theorem or a Definition *) + | StartTheoremProof k -> CT_coerce_THM_to_DEFN_OR_THM (xlate_thm k) + | StartDefinitionBody k -> CT_coerce_DEFN_to_DEFN_OR_THM (xlate_defn k);; + +(* let xlate_var x = CT_var (match x with | "HYPOTHESES" -> "Hypothesis" | "HYPOTHESIS" -> "Hypothesis" @@ -1466,15 +2236,28 @@ let xlate_var x = CT_var (match x with "Axiom" as s -> s | "Parameter" as s -> s | _ -> xlate_error "xlate_var");; +*) +let xlate_var x = CT_var (match x with + | AssumptionParameter -> "Parameter" + | AssumptionAxiom -> "Axiom" + | AssumptionVariable -> "Variable" + | AssumptionHypothesis -> "Hypothesis");; +(* let xlate_dep = function | "DEP" -> CT_dep "Induction for" | "NODEP" -> CT_dep "Minimality for" | _ -> xlate_error "xlate_dep";; +*) +let xlate_dep = + function + | true -> CT_dep "Induction for" + | false -> CT_dep "Minimality for";; let xlate_locn = function +(* | Varg_int n -> CT_coerce_INT_to_INT_OR_LOCN n | Varg_string (CT_string "top") -> CT_coerce_LOCN_to_INT_OR_LOCN (CT_locn "top") @@ -1483,6 +2266,22 @@ let xlate_locn = | Varg_string (CT_string "next") -> CT_coerce_LOCN_to_INT_OR_LOCN (CT_locn "next") | _ -> xlate_error "xlate_locn";; +*) + | GoTo n -> CT_coerce_INT_to_INT_OR_LOCN (CT_int n) + | GoTop -> CT_coerce_LOCN_to_INT_OR_LOCN (CT_locn "top") + | GoPrev -> CT_coerce_LOCN_to_INT_OR_LOCN (CT_locn "prev") + | GoNext -> CT_coerce_LOCN_to_INT_OR_LOCN (CT_locn "next") + +let xlate_search_restr = + function + | SearchOutside [] -> CT_coerce_NONE_to_IN_OR_OUT_MODULES CT_none + | SearchInside (m1::l1) -> + CT_in_modules (CT_id_ne_list(loc_qualid_to_ct_ID m1, + List.map loc_qualid_to_ct_ID l1)) + | SearchOutside (m1::l1) -> + CT_out_modules (CT_id_ne_list(loc_qualid_to_ct_ID m1, + List.map loc_qualid_to_ct_ID l1)) + | SearchInside [] -> xlate_error "bad extra argument for Search" let xlate_check = function @@ -1491,18 +2290,12 @@ let xlate_check = | _ -> xlate_error "xlate_check";; let build_constructors l = - let strip_coerce = - function - | CT_coerce_ID_to_ID_OPT id -> id - | _ -> xlate_error "build_constructors" in - let rec rebind = - function - | [] -> [] - | (CT_binder ((CT_id_opt_ne_list (id, ids)), c)) :: l -> - (List.map (function id_opt -> CT_constr (strip_coerce id_opt, c)) - (id::ids)) @ rebind l in - CT_constr_list (rebind l);; + let f (id,coe,c) = + if coe then xlate_error "TODO: coercions in constructors" + else CT_constr (xlate_ident id, xlate_constr c) in + CT_constr_list (List.map f l) +(* let build_record_field_list l = let build_record_field = function @@ -1515,6 +2308,17 @@ let build_record_field_list l = CT_constr_coercion (id, coerce_iVARG_to_FORMULA c) | _ -> xlate_error "unexpected field in build_record_field_list" in CT_recconstr_list (List.map build_record_field l);; +*) +let build_record_field_list l = + let build_record_field (coe,d) = match d with + | AssumExpr (id,c) -> + if coe then CT_constr_coercion (xlate_ident id, xlate_constr c) + else + CT_coerce_CONSTR_to_RECCONSTR + (CT_constr (xlate_ident id, xlate_constr c)) + | DefExpr (id,c,topt) -> + xlate_error "TODO: manifest fields in Record" in + CT_recconstr_list (List.map build_record_field l);; let xlate_ast = let rec xlate_ast_aux = @@ -1540,6 +2344,7 @@ let xlate_ast = (CT_id_list (List.map (function s -> CT_ident s) sl)) in xlate_ast_aux;; +(* let get_require_flags impexp spec = let ct_impexp = match impexp with @@ -1553,18 +2358,38 @@ let get_require_flags impexp spec = | CT_string "IMPLEMENTATION" -> ctv_SPEC_OPT_NONE | CT_string s -> xlate_error ("unexpected Require specification flag " ^ s) in ct_impexp, ct_spec;; +*) + +let get_require_flags impexp spec = + let ct_impexp = + match impexp with + | false -> CT_import + | true -> CT_export in + let ct_spec = + match spec with + | None -> ctv_SPEC_OPT_NONE + | Some true -> CT_spec + | Some false -> ctv_SPEC_OPT_NONE in + ct_impexp, ct_spec;; let cvt_optional_eval_for_definition c1 optional_eval = match optional_eval with - None -> ct_coerce_FORMULA_to_DEF_BODY c1 - | Some (Targ_redexp red_com) -> + None -> ct_coerce_FORMULA_to_DEF_BODY (xlate_constr c1) + | Some red -> CT_coerce_EVAL_CMD_to_DEF_BODY( CT_eval(CT_coerce_NONE_to_INT_OPT CT_none, - red_com, - c1)) - | _ -> xlate_error - "bad extra argument (tactic?) for Definition";; + xlate_red_tactic red, + xlate_constr c1)) + +let cvt_vernac_binder = function + | (id,c) -> + CT_binder(CT_id_opt_ne_list (xlate_ident_opt (Some id),[]),xlate_constr c) +let cvt_vernac_binders args = + CT_binder_list(List.map cvt_vernac_binder args) + + +(* let rec cvt_varg = function | Node (_, "VERNACARGLIST", l) -> Varg_varglist (List.map cvt_varg l) @@ -1611,16 +2436,27 @@ let rec cvt_varg = (string_of_node_loc it)) | the_node -> failwith ("cvt_varg : " ^ (string_of_node_loc the_node)) and xlate_vernac = +*) +let xlate_vernac = function +(* | Node(_, "TACDEF", [Nvar(_,id); Node(_,"AST", [Node(_,"FUN", [Node(_,"FUNVAR", largs); tac])])]) -> - CT_tactic_definition(CT_ident id, - CT_id_list(List.map xlate_id largs), +*) + | VernacDeclareTacticDefinition (loc,[(_,id),TacFun (largs,tac)]) -> + let fst, rest = xlate_largs_to_id_unit largs in + let extract = function CT_unit -> xlate_error "TODO: void parameter" + | CT_coerce_ID_to_ID_UNIT x -> x in + let largs = List.map extract (fst::rest) in + CT_tactic_definition(xlate_ident id, + (* TODO, replace CT_id_list by CT_id_unit_list *) + CT_id_list largs, xlate_tactic tac) +(* | Node(_, "TACDEF", Nvar(_, id):: ((Node(_, "AST",[Node(_, "REC", [vc])])::tl) as the_list)) -> let x_rec_tacs = @@ -1639,11 +2475,37 @@ and xlate_vernac = let fst, others = match x_rec_tacs with fst::others -> fst, others | _ -> assert false in +*) + | VernacDeclareTacticDefinition + (loc,((id,TacFunRec (largs,tac))::_ as the_list)) -> + let x_rec_tacs = + List.map + (function + | ((_,x),TacFunRec ((_,fst),(argl,tac))) -> + let fst, rest = xlate_largs_to_id_unit ((Some fst)::argl) in + CT_rec_tactic_fun(xlate_ident x, + CT_id_unit_list(fst, rest), + xlate_tactic tac) + | ((_,x),tac) -> + CT_rec_tactic_fun(xlate_ident x, + (* Pas clair... *) + CT_id_unit_list (xlate_id_unit (Some x), []), + xlate_tactic tac)) the_list in + let fst, others = match x_rec_tacs with + fst::others -> fst, others + | _ -> assert false in CT_rec_tactic_definition(CT_rec_tactic_fun_list(fst, others)) +(* | Node(_, "TACDEF", [Nvar(_,id);Node(_,"AST",[tac])]) -> - CT_tactic_definition(CT_ident id, CT_id_list[], xlate_tactic tac) +*) + | VernacDeclareTacticDefinition (_,[(_,id),tac]) -> + CT_tactic_definition(xlate_ident id, CT_id_list[], xlate_tactic tac) + | VernacDeclareTacticDefinition (loc,_) -> xlate_error "Shouldn't occur" +(* | Node (_, s, l) -> (match s, List.map cvt_varg l with +*) +(* | "LoadFile", ((Varg_string verbose) :: ((Varg_string s) :: [])) -> CT_load ( (match verbose with @@ -1652,6 +2514,14 @@ and xlate_vernac = | CT_string s -> xlate_error ("expecting the keyword Verbose only :" ^ s)), CT_coerce_STRING_to_ID_OR_STRING s) +*) + | VernacLoad (verbose,s) -> + CT_load ( + (match verbose with + | false -> CT_coerce_NONE_to_VERBOSE_OPT CT_none + | true -> CT_verbose), + CT_coerce_STRING_to_ID_OR_STRING (CT_string s)) +(* | "Eval", ((Varg_tactic_arg (Targ_redexp tac)) :: ((Varg_constr f) :: tail)) -> let numopt = @@ -1659,23 +2529,44 @@ and xlate_vernac = | (Varg_int i) :: [] -> CT_coerce_INT_to_INT_OPT i | [] -> CT_coerce_NONE_to_INT_OPT CT_none | _ -> xlate_error "Eval expects an optional integer" in - CT_coerce_EVAL_CMD_to_COMMAND(CT_eval (numopt, tac, f)) +*) + | VernacCheckMayEval (Some red, numopt, f) -> + let red = xlate_red_tactic red in + CT_coerce_EVAL_CMD_to_COMMAND + (CT_eval (xlate_int_opt numopt, red, xlate_constr f)) +(* | "PWD", [] -> CT_pwd | "CD", ((Varg_string str) :: []) -> CT_cd (ctf_STRING_OPT_SOME str) | "CD", [] -> CT_cd ctf_STRING_OPT_NONE | "ADDPATH", ((Varg_string str) :: []) -> CT_addpath str | "RECADDPATH", ((Varg_string str) :: []) -> CT_recaddpath str | "DELPATH", ((Varg_string str) :: []) -> CT_delpath str - | "PrintPath", [] -> CT_print_loadpath | "QUIT", [] -> CT_quit - | (*ML commands *) - "AddMLPath", ((Varg_string str) :: []) -> CT_ml_add_path str +*) + | VernacChdir (Some str) -> CT_cd (ctf_STRING_OPT_SOME (CT_string str)) + | VernacChdir None -> CT_cd ctf_STRING_OPT_NONE + | VernacAddLoadPath (false,str,None) -> CT_addpath (CT_string str) + | VernacAddLoadPath (true,str,None) -> CT_recaddpath (CT_string str) + | VernacAddLoadPath (_,str,Some x) -> + xlate_error"TODO: Add (Rec) LoadPath as" + | VernacRemoveLoadPath str -> CT_delpath (CT_string str) + | VernacToplevelControl Quit -> CT_quit + | VernacToplevelControl _ -> xlate_error "TODO?: Drop/ProtectedToplevel" + (*ML commands *) +(* + | "AddMLPath", ((Varg_string str) :: []) -> CT_ml_add_path str | "RecAddMLPath", ((Varg_string str) :: []) -> CT_rec_ml_add_path str - | "PrintMLPath", [] -> CT_ml_print_path - | "PrintMLModules", [] -> CT_ml_print_modules - | "DeclareMLModule", (str :: l) -> CT_ml_declare_modules (CT_string_ne_list (strip_varg_string str, List.map strip_varg_string l)) + | "DeclareMLModule", (str :: l) -> +*) + | VernacAddMLPath (false,str) -> CT_ml_add_path (CT_string str) + | VernacAddMLPath (true,str) -> CT_rec_ml_add_path (CT_string str) + | VernacDeclareMLModule [] -> failwith "" + | VernacDeclareMLModule (str :: l) -> + CT_ml_declare_modules + (CT_string_ne_list (CT_string str, List.map (fun x -> CT_string x) l)) +(* | "GOAL", [] -> CT_proof_no_op | "GOAL", (c :: []) -> CT_coerce_THEOREM_GOAL_to_COMMAND (CT_goal (coerce_iVARG_to_FORMULA c)) | (*My attempt at getting all variants of Abort to use abort node *) @@ -1690,6 +2581,17 @@ and xlate_vernac = | "FOCUS", [] -> CT_focus (CT_coerce_NONE_to_INT_OPT CT_none) | "FOCUS", [Varg_int n] -> CT_focus (CT_coerce_INT_to_INT_OPT n) | "UNFOCUS", [] -> CT_unfocus +*) + | VernacStartProof (_, None, c, _, _) -> + CT_coerce_THEOREM_GOAL_to_COMMAND (CT_goal (xlate_constr c)) + | VernacAbort (Some id) -> CT_abort(ctf_ID_OPT_OR_ALL_SOME(xlate_ident id)) + | VernacAbort None -> CT_abort ctv_ID_OPT_OR_ALL_NONE + | VernacAbortAll -> CT_abort ctv_ID_OPT_OR_ALL_ALL + | VernacRestart -> CT_restart + | VernacSolve (n, tac) -> CT_solve (CT_int n, xlate_tactic tac) + | VernacFocus nopt -> CT_focus (xlate_int_opt nopt) + | VernacUnfocus -> CT_unfocus +(* | "HintRewrite", [orient; formula_list; Varg_ident base; Varg_tactic t] -> let ct_orient = match orient with | Varg_string (CT_string "LR") -> CT_lr @@ -1698,7 +2600,23 @@ and xlate_vernac = let f_ne_list = match formula_list with Varg_constrlist (fst::rest) -> CT_formula_ne_list(fst,rest) | _ -> assert false in - CT_hintrewrite(ct_orient, f_ne_list, base, t) +*) + | VernacExtend ("HintRewrite", orient :: formula_list :: base :: t) -> + let orient = out_gen rawwit_bool orient in + let formula_list = out_gen (wit_list1 (rawwit_constr)) formula_list in + let base = out_gen rawwit_pre_ident base in + let t = match t with + | [] -> TacId + | [t] -> out_gen rawwit_tactic t + | _ -> failwith "" in + let ct_orient = match orient with + | true -> CT_lr + | false -> CT_rl in + let f_ne_list = match List.map xlate_constr formula_list with + (fst::rest) -> CT_formula_ne_list(fst,rest) + | _ -> assert false in + CT_hintrewrite(ct_orient, f_ne_list, CT_ident base, xlate_tactic t) +(* | "HintResolve", ((Varg_ident id_name) :: ((Varg_varglist dbnames) :: ((Varg_constr c)::[]))) -> @@ -1731,6 +2649,24 @@ and xlate_vernac = Varg_tactic t] -> CT_hint(id_name, CT_id_list (List.map coerce_iVARG_to_ID dbnames), CT_extern(n, c, t)) +*) + | VernacHints (dbnames,h) -> + let dblist = CT_id_list(List.map (fun x -> CT_ident x) dbnames) in + (match h with + | HintsResolve [Some id_name, c] -> (* = Old HintResolve *) + CT_hint(xlate_ident id_name, dblist, CT_resolve (xlate_constr c)) + | HintsImmediate [Some id_name, c] -> (* = Old HintImmediate *) + CT_hint(xlate_ident id_name, dblist, CT_immediate(xlate_constr c)) + | HintsUnfold [Some id_name, q] -> (* = Old HintUnfold *) + CT_hint(xlate_ident id_name, dblist, + CT_unfold_hint (loc_qualid_to_ct_ID q)) + | HintsConstructors (id_name, q) -> + CT_hint(xlate_ident id_name, dblist, + CT_constructors (loc_qualid_to_ct_ID q)) + | HintsExtern (id_name, n, c, t) -> + CT_hint(xlate_ident id_name, dblist, + CT_extern(CT_int n, xlate_constr c, xlate_tactic t)) +(* | "HintsResolve", (Varg_varglist(dbnames)::(Varg_ident n1) :: names) -> CT_hints(CT_ident "Resolve", @@ -1746,7 +2682,44 @@ and xlate_vernac = CT_hints(CT_ident "Unfold", CT_id_ne_list(n1, List.map coerce_iVARG_to_ID names), CT_id_list(List.map coerce_iVARG_to_ID dbnames)) +*) + | HintsResolve l -> (* = Old HintsResolve *) + let l = List.map + (function + (None,Coqast.Node(_,"QUALID",l)) -> Astterm.interp_qualid l + | _ -> failwith "") l in + let n1, names = match List.map tac_qualid_to_ct_ID l with + n1 :: names -> n1, names + | _ -> failwith "" in + CT_hints(CT_ident "Resolve", + CT_id_ne_list(n1, names), + dblist) + | HintsImmediate l -> (* = Old HintsImmediate *) + let l = List.map + (function + (None,Coqast.Node(_,"QUALID",l)) -> Astterm.interp_qualid l + | _ -> failwith "") l in + let n1, names = match List.map tac_qualid_to_ct_ID l with + n1 :: names -> n1, names + | _ -> failwith "" in + CT_hints(CT_ident "Immediate", + CT_id_ne_list(n1, names), + dblist) + | HintsUnfold l -> (* = Old HintsUnfold *) + let l = List.map + (function + (None,qid) -> loc_qualid_to_ct_ID qid + | _ -> failwith "") l in + let n1, names = match l with + n1 :: names -> n1, names + | _ -> failwith "" in + CT_hints(CT_ident "Unfold", + CT_id_ne_list(n1, names), + dblist)) +(* Obsolete | "BY", ((Varg_tactic tcom) :: []) -> xlate_error "BY not implemented" +*) +(* | (*My attempt to get all variants of Save to use the same node *) "SaveNamed", [] -> CT_save (CT_coerce_THM_to_THM_OPT (CT_thm "Theorem"), ctv_ID_OPT_NONE) @@ -1764,33 +2737,101 @@ and xlate_vernac = CT_save (CT_coerce_THM_to_THM_OPT (CT_thm kind_string), ctf_ID_OPT_SOME s) | "SaveAnonymous", [Varg_ident s] -> CT_save (CT_coerce_THM_to_THM_OPT (CT_thm "Theorem"), ctf_ID_OPT_SOME s) +*) + | VernacEndProof (true,None) -> + CT_save (CT_coerce_THM_to_THM_OPT (CT_thm "Theorem"), ctv_ID_OPT_NONE) + | VernacEndProof (false,None) -> + CT_save (CT_coerce_THM_to_THM_OPT (CT_thm "Definition"), ctv_ID_OPT_NONE) + | VernacEndProof (b,Some (s, Some kind)) -> + CT_save (CT_coerce_THM_to_THM_OPT (xlate_thm kind), + ctf_ID_OPT_SOME (xlate_ident s)) + | VernacEndProof (b,Some (s,None)) -> + CT_save (CT_coerce_THM_to_THM_OPT (CT_thm "Theorem"), + ctf_ID_OPT_SOME (xlate_ident s)) +(* | "TRANSPARENT", (id :: idl) -> - CT_transparent(CT_id_ne_list(strip_varg_ident id, - List.map strip_varg_ident idl)) +*) + | VernacSetOpacity (false, id :: idl) -> + CT_transparent(CT_id_ne_list(loc_qualid_to_ct_ID id, + List.map loc_qualid_to_ct_ID idl)) +(* | "OPAQUE", (id :: idl) - -> CT_opaque (CT_id_ne_list(strip_varg_ident id, - List.map strip_varg_ident idl)) +*) + | VernacSetOpacity (true, id :: idl) + -> CT_opaque (CT_id_ne_list(loc_qualid_to_ct_ID id, + List.map loc_qualid_to_ct_ID idl)) + | VernacSetOpacity (_, []) -> xlate_error "Shouldn't occur" +(* No longer supported | "WriteModule", ((Varg_ident id) :: []) -> CT_write_module (id, CT_coerce_NONE_to_STRING_OPT CT_none) +*) +(* | "UNDO", ((Varg_int n) :: []) -> CT_undo (CT_coerce_INT_to_INT_OPT n) - | "SHOW", [] -> CT_show_goal (CT_coerce_NONE_to_INT_OPT CT_none) - | "SHOW", ((Varg_int n) :: []) -> CT_show_goal (CT_coerce_INT_to_INT_OPT n) +*) + | VernacUndo n -> CT_undo (CT_coerce_INT_to_INT_OPT (CT_int n)) +(* + | "SHOW", [] -> + | "SHOW", ((Varg_int n) :: []) -> +*) + | VernacShow (ShowGoal nopt) -> CT_show_goal (xlate_int_opt nopt) +(* | "ShowNode", [] -> CT_show_node | "ShowProof", [] -> CT_show_proof | "ShowTree", [] -> CT_show_tree | "ShowScript", [] -> CT_show_script | "ShowProofs", [] -> CT_show_proofs | "SHOWIMPL", [] -> CT_show_implicits +*) + | VernacShow ShowNode -> CT_show_node + | VernacShow ShowProof -> CT_show_proof + | VernacShow ShowTree -> CT_show_tree + | VernacShow ShowProofNames -> CT_show_proofs + | VernacShow (ShowIntros _|ShowGoalImplicitly _|ShowExistentials|ShowScript) + -> xlate_error "TODO: Show Intro/Intros/Implicits/Existentials/Script" +(* | "Go", (arg :: []) -> CT_go (xlate_locn arg) +*) + | VernacGo arg -> CT_go (xlate_locn arg) +(* | "ExplainProof", l -> - CT_explain_proof (CT_int_list (List.map strip_varg_int l)) +*) + | VernacShow ExplainProof l -> + CT_explain_proof (CT_int_list (List.map (fun x -> CT_int x) l)) +(* | "ExplainProofTree", l -> - CT_explain_prooftree (CT_int_list (List.map strip_varg_int l)) - | "PrintHint", [] -> CT_print_hint (CT_coerce_NONE_to_ID_OPT CT_none) - | "PrintHintDb", [Varg_ident id] -> CT_print_hintdb id +*) + | VernacShow ExplainTree l -> + CT_explain_prooftree (CT_int_list (List.map (fun x -> CT_int x) l)) +(* | "CheckGuard", [] -> CT_guarded - | "PrintHintId", ((Varg_ident id) :: []) -> - CT_print_hint (CT_coerce_ID_to_ID_OPT id) +*) + | VernacCheckGuard -> CT_guarded + | VernacPrint p -> + (match p with + PrintFullContext -> CT_print_all + | PrintName id -> CT_print_id (loc_qualid_to_ct_ID id) + | PrintOpaqueName id -> CT_print_opaqueid (loc_qualid_to_ct_ID id) + | PrintSectionContext id -> CT_print_section (loc_qualid_to_ct_ID id) + | PrintModules -> CT_print_modules + | PrintGrammar (phylum, name) -> CT_print_grammar CT_grammar_none + | PrintHintDb -> CT_print_hint (CT_coerce_NONE_to_ID_OPT CT_none) + | PrintHintDbName id -> CT_print_hintdb (CT_ident id) + | PrintHint id -> + CT_print_hint (CT_coerce_ID_to_ID_OPT (loc_qualid_to_ct_ID id)) + | PrintLoadPath -> CT_print_loadpath + | PrintMLLoadPath -> CT_ml_print_path + | PrintMLModules -> CT_ml_print_modules + | PrintGraph -> CT_print_graph + | PrintClasses -> CT_print_classes + | PrintCoercions -> CT_print_coercions + | PrintCoercionPaths (id1, id2) -> + CT_print_path (xlate_class id1, xlate_class id2) + | PrintInspect n -> CT_inspect (CT_int n) + | PrintUniverses _ -> xlate_error "TODO: Dump Universes" + | PrintHintGoal -> xlate_error "TODO: Print Hint" + | PrintLocalContext -> xlate_error "TODO: Print" + | PrintTables -> xlate_error "TODO: Print Tables") +(* | "PrintAll", [] -> CT_print_all | "PrintId", ((Varg_ident id) :: []) -> CT_print_id id | "PrintOpaqueId", ((Varg_ident id) :: []) -> CT_print_opaqueid id @@ -1799,22 +2840,51 @@ and xlate_vernac = | "PrintModules", [] -> CT_print_modules | "PrintGrammar", ((Varg_ident phylum) :: ((Varg_ident name) :: [])) -> CT_print_grammar CT_grammar_none + | "PrintHint", [] -> CT_print_hint (CT_coerce_NONE_to_ID_OPT CT_none) + | "PrintHintDb", [Varg_ident id] -> CT_print_hintdb id + | "PrintHintId", ((Varg_ident id) :: []) -> + | "PrintPath", [] -> CT_print_loadpath + | "PrintMLPath", [] -> CT_ml_print_path + | "PrintMLModules", [] -> CT_ml_print_modules + | "PrintGRAPH", [] -> CT_print_graph + | "PrintCLASSES", [] -> CT_print_classes + | "PrintCOERCIONS", [] -> CT_print_coercions + | "PrintPATH", ((Varg_ident id1) :: ((Varg_ident id2) :: [])) -> + CT_print_path (id1, id2) + | "INSPECT", ((Varg_int n) :: []) -> CT_inspect n +*) +(* No longer supported | "BeginModule", ((Varg_ident id) :: []) -> CT_module id +*) +(* | "BeginSection", ((Varg_ident id) :: []) -> - CT_coerce_SECTION_BEGIN_to_COMMAND (CT_section id) - | "EndSection", ((Varg_ident id) :: []) -> CT_section_end id +*) + | VernacBeginSection id -> + CT_coerce_SECTION_BEGIN_to_COMMAND (CT_section (xlate_ident id)) +(* + | "EndSection", ((Varg_ident id) :: []) -> +*) + | VernacEndSection id -> CT_section_end (xlate_ident id) +(* | "StartProof", ((Varg_string (CT_string kind)) :: ((Varg_ident s) :: (c :: []))) -> +*) + | VernacStartProof (kind, Some s, c, _, _) -> CT_coerce_THEOREM_GOAL_to_COMMAND( - CT_theorem_goal (xlate_defn_or_thm kind, s, coerce_iVARG_to_FORMULA c)) + CT_theorem_goal (xlate_defn_or_thm kind, xlate_ident s,xlate_constr c)) +(* | (*My attempt: suspend and resume as separate nodes *) "SUSPEND", [] -> CT_suspend | "RESUME", ((Varg_ident id) :: []) -> CT_resume (ctf_ID_OPT_SOME id) | "RESUME", [] -> CT_resume ctv_ID_OPT_NONE - | (*formerly | ("SUSPEND", []) -> suspend(CT_true) + (*formerly | ("SUSPEND", []) -> suspend(CT_true) | ("RESUME", []) -> suspend(CT_false) *) - "DEFINITION", +*) + | VernacSuspend -> CT_suspend + | VernacResume idopt -> CT_resume (xlate_ident_opt idopt) +(* + | "DEFINITION", (* reference : toplevel/vernacentries.ml *) (Varg_string (CT_string kind):: Varg_ident s :: Varg_constr c :: rest) -> let typ_opt, red_option = match rest with @@ -1827,17 +2897,29 @@ and xlate_vernac = CT_coerce_FORMULA_to_FORMULA_OPT (CT_coerce_SORT_TYPE_to_FORMULA b), None | _ -> assert false in - CT_definition - (xlate_defn kind, s, - cvt_optional_eval_for_definition c red_option, typ_opt) +*) + | VernacDefinition (kind,s,red_option,c,typ_opt,_) -> + CT_definition + (xlate_defn kind, xlate_ident s, + cvt_optional_eval_for_definition c red_option, + xlate_constr_opt typ_opt) +(* | "VARIABLE", ((Varg_string (CT_string kind)) :: ((Varg_binderlist b) :: [])) -> CT_variable (xlate_var kind, b) | "PARAMETER", ((Varg_string (CT_string kind)) :: ((Varg_binderlist b) :: [])) -> CT_variable (xlate_var kind, b) +*) + | VernacAssumption (kind, b) -> + CT_variable (xlate_var kind, cvt_vernac_binders b) +(* | "Check", ((Varg_string (CT_string kind)) :: (c :: [])) -> CT_check (coerce_iVARG_to_FORMULA c) +*) + | VernacCheckMayEval (None, numopt, c) -> + CT_check (xlate_constr c) +(* | "SearchPattern",Varg_constr c::l -> (match l with | [] -> CT_search_pattern(c, @@ -1864,7 +2946,15 @@ and xlate_vernac = | _ -> xlate_error "bad extra argument for Search") in CT_search(id, modifier) | _ -> xlate_error "bad argument list for Search") - | "INSPECT", ((Varg_int n) :: []) -> CT_inspect n +*) + | VernacSearch (s,x) -> + (match s with + | SearchPattern c -> + CT_search_pattern(xlate_constr c, xlate_search_restr x) + | SearchHead id -> + CT_search(loc_qualid_to_ct_ID id, xlate_search_restr x) + | SearchRewrite c -> xlate_error "TODO: SearchRewrite") +(* | (*Record from tactics/Record.v *) "RECORD", ((Varg_string coercion_or_not) :: ((Varg_ident s) :: @@ -1887,6 +2977,21 @@ and xlate_vernac = | _ -> assert false), record_constructor, build_record_field_list field_list) +*) + | (*Record from tactics/Record.v *) + VernacRecord + (add_coercion, s, binders, c1, rec_constructor_or_none, field_list) -> + let record_constructor = xlate_ident_opt rec_constructor_or_none in +(* match rec_constructor_or_none with + | None -> CT_coerce_NONE_to_ID_OPT CT_none + | Some id -> CT_coerce_ID_to_ID_OPT id in +*) CT_record + ((if add_coercion then CT_coercion_atm else + CT_coerce_NONE_to_COERCION_OPT(CT_none)), + xlate_ident s, cvt_vernac_binders binders, xlate_sort c1, record_constructor, + build_record_field_list field_list) + +(* TODO | (*Inversions from tactics/Inv.v *) "MakeSemiInversionLemmaFromHyp", ((Varg_int n) :: ((Varg_ident id1) :: ((Varg_ident id2) :: []))) -> @@ -1914,6 +3019,8 @@ and xlate_vernac = ((Varg_ident id) :: (c :: ((Varg_sorttype sort) :: []))) -> CT_derive_depinversion (CT_inv_clear, id, coerce_iVARG_to_FORMULA c, sort) +*) +(* | "ONEINDUCTIVE", ((Varg_string (CT_string f)) :: @@ -1953,6 +3060,17 @@ and xlate_vernac = | _ -> xlate_error "mutual inductive" in CT_mind_decl (CT_co_ind co_or_ind, CT_ind_spec_list (List.map strip_mutind lmi)) +*) + | VernacInductive (isind, lmi) -> + let co_or_ind = if isind then "Inductive" else "CoInductive" in + let strip_mutind (s, parameters, c, constructors) = + CT_ind_spec + (xlate_ident s, cvt_vernac_binders parameters, xlate_constr c, + build_constructors constructors) in + CT_mind_decl + (CT_co_ind co_or_ind, CT_ind_spec_list (List.map strip_mutind lmi)) + +(* | "MUTUALRECURSIVE", ((Varg_varglist (lm :: lmi)) :: []) -> let strip_mutrec = function @@ -1965,6 +3083,19 @@ and xlate_vernac = | _ -> xlate_error "mutual recursive" in CT_fix_decl (CT_fix_rec_list (strip_mutrec lm, List.map strip_mutrec lmi)) +*) + | VernacFixpoint [] -> xlate_error "mutual recursive" + | VernacFixpoint (lm :: lmi) -> + let strip_mutrec (fid, bl, arf, ardef) = + match cvt_vernac_binders bl with + | CT_binder_list (b :: bl) -> + CT_fix_rec (xlate_ident fid, CT_binder_ne_list (b, bl), + xlate_constr arf, xlate_constr ardef) + | _ -> xlate_error "mutual recursive" in + CT_fix_decl + (CT_fix_rec_list (strip_mutrec lm, List.map strip_mutrec lmi)) + +(* | "MUTUALCORECURSIVE", ((Varg_varglist (lm :: lmi)) :: []) -> let strip_mutcorec = function @@ -1974,6 +3105,15 @@ and xlate_vernac = | _ -> xlate_error "mutual corecursive" in CT_cofix_decl (CT_cofix_rec_list (strip_mutcorec lm, List.map strip_mutcorec lmi)) +*) + | VernacCoFixpoint [] -> xlate_error "mutual corecursive" + | VernacCoFixpoint (lm :: lmi) -> + let strip_mutcorec (fid, arf, ardef) = + CT_cofix_rec (xlate_ident fid, xlate_constr arf, xlate_constr ardef) in + CT_cofix_decl + (CT_cofix_rec_list (strip_mutcorec lm, List.map strip_mutcorec lmi)) + +(* | "INDUCTIONSCHEME", ((Varg_varglist (lm :: lmi)) :: []) -> let strip_ind = function @@ -1985,6 +3125,17 @@ and xlate_vernac = | _ -> xlate_error "induction scheme" in CT_ind_scheme (CT_scheme_spec_list (strip_ind lm, List.map strip_ind lmi)) +*) + | VernacScheme [] -> xlate_error "induction scheme" + | VernacScheme (lm :: lmi) -> + let strip_ind (id, depstr, inde, sort) = + CT_scheme_spec + (xlate_ident id, xlate_dep depstr, + CT_coerce_ID_to_FORMULA (loc_qualid_to_ct_ID inde), + xlate_sort sort) in + CT_ind_scheme + (CT_scheme_spec_list (strip_ind lm, List.map strip_ind lmi)) +(* | "SyntaxMacro", ((Varg_ident id) :: (c :: [])) -> CT_syntax_macro @@ -1992,31 +3143,53 @@ and xlate_vernac = | "SyntaxMacro", ((Varg_ident id) :: (c :: ((Varg_int n) :: []))) -> CT_syntax_macro (id, coerce_iVARG_to_FORMULA c, CT_coerce_INT_to_INT_OPT n) +*) + | VernacSyntacticDefinition (id, c, nopt) -> + CT_syntax_macro (xlate_ident id, xlate_constr c, xlate_int_opt nopt) + +(* Obsolete | "ABSTRACTION", ((Varg_ident id) :: (c :: l)) -> CT_abstraction (id, coerce_iVARG_to_FORMULA c, CT_int_list (List.map strip_varg_int l)) +*) +(* | "Require", ((Varg_string impexp) :: ((Varg_string spec) :: ((Varg_ident id) :: []))) -> +*) + | VernacRequire (None, spec, lid) -> xlate_error "TODO: Read Module" + | VernacRequire (Some impexp, spec, [id]) -> let ct_impexp, ct_spec = get_require_flags impexp spec in - CT_require (ct_impexp, ct_spec, id, CT_coerce_NONE_to_STRING_OPT CT_none) + CT_require (ct_impexp, ct_spec, loc_qualid_to_ct_ID id, + CT_coerce_NONE_to_STRING_OPT CT_none) + | VernacRequire (_,_,([]|_::_::_)) -> + xlate_error "TODO: general form of future Require" +(* | "RequireFrom", ((Varg_string impexp) :: ((Varg_string spec) :: ((Varg_ident id) :: ((Varg_string filename) :: [])))) -> +*) + | VernacRequireFrom (impexp, spec, id, filename) -> let ct_impexp, ct_spec = get_require_flags impexp spec in CT_require - (ct_impexp, ct_spec, id, CT_coerce_STRING_to_STRING_OPT filename) + (ct_impexp, ct_spec, xlate_ident id, + CT_coerce_STRING_to_STRING_OPT (CT_string filename)) +(* | "SYNTAX", ((Varg_ident phylum) :: ((Varg_ident s) :: (x :: (y :: l)))) -> - xlate_error "SYNTAX not implemented" - | (*Two versions of the syntax node with and without the binder list. *) +*) + | VernacSyntax (phylum, l) -> xlate_error "SYNTAX not implemented" + (*Two versions of the syntax node with and without the binder list. *) (*Need to update the metal file and ascent.mli first! | ("SYNTAX", [Varg_ident phy; Varg_ident s; spatarg; unparg; blist]) -> (syntaxop phy s spatarg unparg blist) | ("SYNTAX", [Varg_ident phy; Varg_ident s; spatarg; unparg]) -> (syntaxop phy s spatarg unparg coerce_ID_OPT_to_FORMULA_OPT(CT_coerce_NONE_to_ID_OPT(CT_none)))*) +(* Token is obsolete (automatically done by Grammar) and with no effects "TOKEN", ((Varg_string str) :: []) -> CT_token str +*) +(* | "INFIX", ((Varg_ast (CT_coerce_ID_OR_STRING_to_AST (CT_coerce_STRING_to_ID_OR_STRING @@ -2029,11 +3202,26 @@ and xlate_vernac = | "NONA" -> CT_nona | "NONE" -> CT_coerce_NONE_to_ASSOC CT_none | _ -> xlate_error "infix1"), n, str, id) +*) + | VernacInfix (str_assoc, n, str, id) -> + CT_infix ( + (match str_assoc with + | Some Gramext.LeftA -> CT_lefta + | Some Gramext.RightA -> CT_righta + | Some Gramext.NonA -> CT_nona + | None -> CT_coerce_NONE_to_ASSOC CT_none), + CT_int n, CT_string str, loc_qualid_to_ct_ID id) +(* | "GRAMMAR", (ge :: []) -> xlate_error "GRAMMAR not implemented" +*) + | VernacGrammar _ -> xlate_error "GRAMMAR not implemented" +(* Undo and Hyps Limit are now handled through the global options entries | "SETUNDO", ((Varg_int n) :: []) -> CT_setundo n | "UNSETUNDO", [] -> CT_unsetundo | "SETHYPSLIMIT", ((Varg_int n) :: []) -> CT_sethyp n | "UNSETHYPSLIMIT", [] -> CT_unsethyp +*) +(* | "COERCION", ((Varg_string (CT_string s)) :: ((Varg_string (CT_string str)) :: @@ -2049,12 +3237,32 @@ and xlate_vernac = | "" -> CT_coerce_NONE_to_LOCAL_OPT CT_none | _ -> xlate_error "unknown flag for a coercion2" in CT_coercion (local_opt, id_opt, id1, id2, id3) +*) + | VernacCoercion (s, id1, id2, id3) -> + let id_opt = CT_coerce_NONE_to_IDENTITY_OPT CT_none in + let local_opt = + match s with + (* Cannot decide whether it is a global or a Local but at toplevel *) + | Nametab.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none + | Nametab.DischargeAt _ -> CT_local + | Nametab.NotDeclare -> assert false in + CT_coercion (local_opt, id_opt, loc_qualid_to_ct_ID id1, + xlate_class id2, xlate_class id3) + + | VernacIdentityCoercion (s, id1, id2, id3) -> + let id_opt = CT_identity in + let local_opt = + match s with + (* Cannot decide whether it is a global or a Local but at toplevel *) + | Nametab.NeverDischarge -> CT_coerce_NONE_to_LOCAL_OPT CT_none + | Nametab.DischargeAt _ -> CT_local + | Nametab.NotDeclare -> assert false in + CT_coercion (local_opt, id_opt, xlate_ident id1, + xlate_class id2, xlate_class id3) +(* Not supported | "CLASS", (_ :: ((Varg_ident id1) :: [])) -> CT_class id1 - | "PrintGRAPH", [] -> CT_print_graph - | "PrintCLASSES", [] -> CT_print_classes - | "PrintCOERCIONS", [] -> CT_print_coercions - | "PrintPATH", ((Varg_ident id1) :: ((Varg_ident id2) :: [])) -> - CT_print_path (id1, id2) +*) +(* Natural entries are currently not supported | "SelectLanguageText", ((Varg_ident id) :: []) -> CT_set_natural id | "PrintText", ((Varg_ident id) :: []) -> CT_print_natural id | "AddTextParamOmit", ((Varg_ident id) :: []) -> @@ -2079,9 +3287,17 @@ and xlate_vernac = CT_remove_natural_feature (CT_nat_transparent, id) | "PrintTextParamImmediate", [] -> CT_print_natural_feature CT_nat_transparent +*) +(* | "ResetName", ((Varg_ident id) :: []) -> CT_reset id - | "ResetSection", ((Varg_ident id) :: []) -> CT_reset_section id | "ResetInitial", [] -> CT_restore_state (CT_ident "Initial") +*) + | VernacResetName id -> CT_reset (xlate_ident id) + | VernacResetInitial -> CT_restore_state (CT_ident "Initial") +(* No longer supported + | "ResetSection", ((Varg_ident id) :: []) -> CT_reset_section id +*) +(* Omega flags are handled through the global options entries | "OmegaFlag", ((Varg_string (CT_string s)) :: []) -> let fst_code = code (get s 0) in let @@ -2096,13 +3312,26 @@ and xlate_vernac = | _ -> CT_omega_flag (set_or_unset, CT_coerce_STRING_to_OMEGA_FEATURE (CT_string s))) - | s, l -> +*) + | VernacExtend (s, l) -> CT_user_vernac - (CT_ident s, CT_varg_list (List.map coerce_iVARG_to_VARG l))) - | _ -> xlate_error "xlate_vernac";; + (CT_ident s, CT_varg_list (List.map coerce_genarg_to_VARG l)) + | VernacDebug b -> xlate_error "TODO: Debug On/Off" + + | VernacList l -> xlate_error "Not treated here" + | (VernacLocate _|VernacGlobalCheck _|VernacPrintOption _| + VernacMemOption (_, _)|VernacRemoveOption (_, _)|VernacAddOption (_, _)| + VernacSetOption (_, _)|VernacUnsetOption _|VernacDeclareImplicits (_, _)| + VernacHintDestruct (_, _, _, _, _)|VernacBack _|VernacRestoreState _| + VernacWriteState _|VernacSolveExistential (_, _)|VernacCanonical _| + VernacImport (_, _)|VernacExactProof _|VernacDistfix (_, _, _, _)| + VernacTacticGrammar _|VernacVar _|VernacTime _|VernacNop|VernacComments _) + -> xlate_error "TODO: vernac" let xlate_vernac_list = function - | Node (_, "vernac_list", (v :: l)) -> - CT_command_list (xlate_vernac v, List.map xlate_vernac l) - | _ -> xlate_error "xlate_command_list";; + | VernacList (v::l) -> + CT_command_list + (xlate_vernac (snd v), List.map (fun (_,x) -> xlate_vernac x) l) + | VernacList [] -> xlate_error "xlate_command_list" + | _ -> xlate_error "Not a list of commands";; diff --git a/contrib/interface/xlate.mli b/contrib/interface/xlate.mli index b93635e47..0e9fd223f 100644 --- a/contrib/interface/xlate.mli +++ b/contrib/interface/xlate.mli @@ -1,12 +1,12 @@ open Ascent;; -val xlate_vernac : Ctast.t -> ct_COMMAND;; -val xlate_tactic : Ctast.t -> ct_TACTIC_COM;; +val xlate_vernac : Vernacexpr.vernac_expr -> ct_COMMAND;; +val xlate_tactic : Tacexpr.raw_tactic_expr -> ct_TACTIC_COM;; val xlate_formula : Ctast.t -> ct_FORMULA;; val xlate_int : Ctast.t -> ct_INT;; val xlate_string : Ctast.t -> ct_STRING;; -val xlate_id : Ctast.t -> ct_ID;; -val xlate_vernac_list : Ctast.t -> ct_COMMAND_LIST;; +val xlate_ident : Names.identifier -> ct_ID;; +val xlate_vernac_list : Vernacexpr.vernac_expr -> ct_COMMAND_LIST;; val g_nat_syntax_flag : bool ref;; val set_flags : (unit -> unit) ref;; diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index f424f79db..74e33cdfb 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -151,9 +151,4 @@ let import = import let env_of_safe_env e = e -(* Exported typing functions *) - -let typing env c = - let (j,cst) = safe_infer env c in - let _ = add_constraints cst env in - j +let typing = Typeops.typing diff --git a/kernel/term.ml b/kernel/term.ml index 62f21de5f..95968f6c8 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -353,6 +353,9 @@ let isRel c = match kind_of_term c with Rel _ -> true | _ -> false (* Tests if a variable *) let isVar c = match kind_of_term c with Var _ -> true | _ -> false +(* Tests if an inductive *) +let isInd c = match kind_of_term c with Ind _ -> true | _ -> false + (* Destructs the product (x:t1)t2 *) let destProd c = match kind_of_term c with | Prod (x,t1,t2) -> (x,t1,t2) diff --git a/kernel/term.mli b/kernel/term.mli index 73f29ddbe..643d06443 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -222,6 +222,7 @@ val kind_of_type : types -> (constr, types) kind_of_type val isRel : constr -> bool val isVar : constr -> bool +val isInd : constr -> bool val isEvar : constr -> bool val isMeta : constr -> bool val isSort : constr -> bool diff --git a/kernel/typeops.ml b/kernel/typeops.ml index 4606734fe..6da58adbc 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -477,3 +477,10 @@ let infer_local_decls env decls = push_rel d env, add_rel_decl d l, Constraint.union cst1 cst2 | [] -> env, empty_rel_context, Constraint.empty in inferec env decls + +(* Exported typing functions *) + +let typing env c = + let (j,cst) = infer env c in + let _ = add_constraints cst env in + j diff --git a/kernel/typeops.mli b/kernel/typeops.mli index 6980c9759..94d4e0844 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -89,3 +89,7 @@ val judge_of_case : env -> case_info (* Typecheck general fixpoint (not checking guard conditions) *) val type_fixpoint : env -> name array -> types array -> unsafe_judgment array -> constraints + +(* Kernel safe typing but applicable to partial proofs *) +val typing : env -> constr -> unsafe_judgment + diff --git a/kernel/univ.ml b/kernel/univ.ml index c5b998380..9242e9d8a 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -112,8 +112,8 @@ let declare_univ u g = else g -(* The universes of Prop and Set: Type_0, Type_1 and the - resulting graph. *) +(* When typing Prop and Set, there is no constraint on the level, + hence the definition of prop_univ *) let initial_universes = UniverseMap.empty let prop_univ = Max ([],[]) @@ -406,17 +406,20 @@ let num_edges g = UniverseMap.fold (fun _ a n -> n + (reln_len a)) g 0 let pr_arc = function - | Canonical {univ=u; gt=gt; ge=ge} -> - hov 2 - (pr_uni_level u ++ spc () ++ - prlist_with_sep pr_spc (fun v -> str ">" ++ pr_uni_level v) gt ++ - prlist_with_sep pr_spc (fun v -> str ">=" ++ pr_uni_level v) ge) + | Canonical {univ=u; gt=[]; ge=[]} -> + mt () + | Canonical {univ=u; gt=gt; ge=ge} -> + pr_uni_level u ++ str " " ++ + v 0 + (prlist_with_sep pr_spc (fun v -> str "> " ++ pr_uni_level v) gt ++ + prlist_with_sep pr_spc (fun v -> str ">= " ++ pr_uni_level v) ge) ++ + fnl () | Equiv (u,v) -> - pr_uni_level u ++ str "=" ++ pr_uni_level v + pr_uni_level u ++ str " = " ++ pr_uni_level v ++ fnl () let pr_universes g = let graph = UniverseMap.fold (fun k a l -> (k,a)::l) g [] in - prlist_with_sep pr_fnl (function (_,a) -> pr_arc a) graph + prlist (function (_,a) -> pr_arc a) graph (* Dumping constrains to a file *) diff --git a/lib/options.ml b/lib/options.ml index f1e40a89e..b60210fde 100644 --- a/lib/options.ml +++ b/lib/options.ml @@ -45,8 +45,7 @@ let if_verbose f x = if not !silent then f x (* The number of printed hypothesis in a goal *) let print_hyps_limit = ref (None : int option) -let set_print_hyps_limit n = print_hyps_limit := Some n -let unset_print_hyps_limit () = print_hyps_limit := None +let set_print_hyps_limit n = print_hyps_limit := n let print_hyps_limit () = !print_hyps_limit (* A list of the areas of the system where "unsafe" operation diff --git a/lib/options.mli b/lib/options.mli index efc8617de..53e796f9a 100644 --- a/lib/options.mli +++ b/lib/options.mli @@ -28,8 +28,8 @@ val silently : ('a -> 'b) -> 'a -> 'b val if_silent : ('a -> unit) -> 'a -> unit val if_verbose : ('a -> unit) -> 'a -> unit -val set_print_hyps_limit : int -> unit -val unset_print_hyps_limit : unit -> unit +(* If [None], no limit *) +val set_print_hyps_limit : int option -> unit val print_hyps_limit : unit -> int option val add_unsafe : string -> unit diff --git a/lib/pp_control.ml b/lib/pp_control.ml index 2e8476790..78b78020d 100644 --- a/lib/pp_control.ml +++ b/lib/pp_control.ml @@ -95,6 +95,8 @@ let deep_ft = with_output_to stdout let _ = set_gp deep_ft deep_gp (* For parametrization through vernacular *) -let set_depth_boxes v = Format.pp_set_max_boxes std_ft v -let get_depth_boxes () = Format.pp_get_max_boxes std_ft () +let default = Format.pp_get_max_boxes std_ft () +let get_depth_boxes () = Some (Format.pp_get_max_boxes std_ft ()) +let set_depth_boxes v = + Format.pp_set_max_boxes std_ft (match v with None -> default | Some v -> v) diff --git a/lib/pp_control.mli b/lib/pp_control.mli index 2551726df..548b98a7f 100644 --- a/lib/pp_control.mli +++ b/lib/pp_control.mli @@ -42,5 +42,5 @@ val deep_ft : Format.formatter (*s For parametrization through vernacular. *) -val set_depth_boxes : int -> unit -val get_depth_boxes : unit -> int +val set_depth_boxes : int option -> unit +val get_depth_boxes : unit -> int option diff --git a/lib/util.ml b/lib/util.ml index cd575bf08..3354bba53 100644 --- a/lib/util.ml +++ b/lib/util.ml @@ -22,6 +22,8 @@ let errorlabstrm l pps = raise (UserError(l,pps)) (* raising located exceptions *) type loc = int * int +type 'a located = loc * 'a +let dummy_loc = (0,0) let anomaly_loc (loc,s,strm) = Stdpp.raise_with_loc loc (Anomaly (s,strm)) let user_err_loc (loc,s,strm) = Stdpp.raise_with_loc loc (UserError (s,strm)) let invalid_arg_loc (loc,s) = Stdpp.raise_with_loc loc (Invalid_argument s) diff --git a/lib/util.mli b/lib/util.mli index 6bbe609cb..e9cd97a7c 100644 --- a/lib/util.mli +++ b/lib/util.mli @@ -27,7 +27,9 @@ val error : string -> 'a val errorlabstrm : string -> std_ppcmds -> 'a type loc = int * int +type 'a located = loc * 'a +val dummy_loc : loc val anomaly_loc : loc * string * std_ppcmds -> 'a val user_err_loc : loc * string * std_ppcmds -> 'a val invalid_arg_loc : loc * string -> 'a diff --git a/library/declare.ml b/library/declare.ml index 6b2f9a6ae..fc80cfda9 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -35,10 +35,7 @@ open Safe_typing (useful only for persistent constructions), is the length of the section part in [dir] *) -type strength = - | NotDeclare - | DischargeAt of dir_path * int - | NeverDischarge +open Nametab let depth_of_strength = function | DischargeAt (sp',n) -> n diff --git a/library/declare.mli b/library/declare.mli index b7b305cfa..07b9d98b6 100644 --- a/library/declare.mli +++ b/library/declare.mli @@ -26,13 +26,7 @@ open Nametab reset works properly --- and will fill some global tables such as [Nametab] and [Impargs]. *) -type strength = - | NotDeclare - | DischargeAt of dir_path * int - | NeverDischarge - -val is_less_persistent_strength : strength * strength -> bool -val strength_min : strength * strength -> strength +open Nametab type section_variable_entry = | SectionLocalDef of constr * types option @@ -51,6 +45,10 @@ val declare_constant : identifier -> constant_declaration -> constant val redeclare_constant : constant -> Cooking.recipe * strength -> unit +(* +val declare_parameter : identifier -> constr -> constant +*) + (* [declare_mind me] declares a block of inductive types with their constructors in the current section; it returns the path of the whole block *) @@ -61,6 +59,8 @@ val out_inductive : Libobject.obj -> mutual_inductive_entry val make_strength_0 : unit -> strength val make_strength_1 : unit -> strength val make_strength_2 : unit -> strength +val is_less_persistent_strength : strength * strength -> bool +val strength_min : strength * strength -> strength (*s Corresponding test and access functions. *) diff --git a/library/goptions.ml b/library/goptions.ml index f919b12e8..b4056472b 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -29,11 +29,11 @@ let nickname = function | SecondaryTable (s1,s2) -> s1^" "^s2 let error_undeclared_key key = - error ((nickname key)^": no such table or option") + error ((nickname key)^": no table or option of this type") type value = | BoolValue of bool - | IntValue of int + | IntValue of int option | StringValue of string | IdentValue of global_reference @@ -55,11 +55,10 @@ module MakeTable = type key val table : (string * key table_of_A) list ref val encode : key -> t - val check : t -> unit val printer : t -> std_ppcmds val key : option_name val title : string - val member_message : key -> bool -> string + val member_message : t -> bool -> std_ppcmds val synchronous : bool end) -> struct @@ -113,18 +112,17 @@ module MakeTable = (hov 0 (if MySet.is_empty table then str "None" ++ fnl () else MySet.fold - (fun a b -> printer a ++ spc () ++ b) table (mt ())))) + (fun a b -> printer a ++ spc () ++ b) + table (mt ()) ++ fnl ()))) class table_of_A () = object - method add x = - let c = A.encode x in - A.check c; - add_option c + method add x = add_option (A.encode x) method remove x = remove_option (A.encode x) method mem x = - let answer = MySet.mem (A.encode x) !t in - msg (str (A.member_message x answer)) + let y = A.encode x in + let answer = MySet.mem y !t in + msg (A.member_message y answer ++ fnl ()) method print = print_table A.title A.printer !t end @@ -139,10 +137,9 @@ let get_string_table k = List.assoc (nickname k) !string_table module type StringConvertArg = sig - val check : string -> unit val key : option_name val title : string - val member_message : string -> bool -> string + val member_message : string -> bool -> std_ppcmds val synchronous : bool end @@ -152,8 +149,7 @@ struct type key = string let table = string_table let encode x = x - let check = A.check - let printer s = (str s) + let printer = str let key = A.key let title = A.title let member_message = A.member_message @@ -163,29 +159,27 @@ end module MakeStringTable = functor (A : StringConvertArg) -> MakeTable (StringConvert(A)) -let ident_table = ref [] +let ref_table = ref [] -let get_ident_table k = List.assoc (nickname k) !ident_table +let get_ref_table k = List.assoc (nickname k) !ref_table -module type IdentConvertArg = +module type RefConvertArg = sig type t - val encode : global_reference -> t - val check : t -> unit + val encode : qualid located -> t val printer : t -> std_ppcmds val key : option_name val title : string - val member_message : global_reference -> bool -> string + val member_message : t -> bool -> std_ppcmds val synchronous : bool end -module IdentConvert = functor (A : IdentConvertArg) -> +module RefConvert = functor (A : RefConvertArg) -> struct type t = A.t - type key = global_reference - let table = ident_table + type key = qualid located + let table = ref_table let encode = A.encode - let check = A.check let printer = A.printer let key = A.key let title = A.title @@ -193,8 +187,8 @@ struct let synchronous = A.synchronous end -module MakeIdentTable = - functor (A : IdentConvertArg) -> MakeTable (IdentConvert(A)) +module MakeRefTable = + functor (A : RefConvertArg) -> MakeTable (RefConvert(A)) (****************************************************************************) (* 2- Options *) @@ -222,7 +216,7 @@ let check_key key = try error "Sorry, this option name is already used" with Not_found -> if List.mem_assoc (nickname key) !string_table - or List.mem_assoc (nickname key) !ident_table + or List.mem_assoc (nickname key) !ref_table then error "Sorry, this option name is already used" open Summary @@ -301,7 +295,8 @@ let msg_option_value (name,v) = match v with | BoolValue true -> str "true" | BoolValue false -> str "false" - | IntValue n -> int n + | IntValue (Some n) -> int n + | IntValue None -> str "undefined" | StringValue s -> str s | IdentValue r -> pr_global_env (Global.env()) r @@ -343,7 +338,7 @@ let print_tables () = !string_table (mt ()) ++ List.fold_right (fun (nickkey,_) p -> p ++ str (" "^nickkey) ++ fnl ()) - !ident_table (mt ()) ++ + !ref_table (mt ()) ++ fnl () ) diff --git a/library/goptions.mli b/library/goptions.mli index 8f810a266..d43a4283e 100644 --- a/library/goptions.mli +++ b/library/goptions.mli @@ -54,6 +54,7 @@ (*i*) open Pp +open Util open Names open Term open Nametab @@ -68,26 +69,22 @@ type option_name = val error_undeclared_key : option_name -> 'a - (*s Tables. *) (* The functor [MakeStringTable] declares a table containing objects - of type [string]; the function [check] is for testing if a given - object is allowed to be added to the table; the function - [member_message] say what to print when invoking the "Test Toto - Titi foo." command; at the end [title] is the table name printed - when invoking the "Print Toto Titi." command; [active] is roughly - the internal version of the vernacular "Test ...": it tells if a - given object is in the table; [elements] returns the list of - elements of the table *) + of type [string]; the function [member_message] say what to print + when invoking the "Test Toto Titi foo." command; at the end [title] + is the table name printed when invoking the "Print Toto Titi." + command; [active] is roughly the internal version of the vernacular + "Test ...": it tells if a given object is in the table; [elements] + returns the list of elements of the table *) module MakeStringTable : functor (A : sig - val check : string -> unit val key : option_name val title : string - val member_message : string -> bool -> string + val member_message : string -> bool -> std_ppcmds val synchronous : bool end) -> sig @@ -95,29 +92,25 @@ sig val elements : unit -> string list end -(* The functor [MakeIdentTable] declares a new table of [global_reference]; - for generality, identifiers may be internally encode in other type - which is [A.t] through an encoding function [encode : global_reference -> - A.t] (typically, [A.t] is the persistent [global_reference] associated - to the currentdenotation of the [global_reference] and the encoding function - is the globalization function); the function [check] is for testing - if a given object is allowed to be added to the table; the function +(* The functor [MakeRefTable] declares a new table of objects of type + [A.t] practically denoted by [qualid]; the encoding function + [encode : qualid -> A.t] is typically a globalization function, + possibly with some restriction checks; the function [member_message] say what to print when invoking the "Test Toto Titi foo." command; at the end [title] is the table name printed when invoking the "Print Toto Titi." command; [active] is roughly the internal version of the vernacular "Test ...": it tells if a - given object is in the table. *) + given object is in the table. *) -module MakeIdentTable : +module MakeRefTable : functor (A : sig type t - val encode : global_reference -> t - val check : t -> unit + val encode : qualid located -> t val printer : t -> std_ppcmds val key : option_name val title : string - val member_message : global_reference -> bool -> string + val member_message : t -> bool -> std_ppcmds val synchronous : bool end) -> sig @@ -145,7 +138,7 @@ type 'a option_sig = { type 'a write_function = 'a -> unit -val declare_int_option : int option_sig -> int write_function +val declare_int_option : int option option_sig -> int option write_function val declare_bool_option : bool option_sig -> bool write_function val declare_string_option: string option_sig -> string write_function @@ -159,14 +152,14 @@ val get_string_table : mem : string -> unit; print : unit > -val get_ident_table : +val get_ref_table : option_name -> - < add : global_reference -> unit; - remove : global_reference -> unit; - mem : global_reference -> unit; + < add : qualid located -> unit; + remove : qualid located -> unit; + mem : qualid located -> unit; print : unit > -val set_int_option_value : option_name -> int -> unit +val set_int_option_value : option_name -> int option -> unit val set_bool_option_value : option_name -> bool -> unit val set_string_option_value : option_name -> string -> unit diff --git a/library/impargs.ml b/library/impargs.ml index 86acf1687..af6bfd6b7 100644 --- a/library/impargs.ml +++ b/library/impargs.ml @@ -77,7 +77,7 @@ let with_implicits b f x = raise e end -let implicitely f = with_implicits true f +let implicitly f = with_implicits true f let using_implicits = function | No_impl -> with_implicits false diff --git a/library/impargs.mli b/library/impargs.mli index 470f3d0c3..022a38230 100644 --- a/library/impargs.mli +++ b/library/impargs.mli @@ -20,7 +20,7 @@ open Nametab val make_implicit_args : bool -> unit val is_implicit_args : unit -> bool -val implicitely : ('a -> 'b) -> 'a -> 'b +val implicitly : ('a -> 'b) -> 'a -> 'b val with_implicits : bool -> ('a -> 'b) -> 'a -> 'b (*s An [implicits_list] is a list of positions telling which arguments diff --git a/library/library.ml b/library/library.ml index 47beca5f3..bc25eb617 100644 --- a/library/library.ml +++ b/library/library.ml @@ -435,7 +435,7 @@ let locate_qualified_library qid = (LibInPath, extend_dirpath (find_logical_path path) base, file) with Not_found -> raise LibNotFound -let rec_intern_qualified_library qid = +let rec_intern_qualified_library (loc,qid) = try let (_,dir,f) = locate_qualified_library qid in rec_intern_module (dir,f); @@ -443,19 +443,18 @@ let rec_intern_qualified_library qid = with | LibUnmappedDir -> let prefix, id = repr_qualid qid in - errorlabstrm "rec_intern_qualified_library" - (str ("Cannot load "^(string_of_id id)^":") ++ spc () ++ - str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ fnl ()) + user_err_loc (loc, "rec_intern_qualified_library", + str ("Cannot load "^(string_of_id id)^":") ++ spc () ++ + str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ + fnl ()) | LibNotFound -> - errorlabstrm "rec_intern_qualified_library" - (str"Cannot find module " ++ pr_qualid qid ++ str" in loadpath") + user_err_loc (loc, "rec_intern_qualified_library", + str"Cannot find module " ++ pr_qualid qid ++ str" in loadpath") -let rec_intern_module_from_file qid f = +let rec_intern_module_from_file idopt f = (* A name is specified, we have to check it contains module id *) - let prefix, id = repr_qualid qid in - assert (repr_dirpath prefix = []); let _, f = System.find_file_in_path (get_load_path ()) (f^".vo") in - rec_intern_by_filename_only (Some id) f + rec_intern_by_filename_only idopt f (*s [require_module] loads and opens a module. This is a synchronized operation. *) @@ -493,17 +492,19 @@ let require_module spec qidl export = add_anonymous_leaf (in_require (modrefl,Some export)); add_frozen_state () -let require_module_from_file spec qid file export = - let modref = rec_intern_module_from_file qid file in +let require_module_from_file spec idopt file export = + let modref = rec_intern_module_from_file idopt file in add_anonymous_leaf (in_require ([modref],Some export)); add_frozen_state () -let import_module export qid = +let import_module export (loc,qid) = let modref = try Nametab.locate_loaded_library qid with Not_found -> - error ((Nametab.string_of_qualid qid)^" not loaded") in + user_err_loc + (loc,"import_module", + str ((Nametab.string_of_qualid qid)^" not loaded")) in add_anonymous_leaf (in_require ([modref],Some export)) let read_module qid = diff --git a/library/library.mli b/library/library.mli index 5b51780e1..c29299331 100644 --- a/library/library.mli +++ b/library/library.mli @@ -9,6 +9,7 @@ (*i $Id$ i*) (*i*) +open Util open Names open Libobject (*i*) @@ -20,9 +21,9 @@ open Libobject provides a high level function [require] which corresponds to the vernacular command [Require]. *) -val read_module : Nametab.qualid -> unit +val read_module : Nametab.qualid located -> unit val read_module_from_file : System.physical_path -> unit -val import_module : bool -> Nametab.qualid -> unit +val import_module : bool -> Nametab.qualid located -> unit val module_is_loaded : dir_path -> bool val module_is_opened : dir_path -> bool @@ -39,10 +40,10 @@ val fmt_modules_state : unit -> Pp.std_ppcmds exported. *) val require_module : - bool option -> Nametab.qualid list -> bool -> unit + bool option -> Nametab.qualid located list -> bool -> unit val require_module_from_file : - bool option -> Nametab.qualid -> string -> bool -> unit + bool option -> identifier option -> string -> bool -> unit (*s [save_module_to s f] saves the current environment as a module [s] in the file [f]. *) diff --git a/library/nametab.ml b/library/nametab.ml index 99524bde1..b28506f91 100755 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -139,6 +139,8 @@ let push_idtree extract_dirpath tab n dir id o = let push_long_names_ccipath = push_idtree dirpath_of_extended_ref the_ccitab let push_short_name_ccipath = push_idtree dirpath_of_extended_ref the_ccitab +let push_long_names_objpath = + push_idtree (fun sp -> let (p,_) = repr_path sp in p) the_objtab let push_short_name_objpath = push_idtree (fun sp -> let (p,_) = repr_path sp in p) the_objtab @@ -181,6 +183,12 @@ let push_short_name_object sp = let _, s = repr_qualid (qualid_of_sp sp) in push_short_name_objpath 0 empty_dirpath s sp +(* This is for tactic definition names *) + +let push_tactic_path sp = + let dir, s = repr_qualid (qualid_of_sp sp) in + push_long_names_objpath 0 dir s sp + (* This is to remember absolute Section/Module names and to avoid redundancy *) let push_section fulldir = let dir, s = split_dirpath fulldir in @@ -213,6 +221,10 @@ let locate_obj qid = let (dir,id) = repr_qualid qid in locate_in_tree (Idmap.find id !the_objtab) dir +let locate_tactic qid = + let (dir,id) = repr_qualid qid in + locate_in_tree (Idmap.find id !the_objtab) dir + (* Actually, this table has only two levels, since only basename and *) (* fullname are registered *) let push_loaded_library fulldir = @@ -259,13 +271,13 @@ let absolute_reference sp = let locate_in_absolute_module dir id = absolute_reference (make_path dir id) -let global loc qid = +let global (loc,qid) = try match extended_locate qid with | TrueGlobal ref -> ref | SyntacticDef _ -> - error - ("Unexpected reference to a syntactic definition: " - ^(string_of_qualid qid)) + user_err_loc (loc,"global", + str "Unexpected reference to a syntactic definition: " ++ + pr_qualid qid) with Not_found -> error_global_not_found_loc loc qid @@ -299,6 +311,13 @@ let pr_global_env env ref = let s = string_of_qualid (shortest_qualid_of_global env ref) in (str s) +let global_inductive (loc,qid as locqid) = + match global locqid with + | IndRef ind -> ind + | ref -> + user_err_loc (loc,"global_inductive", + pr_qualid qid ++ spc () ++ str "is not an inductive type") + (********************************************************************) (********************************************************************) @@ -330,3 +349,8 @@ let _ = Summary.unfreeze_function = unfreeze; Summary.init_function = init; Summary.survive_section = false } + +type strength = + | NotDeclare + | DischargeAt of dir_path * int + | NeverDischarge diff --git a/library/nametab.mli b/library/nametab.mli index adb764fcb..d71e3e379 100755 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -67,6 +67,10 @@ val push_syntactic_definition : section_path -> unit val push_short_name_syntactic_definition : section_path -> unit val push_short_name_object : section_path -> unit +(*s Register visibility of absolute paths by short names *) +val push_tactic_path : section_path -> unit +val locate_tactic : qualid -> section_path + (*s Register visibility by all qualifications *) val push_section : dir_path -> unit @@ -82,7 +86,10 @@ val locate : qualid -> global_reference (* This function is used to transform a qualified identifier into a global reference, with a nice error message in case of failure *) -val global : loc -> qualid -> global_reference +val global : qualid located -> global_reference + +(* The same for inductive types *) +val global_inductive : qualid located -> inductive (* This locates also syntactic definitions *) val extended_locate : qualid -> extended_global_reference @@ -109,3 +116,8 @@ val locate_in_absolute_module : dir_path -> identifier -> global_reference val push_loaded_library : dir_path -> unit val locate_loaded_library : qualid -> dir_path + +type strength = + | NotDeclare + | DischargeAt of dir_path * int + | NeverDischarge diff --git a/parsing/ast.ml b/parsing/ast.ml index ac5c149e4..1a4be70d9 100755 --- a/parsing/ast.ml +++ b/parsing/ast.ml @@ -12,7 +12,8 @@ open Pp open Util open Names open Coqast -open Pcoq +open Tacexpr +open Genarg let isMeta s = String.length s <> 0 & s.[0]='$' @@ -30,6 +31,36 @@ let loc = function | Path (loc,_) -> loc | Dynamic (loc,_) -> loc +type entry_type = + | PureAstType + | IntAstType + | IdentAstType + | AstListType + | TacticAtomAstType + | ThmTokenAstType + | DynamicAstType + | ReferenceAstType + | GenAstType of Genarg.argument_type + +(* patterns of ast *) +type astpat = + | Pquote of t + | Pmeta of string * tok_kind + | Pnode of string * patlist + | Pslam of identifier option * astpat + | Pmeta_slam of string * astpat + +and patlist = + | Pcons of astpat * patlist + | Plmeta of string + | Pnil + +and tok_kind = Tnum | Tid | Tstr | Tpath | Tvar | Tany | Tlist + +type pat = + | AstListPat of patlist + | PureAstPat of astpat + (* building a node with dummy location *) let ope(op,l) = Node(dummy_loc,op,l) @@ -66,38 +97,42 @@ let nvar_of_ast = function | Nvar(_,s) -> s | ast -> invalid_arg_loc (loc ast, "Ast.nvar_of_ast") +let meta_of_ast = function + | Nmeta(_,s) -> s + | ast -> invalid_arg_loc (loc ast, "Ast.nvar_of_ast") + let id_of_ast = function | Id(_,s) -> s | ast -> invalid_arg_loc (loc ast, "Ast.nvar_of_ast") -(* patterns of ast *) -type pat = - | Pquote of t - | Pmeta of string * tok_kind - | Pnode of string * patlist - | Pslam of identifier option * pat - | Pmeta_slam of string * pat - -and patlist = - | Pcons of pat * patlist - | Plmeta of string - | Pnil - -and tok_kind = Tnum | Tid | Tstr | Tpath | Tvar | Tany | Tlist - (* semantic actions of grammar rules *) type act = - | Aast of pat + | Act of pat + | ActCase of act * (pat * act) list + | ActCaseList of act * (pat * act) list + +(* +type act = + | Aast of typed_ast | Aastlist of patlist - | Acase of act * (pat * act) list + | Acase of act * (Coqast.t * act) list + | Atypedcase of act * (typed_ast * act) list | Acaselist of act * (patlist * act) list +*) (* values associated to variables *) -type v = - | Vast of t - | Vastlist of t list +type typed_ast = + | AstListNode of Coqast.t list + | PureAstNode of Coqast.t -type env = (string * v) list +type ast_action_type = ETast | ETastl + +type grammar_action = + | SimpleAction of loc * typed_ast + | CaseAction of + loc * grammar_action * ast_action_type * (t list * grammar_action) list + +type env = (string * typed_ast) list (* Pretty-printing *) let rec print_ast ast = @@ -147,18 +182,17 @@ let rec print_astpat = function (str"[" ++ str(string_of_id id) ++ str"]" ++ cut () ++ print_astpat b) and print_astlpat = function - | Pnil -> (mt ()) + | Pnil -> mt () | Pcons(h,Pnil) -> hov 1 (print_astpat h) | Pcons(h,t) -> hov 1 (print_astpat h ++ spc () ++ print_astlpat t) - | Plmeta(s) -> (str"| " ++ str s) + | Plmeta(s) -> str"| " ++ str s let print_val = function - | Vast a -> print_ast a - | Vastlist al -> + | PureAstNode a -> print_ast a + | AstListNode al -> hov 1 (str"[" ++ prlist_with_sep pr_spc print_ast al ++ str"]") - (* Ast values environments *) let grammar_type_error (loc,s) = @@ -182,15 +216,55 @@ let rec coerce_to_var = function | Nvar(_,id) as var -> var | Nmeta(_,id) as var -> var | Node(_,"QUALID",[Nvar(_,id) as var]) -> var - | Node(_,"QUALIDARG",[Nvar(_,id) as var]) -> var | ast -> user_err_loc (loc ast,"Ast.coerce_to_var", (str"This expression should be a simple identifier")) let coerce_to_id a = match coerce_to_var a with | Nvar (_,id) -> id -(* | Nmeta(_,id) -> id_of_string id*) - | ast -> invalid_arg "coerce_to_id" + | ast -> user_err_loc + (loc ast,"Ast.coerce_to_id", + str"This expression should be a simple identifier") + +let coerce_qualid_to_id (loc,qid) = match Nametab.repr_qualid qid with + | dir, id when dir = Nameops.empty_dirpath -> id + | _ -> + user_err_loc (loc, "Ast.coerce_qualid_to_id", + str"This expression should be a simple identifier") + +let coerce_reference_to_id = function + | RIdent (_,id) -> id + | RQualid (loc,_) -> + user_err_loc (loc, "Ast.coerce_reference_to_id", + str"This expression should be a simple identifier") + +(* This is to interpret the macro $ABSTRACT used in binders *) +(* $ABSTRACT should occur in this configuration : *) +(* ($ABSTRACT name (s1 a1 ($LIST l1)) ... (s2 an ($LIST ln)) b) *) +(* where li is id11::...::id1p1 and it produces the ast *) +(* (s1' a1 [id11]...[id1p1](... (sn' an [idn1]...[idnpn]b)...)) *) +(* where s1' is overwritten by name if s1 is $BINDER otherwise s1 *) + +let slam_ast (_,fin) id ast = + match id with + | Coqast.Nvar ((deb,_), s) -> Coqast.Slam ((deb,fin), Some s, ast) + | Coqast.Nmeta ((deb,_), s) -> Coqast.Smetalam ((deb,fin), s, ast) + | _ -> invalid_arg "Ast.slam_ast" + +let abstract_binder_ast (_,fin as loc) name a b = + match a with + | Coqast.Node((deb,_),s,d::l) -> + let s' = if s="BINDER" then name else s in + Coqast.Node((deb,fin),s', [d; List.fold_right (slam_ast loc) l b]) + | _ -> invalid_arg "Bad usage of $ABSTRACT macro" + +let abstract_binders_ast loc name a b = + match a with + | Coqast.Node(_,"BINDERS",l) -> + List.fold_right (abstract_binder_ast loc name) l b + | _ -> invalid_arg "Bad usage of $ABSTRACT macro" + +(* Pattern-matching on ast *) let env_assoc_value loc v env = try @@ -202,107 +276,22 @@ let env_assoc_value loc v env = let env_assoc_list sigma (loc,v) = match env_assoc_value loc v sigma with - | Vastlist al -> al - | Vast a -> [a] + | AstListNode al -> al + | PureAstNode a -> [a] let env_assoc sigma k (loc,v) = match env_assoc_value loc v sigma with - | Vast a -> check_cast loc a k; a + | PureAstNode a -> check_cast loc a k; a | _ -> grammar_type_error (loc,"Ast.env_assoc: "^v^" is an ast list") let env_assoc_nvars sigma (dloc,v) = match env_assoc_value dloc v sigma with - | Vastlist al -> List.map coerce_to_id al - | Vast ast -> [coerce_to_id ast] + | AstListNode al -> List.map coerce_to_id al + | PureAstNode ast -> [coerce_to_id ast] let build_lams dloc idl ast = List.fold_right (fun id lam -> Slam(dloc,Some id,lam)) idl ast - -(* Substitution *) - -(* Locations in quoted ast are wrong (they refer to the right hand - side of a grammar rule). A default location dloc is used whenever - we create an ast constructor. Locations in the binding list are trusted. *) - -let rec pat_sub dloc sigma pat = - match pat with - | Pmeta(pv,c) -> env_assoc sigma c (dloc,pv) - | Pmeta_slam(pv,p) -> - let idl = env_assoc_nvars sigma (dloc,pv) in - build_lams dloc idl (pat_sub dloc sigma p) - | Pquote a -> set_loc dloc a - | Pnode(op,pl) -> Node(dloc, op, patl_sub dloc sigma pl) - | Pslam(os,p) -> Slam(dloc, os, pat_sub dloc sigma p) - -and patl_sub dloc sigma pl = - match pl with - | Pnil -> - [] - | Plmeta pv -> - env_assoc_list sigma (dloc,pv) - | Pcons(Pmeta(pv,Tlist), ptl) -> - (env_assoc_list sigma (dloc,pv))@(patl_sub dloc sigma ptl) - | Pcons(p1,ptl) -> - (pat_sub dloc sigma p1)::(patl_sub dloc sigma ptl) - - -(* Converting and checking free meta-variables *) - -type entry_env = (string * entry_type) list - -let type_of_meta env loc pv = - try - List.assoc pv env - with Not_found -> - user_err_loc (loc,"Ast.type_of_meta", - (str"variable " ++ str pv ++ str" is unbound")) - -let check_ast_meta env loc pv = - if (type_of_meta env loc pv) <> ETast then - user_err_loc (loc,"Ast.check_ast_meta", - (str"variable " ++ str pv ++ str" is a List")) - -let rec val_of_ast env ast = - match ast with - | Nmeta(loc,pv) -> - check_ast_meta env loc pv; - Pmeta(pv,Tany) -(* - | Id(loc,pv) when isMeta pv -> - check_ast_meta env loc pv; - Pmeta(pv,Tid) -*) - | Node(_,"$QUOTE",[qast]) -> Pquote (set_loc dummy_loc qast) - | Smetalam(loc,s,a) -> - let _ = type_of_meta env loc s in (* ids are coerced to id lists *) - Pmeta_slam(s, val_of_ast env a) - | (Path _|Num _|Id _|Str _|Nvar _) -> Pquote (set_loc dummy_loc ast) - | Slam(_,os,b) -> Pslam(os, val_of_ast env b) - | Node(loc,op,_) when isMeta op -> - user_err_loc(loc,"Ast.val_of_ast", - (str"no metavariable in operator position.")) - | Node(_,op,args) -> Pnode(op, vall_of_astl env args) - | Dynamic(loc,_) -> - invalid_arg_loc(loc,"val_of_ast: dynamic") - -and vall_of_astl env astl = - match astl with - | (Node(loc,"$LIST",[Nmeta(locv,pv)]))::asttl -> - if type_of_meta env locv pv = ETastl then - if asttl = [] then - Plmeta pv - else - Pcons(Pmeta(pv,Tlist), vall_of_astl env asttl) - else - user_err_loc (loc,"Ast.vall_of_astl", - (str"variable " ++ str pv ++ - str" is not a List")) - | ast::asttl -> - Pcons (val_of_ast env ast, vall_of_astl env asttl) - | [] -> Pnil - - (* Alpha-conversion *) let rec alpha_var id1 id2 = function @@ -330,31 +319,15 @@ let rec alpha alp a1 a2 = let alpha_eq (a1,a2)= alpha [] a1 a2 +let alpha_eq_val (x,y) = x = y +(* let alpha_eq_val = function | (Vast a1, Vast a2) -> alpha_eq (a1,a2) | (Vastlist al1, Vastlist al2) -> (List.length al1 = List.length al2) & List.for_all2 (fun x y -> alpha_eq (x,y)) al1 al2 | _ -> false - -let rec occur_var_ast s = function - | Node(loc,op,args) -> List.exists (occur_var_ast s) args - | Nvar(_,s2) -> s = s2 - | Smetalam _ | Nmeta _ -> anomaly "occur_var: metas should not occur here" - | Slam(_,sopt,body) -> (Some s <> sopt) & occur_var_ast s body - | Id _ | Str _ | Num _ | Path _ -> false - | Dynamic _ -> (* Hum... what to do here *) false - -let rec replace_vars_ast l = function - | Node(loc,op,args) -> Node (loc,op, List.map (replace_vars_ast l) args) - | Nvar(loc,s) as a -> (try Nvar (loc, List.assoc s l) with Not_found -> a) - | Smetalam _ | Nmeta _ -> anomaly "replace_var: metas should not occur here" - | Slam(loc,None,body) -> Slam(loc,None,replace_vars_ast l body) - | Slam(loc,Some s,body) as a -> - if List.mem_assoc s l then a else - Slam(loc,Some s,replace_vars_ast l body) - | Id _ | Str _ | Num _ | Path _ as a -> a - | Dynamic _ as a -> (* Hum... what to do here *) a +*) exception No_match of string @@ -379,7 +352,7 @@ let bind_env sigma var v = let bind_env_ast sigma var ast = try - bind_env sigma var (Vast ast) + bind_env sigma var (PureAstNode ast) with e -> Stdpp.raise_with_loc (loc ast) e @@ -387,7 +360,7 @@ let alpha_define sigma loc ps s = try let _ = List.assoc ps sigma in sigma with Not_found -> - if ps = "$_" then sigma else (ps, Vast(Nvar(loc,s)))::sigma + if ps = "$_" then sigma else (ps, PureAstNode(Nvar(loc,s)))::sigma (* Match an ast with an ast pattern. Returns the new environnement. *) @@ -422,12 +395,27 @@ and amatchl alp sigma spatl astl = | (Pnil, []) -> sigma | (Pcons(pat,patl), ast::asttl) -> amatchl alp (amatch alp sigma pat ast) patl asttl - | (Plmeta pv, _) -> bind_env sigma pv (Vastlist astl) + | (Plmeta pv, _) -> bind_env sigma pv (AstListNode astl) | _ -> raise (No_match "Ast.amatchl") let ast_match = amatch [] + +let typed_ast_match env p a = match (p,a) with + | PureAstPat p, PureAstNode a -> amatch [] env p a + | AstListPat pl, AstListNode al -> amatchl [] env pl al + | _ -> failwith "Ast.typed_ast_match: TODO" + let astl_match = amatchl [] +let ast_first_match pat_of_fun env ast sl = + let rec aux = function + | [] -> None + | h::t -> + (try Some (h, ast_match env (pat_of_fun h) ast) + with (No_match _| Stdpp.Exc_located (_,No_match _)) -> aux t) + in + aux sl + let first_match pat_of_fun env ast sl = let rec aux = function | [] -> None @@ -515,12 +503,157 @@ and patl_of_astl env astl = let (p1,env1) = pat_of_ast env ast in let (ptl,env2) = patl_of_astl env1 asttl in (Pcons (p1,ptl), env2) + +type entry_env = (string * ast_action_type) list -let to_pat env ast = +(* +let to_pat env = function + | AstListNode al -> let p,e = patl_of_astl env al in AstListPat p, e + | PureAstNode a -> let p,e = pat_of_ast env a in PureAstPat p, e +*) + +let to_pat = pat_of_ast + +(* match ast with | Node(_,"ASTPAT",[p]) -> pat_of_ast env p | _ -> invalid_arg_loc (loc ast,"Ast.to_pat") +*) + + +(* Substitution *) +(* Locations in quoted ast are wrong (they refer to the right hand + side of a grammar rule). A default location dloc is used whenever + we create an ast constructor. Locations in the binding list are trusted. *) + +let rec pat_sub dloc sigma pat = + match pat with + | Pmeta(pv,c) -> env_assoc sigma c (dloc,pv) + | Pmeta_slam(pv,p) -> + let idl = env_assoc_nvars sigma (dloc,pv) in + build_lams dloc idl (pat_sub dloc sigma p) + | Pquote a -> set_loc dloc a + | Pnode(op,pl) -> Node(dloc, op, patl_sub dloc sigma pl) + | Pslam(os,p) -> Slam(dloc, os, pat_sub dloc sigma p) + +and patl_sub dloc sigma pl = + match pl with + | Pnil -> + [] + | Plmeta pv -> + env_assoc_list sigma (dloc,pv) + | Pcons(Pmeta(pv,Tlist), ptl) -> + (env_assoc_list sigma (dloc,pv))@(patl_sub dloc sigma ptl) + | Pcons(p1,ptl) -> + (pat_sub dloc sigma p1)::(patl_sub dloc sigma ptl) + +(* Converting and checking free meta-variables *) + +let type_of_meta env loc pv = + try + List.assoc pv env + with Not_found -> + user_err_loc (loc,"Ast.type_of_meta", + (str"variable " ++ str pv ++ str" is unbound")) + +let check_ast_meta env loc pv = + match type_of_meta env loc pv with + | ETast -> () + | _ -> + user_err_loc (loc,"Ast.check_ast_meta", + (str"variable " ++ str pv ++ str" is not of ast type")) + +let rec val_of_ast env = function + | Nmeta(loc,pv) -> + check_ast_meta env loc pv; + Pmeta(pv,Tany) + | Node(_,"$QUOTE",[qast]) -> Pquote (set_loc dummy_loc qast) + | Smetalam(loc,s,a) -> + let _ = type_of_meta env loc s in (* ids are coerced to id lists *) + Pmeta_slam(s, val_of_ast env a) + | (Path _|Num _|Id _|Str _|Nvar _ as ast) -> Pquote (set_loc dummy_loc ast) + | Slam(_,os,b) -> Pslam(os, val_of_ast env b) + | Node(loc,op,_) when isMeta op -> + user_err_loc(loc,"Ast.val_of_ast", + (str"no metavariable in operator position.")) + | Node(_,op,args) -> Pnode(op, vall_of_astl env args) + | Dynamic(loc,_) -> + invalid_arg_loc(loc,"val_of_ast: dynamic") + +and vall_of_astl env = function + | (Node(loc,"$LIST",[Nmeta(locv,pv)]))::asttl -> + if type_of_meta env locv pv = ETastl then + if asttl = [] then + Plmeta pv + else + Pcons(Pmeta(pv,Tlist), vall_of_astl env asttl) + else + user_err_loc + (loc,"Ast.vall_of_astl", + str"variable " ++ str pv ++ str" is not a List") + | ast::asttl -> Pcons (val_of_ast env ast, vall_of_astl env asttl) + | [] -> Pnil +(* +let rec val_of_ast_constr env = function +(* + | ConstrEval (r,c) -> ConstrEvalPat (r,val_of_ast_constr env c) + | ConstrContext (x,c) -> ConstrContextPat (x,val_of_ast_constr env c) +*) + | ConstrTerm c -> ConstrTermPat (val_of_ast env c) +*) +(* +let rec check_pat_meta env = function + | Pquote _ -> () + | Pmeta(s,Tany) -> check_ast_meta env loc s + | Pmeta(s,_) -> anomaly "not well-formed pattern" + | Pmeta_slam(s,b) -> + let _ = type_of_meta env loc s in (* ids are coerced to id lists *) + check_pat_meta env b + | Pslam(_,b) -> check_pat_meta env b + | Pnode(op,Plmeta (locv,pv)) -> + if type_of_meta env locv pv <> ETastl then + user_err_loc (locv,"Ast.vall_of_astl", + [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >]) + | Pnode(op,l) -> check_patlist_meta env l + +and check_patlist_meta env = function + | Plmeta (locv,pv) -> + if type_of_meta env locv pv <> ETastl then + user_err_loc (locv,"Ast.vall_of_astl", + [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >]) + | Pcons(Pmeta(pv,Tlist),l) -> + if l = Pnil then anomaly "not well-formed pattern list"; + if type_of_meta env locv pv <> ETastl then + user_err_loc (locv,"Ast.vall_of_astl", + [< 'sTR"variable "; 'sTR pv; 'sTR" is not a List" >]) + else check_patlist_meta env l + | Pcons(p,l) -> check_pat_meta env p; check_patlist_meta env l + | Pnil -> () + +let check_typed_pat_meta env = function + | AstListPat cl -> check_patlist_meta env cl + | PureAstPat c -> check_pat_meta env c +*) + +let rec occur_var_ast s = function + | Node(loc,op,args) -> List.exists (occur_var_ast s) args + | Nvar(_,s2) -> s = s2 + | Smetalam _ | Nmeta _ -> anomaly "occur_var: metas should not occur here" + | Slam(_,sopt,body) -> (Some s <> sopt) & occur_var_ast s body + | Id _ | Str _ | Num _ | Path _ -> false + | Dynamic _ -> (* Hum... what to do here *) false + +let rec replace_vars_ast l = function + | Node(loc,op,args) -> Node (loc,op, List.map (replace_vars_ast l) args) + | Nvar(loc,s) as a -> (try Nvar (loc, List.assoc s l) with Not_found -> a) + | Smetalam _ | Nmeta _ -> anomaly "replace_var: metas should not occur here" + | Slam(loc,None,body) -> Slam(loc,None,replace_vars_ast l body) + | Slam(loc,Some s,body) as a -> + if List.mem_assoc s l then a else + Slam(loc,Some s,replace_vars_ast l body) + | Id _ | Str _ | Num _ | Path _ as a -> a + | Dynamic _ as a -> (* Hum... what to do here *) a (* Ast with cases and metavariables *) @@ -549,67 +682,63 @@ let caselist_failed loc sigma el pats = brk(1,1) ++ v 0 (prlist_with_sep pr_spc print_astlpat pats) ++ fnl () ++ print_sig sigma) -let rec eval_act dloc sigma act = - match act with - | Aast pat -> Vast (pat_sub dloc sigma pat) - | Aastlist patl -> Vastlist (patl_sub dloc sigma patl) - | Acase(e,ml) -> - (match eval_act dloc sigma e with - | Vast esub -> - (match first_match fst sigma esub ml with - | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a - | _ -> case_failed dloc sigma esub (List.map fst ml)) - | _ -> grammar_type_error (dloc,"Ast.eval_act")) - | Acaselist(e,ml) -> - (match eval_act dloc sigma e with - | Vastlist elsub -> - (match first_matchl fst sigma elsub ml with - | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a - | _ -> caselist_failed dloc sigma elsub (List.map fst ml)) - | _ -> grammar_type_error (dloc,"Ast.eval_act")) +let myfst = function + | PureAstPat p, a -> p + | _ -> error "Expects a pure ast" + +let myfstl = function + | AstListPat p, a -> p + | _ -> error "Expects an ast list" + +let rec eval_act dloc sigma = function + | Act (AstListPat patl) -> AstListNode (patl_sub dloc sigma patl) + | Act (PureAstPat pat) -> PureAstNode (pat_sub dloc sigma pat) + | ActCase(e,ml) -> + (match eval_act dloc sigma e with + | (PureAstNode esub) -> + (match ast_first_match myfst sigma esub ml with + | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a + | _ -> case_failed dloc sigma esub (List.map myfst ml)) + | _ -> grammar_type_error (dloc,"Ast.eval_act")) + | ActCaseList(e,ml) -> + (match eval_act dloc sigma e with + | AstListNode elsub -> + (match first_matchl myfstl sigma elsub ml with + | Some((_,a),sigma_pat) -> eval_act dloc sigma_pat a + | _ -> caselist_failed dloc sigma elsub (List.map myfstl ml)) + | _ -> grammar_type_error (dloc,"Ast.eval_act")) + +let val_of_typed_ast loc env = function + | ETast, PureAstNode c -> PureAstPat (val_of_ast env c) + | ETastl, AstListNode cl -> AstListPat (vall_of_astl env cl) + | (ETast|ETastl), _ -> + invalid_arg_loc (loc,"Ast.act_of_ast: ill-typed") (* TODO: case sur des variables uniquement -> pas de pb de conflit Ast/List *) -let rec act_of_ast vars etyp ast = - match ast with - | Node(_,"$CASE",(a::et::cl)) -> - let atyp = entry_type et in - let pa = act_of_ast vars atyp a in - (match atyp with - | ETast -> - let acl = List.map (case vars etyp) cl in - Acase (pa,acl) - | ETastl -> - let acl = List.map (caselist vars etyp) cl in - Acaselist (pa,acl)) - | Node(_,"ASTLIST",al) when etyp = ETastl -> - Aastlist (vall_of_astl vars al) - | a when etyp = ETast -> - Aast (val_of_ast vars a) - | _ -> - invalid_arg_loc (loc ast,"Ast.act_of_ast: ill-typed") - -and case vars etyp ast = - match ast with - | Node(_,"CASE",[Node(loca,"ASTLIST",pl);a]) -> - (match pl with - | [p] -> - let (apl,penv) = pat_of_ast vars p in - let aa = act_of_ast penv etyp a in - (apl,aa) - | _ -> user_err_loc - (loca,"Ast.case", - (str"case pattern for an ast should be a single ast"))) - | _ -> invalid_arg_loc (loc ast,"Ast.case") +let rec act_of_ast vars etyp = function + | CaseAction (loc,a,atyp,cl) -> + let pa = act_of_ast vars atyp a in + (match atyp with + | ETastl -> + let acl = List.map (caselist vars etyp) cl in + ActCaseList (pa,acl) + | _ -> + let acl = List.map (case loc vars etyp) cl in + ActCase (pa,acl)) + | SimpleAction (loc,a) -> Act (val_of_typed_ast loc vars (etyp,a)) + +and case loc vars etyp = function + | [p],a -> + let (apl,penv) = pat_of_ast vars p in + let aa = act_of_ast penv etyp a in + (PureAstPat apl,aa) + | _ -> + user_err_loc + (loc, "Ast.case", str"case pattern for an ast should be a single ast") -and caselist vars etyp ast = - match ast with - | Node(_,"CASE",[Node(_,"ASTLIST",pl);a]) -> - let (apl,penv) = patl_of_astl vars pl in - let aa = act_of_ast penv etyp a in - (apl,aa) - | _ -> invalid_arg_loc (loc ast,"Ast.caselist") +and caselist vars etyp (pl,a) = + let (apl,penv) = patl_of_astl vars pl in + let aa = act_of_ast penv etyp a in + (AstListPat apl,aa) -let to_act_check_vars env etyp ast = - match ast with - | Node(_,"ASTACT",[a]) -> act_of_ast env etyp a - | _ -> invalid_arg_loc (loc ast,"Ast.to_act_env") +let to_act_check_vars = act_of_ast diff --git a/parsing/ast.mli b/parsing/ast.mli index fae49ac34..fd7581199 100755 --- a/parsing/ast.mli +++ b/parsing/ast.mli @@ -11,13 +11,17 @@ (*i*) open Pp open Names -open Pcoq +open Coqast +open Genarg (*i*) (* Abstract syntax trees. *) val dummy_loc : Coqast.loc val loc : Coqast.t -> Coqast.loc +(* +val vernac_loc : Coqast.vernac_ast -> Coqast.loc +*) (* ast constructors with dummy location *) val ope : string * Coqast.t list -> Coqast.t @@ -38,43 +42,70 @@ val section_path : section_path -> section_path val num_of_ast : Coqast.t -> int val id_of_ast : Coqast.t -> string val nvar_of_ast : Coqast.t -> identifier +val meta_of_ast : Coqast.t -> string (* ast processing datatypes *) +type entry_type = + | PureAstType + | IntAstType + | IdentAstType + | AstListType + | TacticAtomAstType + | ThmTokenAstType + | DynamicAstType + | ReferenceAstType + | GenAstType of Genarg.argument_type + (* patterns of ast *) -type pat = - | Pquote of Coqast.t +type astpat = + | Pquote of t | Pmeta of string * tok_kind | Pnode of string * patlist - | Pslam of identifier option * pat - | Pmeta_slam of string * pat + | Pslam of identifier option * astpat + | Pmeta_slam of string * astpat and patlist = - | Pcons of pat * patlist + | Pcons of astpat * patlist | Plmeta of string | Pnil and tok_kind = Tnum | Tid | Tstr | Tpath | Tvar | Tany | Tlist +type pat = + | AstListPat of patlist + | PureAstPat of astpat + (* semantic actions of grammar rules *) type act = - | Aast of pat - | Aastlist of patlist - | Acase of act * (pat * act) list - | Acaselist of act * (patlist * act) list + | Act of pat + | ActCase of act * (pat * act) list + | ActCaseList of act * (pat * act) list (* values associated to variables *) -type v = - | Vast of Coqast.t - | Vastlist of Coqast.t list +type typed_ast = + | AstListNode of Coqast.t list + | PureAstNode of Coqast.t + +type ast_action_type = ETast | ETastl + +type grammar_action = + | SimpleAction of loc * typed_ast + | CaseAction of + loc * grammar_action * ast_action_type * (t list * grammar_action) list -type env = (string * v) list +type env = (string * typed_ast) list val coerce_to_var : Coqast.t -> Coqast.t -(* val coerce_to_id : Coqast.t -> identifier -*) + +val coerce_qualid_to_id : Nametab.qualid Util.located -> identifier + +val coerce_reference_to_id : Tacexpr.reference_expr -> identifier + +val abstract_binders_ast : + Coqast.loc -> string -> Coqast.t -> Coqast.t -> Coqast.t exception No_match of string @@ -82,35 +113,36 @@ val isMeta : string -> bool val print_ast : Coqast.t -> std_ppcmds val print_astl : Coqast.t list -> std_ppcmds -val print_astpat : pat -> std_ppcmds +val print_astpat : astpat -> std_ppcmds val print_astlpat : patlist -> std_ppcmds (* Meta-syntax operations: matching and substitution *) -type entry_env = (string * entry_type) list +type entry_env = (string * ast_action_type) list val grammar_type_error : Coqast.loc * string -> 'a (* Converting and checking free meta-variables *) -val pat_sub : Coqast.loc -> env -> pat -> Coqast.t -val val_of_ast : entry_env -> Coqast.t -> pat +val pat_sub : Coqast.loc -> env -> astpat -> Coqast.t +val val_of_ast : entry_env -> Coqast.t -> astpat val vall_of_astl : entry_env -> Coqast.t list -> patlist +val pat_of_ast : entry_env -> Coqast.t -> astpat * entry_env + val alpha_eq : Coqast.t * Coqast.t -> bool -val alpha_eq_val : v * v -> bool +val alpha_eq_val : typed_ast * typed_ast -> bool val occur_var_ast : identifier -> Coqast.t -> bool val replace_vars_ast : (identifier * identifier) list -> Coqast.t -> Coqast.t -val bind_env : env -> string -> v -> env -val ast_match : env -> pat -> Coqast.t -> env +val bind_env : env -> string -> typed_ast -> env +val ast_match : env -> astpat -> Coqast.t -> env val astl_match : env -> patlist -> Coqast.t list -> env -val first_match : ('a -> pat) -> env -> Coqast.t -> 'a list -> - ('a * env) option +val first_match : ('a -> astpat) -> env -> t -> 'a list -> ('a * env) option val first_matchl : ('a -> patlist) -> env -> Coqast.t list -> 'a list -> ('a * env) option -val to_pat : entry_env -> Coqast.t -> (pat * entry_env) +val to_pat : entry_env -> Coqast.t -> (astpat * entry_env) -val eval_act : Coqast.loc -> env -> act -> v -val to_act_check_vars : entry_env -> entry_type -> Coqast.t -> act +val eval_act : Coqast.loc -> env -> act -> typed_ast +val to_act_check_vars : entry_env -> ast_action_type -> grammar_action -> act diff --git a/parsing/astterm.ml b/parsing/astterm.ml index fa44a19ac..0985d309f 100644 --- a/parsing/astterm.ml +++ b/parsing/astterm.ml @@ -599,6 +599,15 @@ let ast_to_rawconstr sigma env allow_soapp lvar = | a::args -> (dbrec env a) :: (ast_to_args env args) | [] -> [] + and interp_binding env = function + | Node(_,"BINDING", [Num(_,n);Node(loc,"CONSTR",[c])]) -> + (AnonHyp n,dbrec env c) + | Node(_,"BINDING", [Nvar(loc0,s); Node(loc1,"CONSTR",[c])]) -> + (NamedHyp (ident_of_nvar loc0 s), dbrec env c) + | x -> + errorlabstrm "bind_interp" + (str "Not the expected form in binding" ++ print_ast x) + in dbrec env @@ -630,10 +639,10 @@ let implicits_of_extended_reference = function let warning_globalize qid = warning ("Could not globalize " ^ (string_of_qualid qid)) -let globalize_qualid qid = +let globalize_qualid (loc,qid) = try let ref = Nametab.extended_locate qid in - ast_of_extended_ref_loc dummy_loc ref + ast_of_extended_ref_loc loc ref with Not_found -> if_verbose warning_globalize qid; Termast.ast_of_qualid qid @@ -736,6 +745,10 @@ let constrOut = function anomalylabstrm "constrOut" (str "Not a Dynamic ast: " ++ print_ast ast) +let interp_global_constr env (loc,qid) = + let c,_ = rawconstr_of_qualid (Idset.empty,[]) ([],env) loc qid in + understand Evd.empty env c + let interp_constr sigma env c = understand sigma env (interp_rawconstr sigma env c) diff --git a/parsing/astterm.mli b/parsing/astterm.mli index 6edb387bc..3074ff665 100644 --- a/parsing/astterm.mli +++ b/parsing/astterm.mli @@ -24,7 +24,13 @@ open Pattern val constrIn : constr -> Coqast.t val constrOut : Coqast.t -> constr +(* Interprets global names, including syntactic defs and section variables *) +val interp_global_constr : env -> Nametab.qualid Util.located -> constr + val interp_rawconstr : evar_map -> env -> Coqast.t -> rawconstr +val interp_rawconstr_gen : + evar_map -> env -> (identifier * Impargs.implicits_list) list -> + bool -> identifier list -> Coqast.t -> rawconstr val interp_constr : evar_map -> env -> Coqast.t -> constr val interp_casted_constr : evar_map -> env -> Coqast.t -> types -> constr val interp_type : evar_map -> env -> Coqast.t -> types @@ -75,7 +81,9 @@ val interp_constrpattern : bound idents in grammar or pretty-printing rules) *) val globalize_constr : Coqast.t -> Coqast.t val globalize_ast : Coqast.t -> Coqast.t -val globalize_qualid : Nametab.qualid -> Coqast.t +val globalize_qualid : Nametab.qualid Util.located -> Coqast.t + +val ast_of_extended_ref_loc : loc -> Nametab.extended_global_reference -> Coqast.t (* This transforms args of a qualid keyword into a qualified ident *) (* it does no relocation *) diff --git a/parsing/coqast.ml b/parsing/coqast.ml index 7604818f2..9b65bdfb1 100644 --- a/parsing/coqast.ml +++ b/parsing/coqast.ml @@ -97,3 +97,150 @@ let hcons_ast (hstr,hid,hpath) = let hloc = Hashcons.simple_hcons Hloc.f () in let hast = Hashcons.recursive_hcons Hast.f (hloc,hstr,hid,hpath) in (hast,hloc) + +(* +type 'vernac_ast raw_typed_ast = + | PureAstNode of t + | AstListNode of t list + | PureAstNode of t + | TacticAstNode of tactic_expr + | VernacAstNode of 'vernac_ast + +type entry_type = + | PureAstType + | AstListType + | VernacAstType + | TacticAstType + | VernacAtomAstType + | DynamicAstType + | QualidAstType + | ConstrAstType + | BinderAstType + | ThmTokAstType + | BindingListAstType + +type astpat = + | Pquote of t + | Pmeta of string * tok_kind + | Pnode of string * patlist + | Pslam of identifier option * astpat + | Pmeta_slam of string * astpat + +and patlist = + | Pcons of astpat * patlist + | Plmeta of string + | Pnil + +type 'a syntax_rule = string * 'a raw_typed_ast * t unparsing_hunk list +type 'a raw_syntax_entry_ast = precedence * 'a syntax_rule list + +type grammar_associativity = Gramext.g_assoc option +type 'a raw_grammar_action = + | SimpleAction of loc * 'a raw_typed_ast + | CaseAction of loc * + 'a raw_grammar_action * entry_type option + * (t list * 'a raw_grammar_action) list +type grammar_production = + VTerm of string | VNonTerm of loc * nonterm * string option +type 'a grammar_rule = string * grammar_production list * 'a raw_grammar_action +type 'a raw_grammar_entry_ast = + (loc * string) * entry_type option * grammar_associativity * 'a grammar_rule list + +type evaluable_global_reference_ast = + | AEvalVarRef of identifier + | AEvalConstRef of section_path + +type flags_ast = int + +type red_expr_ast = (t,t,t) Rawterm.red_expr_gen + +type vernac_arg = + | VARG_VARGLIST of vernac_arg list + | VARG_STRING of string + | VARG_NUMBER of int + | VARG_NUMBERLIST of int list + | VARG_IDENTIFIER of identifier + | VARG_QUALID of Nametab.qualid + | VCALL of string * vernac_arg list + | VARG_CONSTR of t + | VARG_CONSTRLIST of t list + | VARG_TACTIC of tactic_expr + | VARG_REDEXP of red_expr_ast + | VARG_BINDER of identifier list * t + | VARG_BINDERLIST of (identifier list * t) list + | VARG_AST of t + | VARG_ASTLIST of t list + | VARG_UNIT + | VARG_DYN of Dyn.t + +type def_kind = DEFINITION | LET | LOCAL | THEOREM | LETTOP | DECL | REMARK + | FACT | LEMMA + | COERCION | LCOERCION | OBJECT | LOBJECT | OBJCOERCION | LOBJCOERCION + | SUBCLASS | LSUBCLASS + +open Nametab + +type declaration_hook = strength -> global_reference -> unit +let add_coercion = ref (fun _ _ -> ()) +let add_subclass = ref (fun _ _ -> ()) +let add_object = ref (fun _ _ -> ()) + +type constr_ast = t +type sort_expr = t + +type simple_binders = identifier * constr_ast +type constructor_ast = identifier * constr_ast +type inductive_ast = + identifier * simple_binders list * constr_ast * constructor_ast list +type fixpoint_ast = + identifier * simple_binders list * constr_ast * constr_ast +type cofixpoint_ast = + identifier * constr_ast * constr_ast + +type local_decl_expr = + | AssumExpr of identifier * constr_ast + | DefExpr of identifier * constr_ast * constr_ast option + +type vernac_atom_ast = + (* Syntax *) + | VernacGrammar of string * vernac_ast raw_grammar_entry_ast list + | VernacSyntax of string * vernac_ast raw_syntax_entry_ast list + | VernacInfix of grammar_associativity * int * string * t + | VernacDistfix of grammar_associativity * int * string * t + (* Gallina *) + | VernacDefinition of (bool * strength) * identifier * t option * constr_ast * constr_ast option * declaration_hook + | VernacStartProof of (bool * strength) * identifier * constr_ast * bool * declaration_hook + | VernacEndProof of bool * strength option * identifier option + | VernacAssumption of strength * (identifier list * constr_ast) list + | VernacInductive of bool * inductive_ast list + | VernacFixpoint of fixpoint_ast list + | VernacCoFixpoint of cofixpoint_ast list + (* Gallina extensions *) + | VernacRecord of bool * identifier * simple_binders list * sort_expr * identifier option * (bool * local_decl_expr) list + + (* Commands *) + | TacticDefinition of loc * (identifier * tactic_expr) list + | VernacSetOpacity of bool * (loc * identifier list) list + | VernacSolve of int * tactic_expr + (* For temporary compatibility *) + | ExtraVernac of t + (* For extension *) + | VernacExtend of string * vernac_arg list + (* For actions in Grammar and patterns in Syntax rules *) + | VernacMeta of loc * string + +and located_vernac_atom_ast = loc * vernac_atom_ast + +and vernac_ast = + | VTime of located_vernac_atom_ast + | VLoad of bool * string + | VernacList of located_vernac_atom_ast list + | VernacVar of identifier + +type pat = vernac_ast raw_pat +type typed_ast = vernac_ast raw_typed_ast +type grammar_action = vernac_ast raw_grammar_action +type grammar_entry_ast = vernac_ast raw_grammar_entry_ast +type syntax_entry_ast = vernac_ast raw_syntax_entry_ast + +*) diff --git a/parsing/coqast.mli b/parsing/coqast.mli index f741574e8..b5ad45c74 100644 --- a/parsing/coqast.mli +++ b/parsing/coqast.mli @@ -40,3 +40,10 @@ val hcons_ast: (string -> string) * (Names.identifier -> Names.identifier) * (section_path -> section_path) -> (t -> t) * (loc -> loc) + +(* +val map_tactic_expr : (t -> t) -> (tactic_expr -> tactic_expr) -> tactic_expr -> tactic_expr +val fold_tactic_expr : + ('a -> t -> 'a) -> ('a -> tactic_expr -> 'a) -> 'a -> tactic_expr -> 'a +val iter_tactic_expr : (tactic_expr -> unit) -> tactic_expr -> unit +*) diff --git a/parsing/coqlib.ml b/parsing/coqlib.ml index 022b6942f..07e439af8 100644 --- a/parsing/coqlib.ml +++ b/parsing/coqlib.ml @@ -27,16 +27,12 @@ let datatypes_module = make_dir ["Coq";"Init";"Datatypes"] let arith_module = make_dir ["Coq";"Arith";"Arith"] let nat_path = make_path datatypes_module (id_of_string "nat") -let myvar_path = - make_path arith_module (id_of_string "My_special_variable") let glob_nat = IndRef (nat_path,0) let glob_O = ConstructRef ((nat_path,0),1) let glob_S = ConstructRef ((nat_path,0),2) -let glob_My_special_variable_nat = ConstRef myvar_path - let eq_path = make_path logic_module (id_of_string "eq") let eqT_path = make_path logic_type_module (id_of_string "eqT") @@ -242,12 +238,11 @@ let coq_ex_pattern_gen ex = let coq_existS_pattern = coq_ex_pattern_gen coq_existS_ref let coq_existT_pattern = coq_ex_pattern_gen coq_existT_ref -(* Patterns "~ ?", "? -> False" and "(?1 -> ?2)" *) +(* Patterns "~ ?" and "? -> False" *) let coq_not_pattern = lazy(PApp(PRef (Lazy.force coq_not_ref), [|PMeta None|])) let imp a b = PProd (Anonymous, a, b) let coq_imp_False_pattern = lazy (imp (PMeta None) (PRef (Lazy.force coq_False_ref))) -let coq_arrow_pattern = lazy (imp (PMeta (Some 1)) (PMeta (Some 2))) (* Pattern "(sumbool (eq ?1 ?2 ?3) ?4)" *) let coq_eqdec_partial_pattern = @@ -272,18 +267,6 @@ let coq_eqdec_pattern = [|PApp (PRef (Lazy.force coq_eq_ref), [| PMeta (Some 1); PRel 2; PRel 1 |])|]) |])))) -(* "(A : ?)(x:A)(? A x x)" and "(x : ?)(? x x)" *) -let name_A = Name (id_of_string "A") -let coq_refl_rel1_pattern = - lazy - (PProd - (name_A, PMeta None, - PProd (x, PRel 1, PApp (PMeta None, [|PRel 2; PRel 1; PRel 1|])))) -let coq_refl_rel2_pattern = - lazy - (PProd (x, PMeta None, PApp (PMeta None, [|PRel 1; PRel 1|]))) - - let build_coq_eq_pattern () = Lazy.force coq_eq_pattern let build_coq_eqT_pattern () = Lazy.force coq_eqT_pattern let build_coq_idT_pattern () = Lazy.force coq_idT_pattern @@ -293,7 +276,4 @@ let build_coq_not_pattern () = Lazy.force coq_not_pattern let build_coq_imp_False_pattern () = Lazy.force coq_imp_False_pattern let build_coq_eqdec_partial_pattern () = Lazy.force coq_eqdec_partial_pattern let build_coq_eqdec_pattern () = Lazy.force coq_eqdec_pattern -let build_coq_arrow_pattern () = Lazy.force coq_arrow_pattern -let build_coq_refl_rel1_pattern () = Lazy.force coq_refl_rel1_pattern -let build_coq_refl_rel2_pattern () = Lazy.force coq_refl_rel2_pattern let build_coq_sig_pattern () = Lazy.force coq_sig_pattern diff --git a/parsing/coqlib.mli b/parsing/coqlib.mli index dc65d7ab8..0b5bb0ec5 100644 --- a/parsing/coqlib.mli +++ b/parsing/coqlib.mli @@ -29,9 +29,6 @@ val glob_nat : global_reference val glob_O : global_reference val glob_S : global_reference -(* Special variable for pretty-printing of constant naturals *) -val glob_My_special_variable_nat : global_reference - (* Equality *) val glob_eq : global_reference val glob_eqT : global_reference @@ -129,12 +126,5 @@ val build_coq_eqdec_partial_pattern : constr_pattern delayed (* ["! (x,y:?1). (sumbool (eq ?1 x y) ~(eq ?1 x y))"] *) val build_coq_eqdec_pattern : constr_pattern delayed -(* ["(A : ?)(x:A)(? A x x)"] and ["(x : ?)(? x x)"] *) -val build_coq_refl_rel1_pattern : constr_pattern delayed -val build_coq_refl_rel2_pattern : constr_pattern delayed - -(* ["(?1 -> ?2)"] *) -val build_coq_arrow_pattern : constr_pattern delayed - (* ["(sig ?1 ?2)"] *) val build_coq_sig_pattern : constr_pattern delayed diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml index fb3039c0e..97cfd77b1 100644 --- a/parsing/egrammar.ml +++ b/parsing/egrammar.ml @@ -10,14 +10,21 @@ open Pp open Util +open Extend open Pcoq open Coqast open Ast -open Extend +open Genarg (* State of the grammar extensions *) -let (grammar_state : grammar_command list ref) = ref [] +type all_grammar_command = + | AstGrammar of grammar_command + | TacticGrammar of + (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) + list + +let (grammar_state : all_grammar_command list ref) = ref [] (* Interpretation of the right hand side of grammar rules *) @@ -38,8 +45,9 @@ let gram_action (name, etyp) env act dloc = try let v = Ast.eval_act dloc env act in match etyp, v with - | (ETast, Vast ast) -> Obj.repr ast - | (ETastl, Vastlist astl) -> Obj.repr astl + | (PureAstType, PureAstNode ast) -> Obj.repr ast + | (AstListType, AstListNode astl) -> Obj.repr astl + | (GenAstType ConstrArgType, PureAstNode ast) -> Obj.repr ast | _ -> grammar_type_error (dloc, "Egrammar.gram_action") with e -> let (loc, exn) = @@ -54,7 +62,7 @@ let gram_action (name, etyp) env act dloc = * is written (with camlp4 conventions): * (fun vi -> .... (fun v1 -> act(v1 .. vi) )..) * where v1..vi are the values generated by non-terminals nt1..nti. - * Since the actions are executed by substituing an environment, + * Since the actions are executed by substituting an environment, * make_act builds the following closure: * * ((fun env -> @@ -72,12 +80,12 @@ let gram_action (name, etyp) env act dloc = let make_act name_typ a pil = let act_without_arg env = Gramext.action (gram_action name_typ env a) and make_prod_item act_tl = function - | (None, _) -> (* parse a non-binding item *) + | None -> (* parse a non-binding item *) (fun env -> Gramext.action (fun _ -> act_tl env)) - | (Some p, ETast) -> (* non-terminal *) - (fun env -> Gramext.action (fun v -> act_tl ((p, Vast v) :: env))) - | (Some p, ETastl) -> (* non-terminal (list) *) - (fun env -> Gramext.action (fun v -> act_tl ((p, Vastlist v) :: env))) + | Some (p, ETast) -> (* non-terminal *) + (fun env -> Gramext.action (fun v -> act_tl((p,PureAstNode v)::env))) + | Some (p, ETastl) -> (* non-terminal *) + (fun env -> Gramext.action (fun v -> act_tl((p,AstListNode v)::env))) in (List.fold_left make_prod_item act_without_arg pil) [] @@ -88,16 +96,22 @@ let make_act name_typ a pil = * annotations are added when type-checking the command, function * Extend.of_ast) *) -let check_entry_type (u,n) typ = - match force_entry_type u n typ with - | Ast e -> Gram.Entry.obj e - | ListAst e -> Gram.Entry.obj e +let get_entry_type (u,n) = + Gram.Entry.obj (object_of_typed_entry (get_entry (get_univ u) n)) + +let rec build_prod_item univ = function + | ProdList0 s -> Gramext.Slist0 (build_prod_item univ s) + | ProdList1 s -> Gramext.Slist1 (build_prod_item univ s) + | ProdOpt s -> Gramext.Sopt (build_prod_item univ s) + | ProdPrimitive nt -> + let eobj = get_entry_type nt in + Gramext.Snterm eobj let symbol_of_prod_item univ = function - | Term tok -> (Gramext.Stoken tok, (None, ETast)) - | NonTerm (nt, nttyp, ovar) -> - let eobj = check_entry_type (qualified_nterm univ nt) nttyp in - (Gramext.Snterm eobj, (ovar, nttyp)) + | Term tok -> (Gramext.Stoken tok, None) + | NonTerm (nt, ovar) -> + let eobj = build_prod_item univ nt in + (eobj, ovar) let make_rule univ etyp rule = let pil = List.map (symbol_of_prod_item univ) rule.gr_production in @@ -113,6 +127,7 @@ let extend_entry univ (te, etyp, ass, rls) = (* Defines new entries. If the entry already exists, check its type *) let define_entry univ {ge_name=n; ge_type=typ; gl_assoc=ass; gl_rules=rls} = + let typ = if typ = ETast then GenAstType ConstrArgType else AstListType in let e = force_entry_type univ n typ in (e,typ,ass,rls) @@ -120,30 +135,99 @@ let define_entry univ {ge_name=n; ge_type=typ; gl_assoc=ass; gl_rules=rls} = let extend_grammar gram = let univ = get_univ gram.gc_univ in let tl = List.map (define_entry univ) gram.gc_entries in - List.iter (extend_entry univ) tl; - grammar_state := gram :: !grammar_state + List.iter (extend_entry univ) tl + +(* Add a grammar rules for tactics *) +type grammar_tactic_production = + | TacTerm of string + | TacNonTerm of loc * (Gramext.g_symbol * argument_type) * string option + +let make_prod_item = function + | TacTerm s -> (Gramext.Stoken (Extend.terminal s), None) + | TacNonTerm (_,(nont,t), po) -> + (nont, option_app (fun p -> (p,t)) po) + +let make_gen_act f pil = + let rec make env = function + | [] -> + Gramext.action (fun loc -> f env) + | None :: tl -> (* parse a non-binding item *) + Gramext.action (fun _ -> make env tl) + | Some (p, t) :: tl -> (* non-terminal *) + Gramext.action (fun v -> make ((p,in_generic t v) :: env) tl) in + make [] (List.rev pil) + +let make_rule univ f g (s',pt) = + let hd = Gramext.Stoken ("IDENT", s') in + let pil = (hd,None) :: List.map g pt in + let (symbs,ntl) = List.split pil in + let act = make_gen_act f ntl in + (symbs, act) + +(* These grammars are not a removable *) +let extend_tactic_grammar s gl = + let univ = get_univ "tactic" in + let make_act l = Tacexpr.TacExtend (s,List.map snd l) in + let rules = List.rev (List.map (make_rule univ make_act make_prod_item) gl) + in Gram.extend Tactic.simple_tactic None [(None, None, rules)] -(* backtrack the grammar extensions *) -let remove_rule univ e rule = - let symbs = - List.map (fun pi -> fst (symbol_of_prod_item univ pi)) rule.gr_production +let extend_vernac_command_grammar s gl = + let univ = get_univ "vernac" in + let make_act l = Vernacexpr.VernacExtend (s,List.map snd l) in + let rules = List.rev (List.map (make_rule univ make_act make_prod_item) gl) + in Gram.extend Vernac_.command None [(None, None, rules)] + +let rec interp_entry_name u s = + let l = String.length s in + if l > 8 & String.sub s 0 3 = "ne_" & String.sub s (l-5) 5 = "_list" then + let t, g = interp_entry_name u (String.sub s 3 (l-8)) in + List1ArgType t, Gramext.Slist1 g + else if l > 5 & String.sub s (l-5) 5 = "_list" then + let t, g = interp_entry_name u (String.sub s 0 (l-5)) in + List0ArgType t, Gramext.Slist0 g + else if l > 4 & String.sub s (l-4) 4 = "_opt" then + let t, g = interp_entry_name u (String.sub s 0 (l-4)) in + OptArgType t, Gramext.Sopt g + else + let n = Extend.rename_command s in + let e = get_entry (get_univ u) n in + let o = object_of_typed_entry e in + let t = match type_of_typed_entry e with + | GenAstType t -> t + | _ -> failwith "Only entries of generic type can be used in alias" in + t, Gramext.Snterm (Pcoq.Gram.Entry.obj o) + +let qualified_nterm current_univ = function + | NtQual (univ, en) -> (rename_command univ, rename_command en) + | NtShort en -> (current_univ, rename_command en) + +let make_vprod_item univ = function + | VTerm s -> (Gramext.Stoken (Extend.terminal s), None) + | VNonTerm (loc, nt, po) -> + let (u,nt) = qualified_nterm univ nt in + let (etyp, e) = interp_entry_name u nt in + e, option_app (fun p -> (p,etyp)) po + +let add_tactic_entries gl = + let univ = get_univ "tactic" in + let make_act s tac l = Tacexpr.TacAlias (s,l,tac) in + let rules = + List.rev (List.map (fun (s,l,tac) -> make_rule univ (make_act s tac) (make_vprod_item "tactic") l) gl) in - match e with - | Ast en -> Gram.delete_rule en symbs - | ListAst en -> Gram.delete_rule en symbs + let tacentry = get_entry (get_univ "tactic") "simple_tactic" in + grammar_extend tacentry None [(None, None, rules)] + +let extend_grammar gram = + (match gram with + | AstGrammar g -> extend_grammar g + | TacticGrammar l -> add_tactic_entries l); + grammar_state := gram :: !grammar_state -let remove_entry univ entry = - let e = get_entry univ entry.ge_name in - List.iter (remove_rule univ e) entry.gl_rules - -let remove_grammar gram = - let univ = get_univ gram.gc_univ in - List.iter (remove_entry univ) (List.rev gram.gc_entries) (* Summary functions: the state of the lexer is included in that of the parser. Because the grammar affects the set of keywords when adding or removing grammar rules. *) -type frozen_t = grammar_command list * Lexer.frozen_t +type frozen_t = all_grammar_command list * Lexer.frozen_t let freeze () = (!grammar_state, Lexer.freeze ()) @@ -153,20 +237,24 @@ let factorize_grams l1 l2 = if l1 == l2 then ([], [], l1) else list_share_tails l1 l2 let number_of_entries gcl = - List.fold_left (fun n gc -> n + (List.length gc.gc_entries)) 0 gcl + List.fold_left + (fun n -> function + | AstGrammar gc -> n + (List.length gc.gc_entries) + | TacticGrammar l -> n + 1) + 0 gcl let unfreeze (grams, lex) = let (undo, redo, common) = factorize_grams !grammar_state grams in - (* List.iter remove_grammar undo;*) remove_grammars (number_of_entries undo); grammar_state := common; Lexer.unfreeze lex; List.iter extend_grammar (List.rev redo) let init_grammar () = - List.iter remove_grammar !grammar_state; + remove_grammars (number_of_entries !grammar_state); grammar_state := [] +let _ = Lexer.init () + let init () = - Lexer.init (); init_grammar () diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli index 83dc3ce65..e7d2cde12 100644 --- a/parsing/egrammar.mli +++ b/parsing/egrammar.mli @@ -11,7 +11,8 @@ (*i*) open Coqast open Ast -open Pcoq +open Coqast +open Vernacexpr open Extend (*i*) @@ -21,9 +22,19 @@ val freeze : unit -> frozen_t val unfreeze : frozen_t -> unit val init : unit -> unit -val extend_grammar : grammar_command -> unit +type all_grammar_command = + | AstGrammar of grammar_command + | TacticGrammar of (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) list + +val extend_grammar : all_grammar_command -> unit + +(* Add grammar rules for tactics *) +type grammar_tactic_production = + | TacTerm of string + | TacNonTerm of loc * (Gramext.g_symbol * Genarg.argument_type) * string option + +val extend_tactic_grammar : + string -> (string * grammar_tactic_production list) list -> unit -val remove_rule : (string * gram_universe) -> typed_entry -> grammar_rule -> - unit -val remove_entry : (string * gram_universe) -> grammar_entry -> unit -val remove_grammar : grammar_command -> unit +val extend_vernac_command_grammar : + string -> (string * grammar_tactic_production list) list -> unit diff --git a/parsing/esyntax.ml b/parsing/esyntax.ml index 9d83c7878..608868ca6 100644 --- a/parsing/esyntax.ml +++ b/parsing/esyntax.ml @@ -13,6 +13,9 @@ open Util open Coqast open Ast open Extend +open Vernacexpr +open Names +open Nametab (*** Syntax keys ***) @@ -74,24 +77,34 @@ let spat_key astp = let se_key se = spat_key se.syn_astpat - (** Syntax entry tables (state of the pretty_printer) **) let from_name_table = ref Gmap.empty let from_key_table = ref Gmapl.empty +let infix_symbols_map = ref Stringmap.empty +let infix_names_map = ref Spmap.empty + (* Summary operations *) -type frozen_t = (string * string, syntax_entry) Gmap.t * - (string * key, syntax_entry) Gmapl.t +type frozen_t = (string * string, astpat syntax_entry) Gmap.t * + (string * key, astpat syntax_entry) Gmapl.t * + section_path Stringmap.t * string list Spmap.t -let freeze () = (!from_name_table, !from_key_table) +let freeze () = + (!from_name_table, !from_key_table, !infix_symbols_map, !infix_names_map) -let unfreeze (fnm,fkm) = +let unfreeze (fnm,fkm,infs,infn) = from_name_table := fnm; - from_key_table := fkm + from_key_table := fkm; + infix_symbols_map := infs; + infix_names_map := infn let init () = from_name_table := Gmap.empty; from_key_table := Gmapl.empty +(* + infix_symbols_map := Stringmap.empty; + infix_names_map := Spmap.empty +*) let find_syntax_entry whatfor gt = let gt_keys = ast_keys gt in @@ -105,7 +118,8 @@ let remove_with_warning name = if Gmap.mem name !from_name_table then begin let se = Gmap.find name !from_name_table in let key = (fst name, se_key se) in - (if !warning_verbose then + if !warning_verbose then + (Options.if_verbose warning ("overriding syntax rule "^(fst name)^":"^(snd name)^".")); from_name_table := Gmap.remove name !from_name_table; from_key_table := Gmapl.remove key se !from_key_table @@ -117,7 +131,7 @@ let add_rule whatfor se = remove_with_warning name; from_name_table := Gmap.add name se !from_name_table; from_key_table := Gmapl.add key se !from_key_table - + let add_ppobject {sc_univ=wf;sc_entries=sel} = List.iter (add_rule wf) sel @@ -126,6 +140,8 @@ let add_ppobject {sc_univ=wf;sc_entries=sel} = List.iter (add_rule wf) sel type std_printer = Coqast.t -> std_ppcmds type unparsing_subfunction = string -> tolerability option -> std_printer +type std_constr_printer = Genarg.constr_ast -> std_ppcmds + (* Module of primitive printers *) module Ppprim = struct @@ -136,10 +152,9 @@ module Ppprim = end (* A printer for the tokens. *) -let token_printer stdpr ast = - match ast with - | Id _ | Num _ | Str _ | Path _ -> print_ast ast - | _ -> stdpr ast +let token_printer stdpr = function + | (Id _ | Num _ | Str _ | Path _ as ast) -> print_ast ast + | a -> stdpr a (* Register the primitive printer for "token". It is not used in syntax/PP*.v, * but any ast matching no PP rule is printed with it. *) @@ -148,11 +163,56 @@ let _ = Ppprim.add ("token",token_printer) (* A primitive printer to do "print as" (to specify a length for a string) *) let print_as_printer stdpr = function - | Node (_, "AS", [Num(_,n); Str(_,s)]) -> (stras (n,s)) + | Node (_, "AS", [Num(_,n); Str(_,s)]) -> stras (n,s) | ast -> stdpr ast let _ = Ppprim.add ("print_as",print_as_printer) +(* Handle infix symbols *) + +let find_infix_symbols sp = + try Spmap.find sp !infix_names_map with Not_found -> [] + +let find_infix_name a = + try Stringmap.find a !infix_symbols_map + with Not_found -> anomaly ("Undeclared symbol: "^a) + +let declare_infix_symbol sp s = + let l = find_infix_symbols sp in + infix_names_map := Spmap.add sp (s::l) !infix_names_map; + infix_symbols_map := Stringmap.add s sp !infix_symbols_map + +let meta_pattern m = Pmeta(m,Tany) + +let make_hunks (lp,rp) s e1 e2 = + let n,s = + if is_letter (s.[String.length s -1]) or is_letter (s.[0]) + then 1,s^" " else 0,s + in + [PH (meta_pattern e1, None, lp); + UNP_BRK (n, 1); RO s; + PH (meta_pattern e2, None, rp)] + +let build_syntax (ref,e1,e2,assoc) = + let sp = match ref with + | TrueGlobal r -> Nametab.sp_of_global (Global.env()) r + | SyntacticDef sp -> sp in + let rec find_symbol = function + | [] -> + let s = match ref with + | TrueGlobal r -> + string_of_qualid (shortest_qualid_of_global (Global.env()) r) + | SyntacticDef sp -> string_of_path sp in + UNP_BOX (PpHOVB 0, + [RO "("; RO s; UNP_BRK (1, 1); PH (meta_pattern e1, None, E); + UNP_BRK (1, 1); PH (meta_pattern e2, None, E); RO ")"]) + | a::l -> + if find_infix_name a = sp then + UNP_BOX (PpHOVB 1, make_hunks assoc a e1 e2) + else + find_symbol l + in find_symbol (find_infix_symbols sp) + (* Print the syntax entry. In the unparsing hunks, the tokens are * printed using the token_printer, unless another primitive printer @@ -162,22 +222,21 @@ let print_syntax_entry whatfor sub_pr env se = let rule_prec = (se.syn_id, se.syn_prec) in let rec print_hunk = function | PH(e,externpr,reln) -> - let printer = - match externpr with (* May branch to an other printer *) - | Some c -> - (try (* Test for a primitive printer *) - (Ppprim.map c) (sub_pr whatfor (Some(rule_prec,reln))) - with Not_found -> - token_printer (sub_pr c (Some(rule_prec,reln)))) - | None -> token_printer (sub_pr whatfor (Some(rule_prec,reln))) - in - printer (Ast.pat_sub Ast.dummy_loc env e) + let node = Ast.pat_sub Ast.dummy_loc env e in + let printer = + match externpr with (* May branch to an other printer *) + | Some c -> + (try (* Test for a primitive printer *) Ppprim.map c + with Not_found -> token_printer ) + | _ -> token_printer in + printer (sub_pr whatfor (Some(rule_prec,reln))) node | RO s -> str s | UNP_TAB -> tab () | UNP_FNL -> fnl () | UNP_BRK(n1,n2) -> brk(n1,n2) | UNP_TBRK(n1,n2) -> tbrk(n1,n2) | UNP_BOX (b,sub) -> ppcmd_of_box b (prlist print_hunk sub) + | UNP_INFIX (sp,e1,e2,assoc) -> print_hunk (build_syntax (sp,e1,e2,assoc)) in prlist print_hunk se.syn_hunks diff --git a/parsing/esyntax.mli b/parsing/esyntax.mli index 40c541889..5ddadb5bc 100644 --- a/parsing/esyntax.mli +++ b/parsing/esyntax.mli @@ -11,6 +11,7 @@ (*i*) open Pp open Extend +open Vernacexpr (*i*) (* Syntax entry tables. *) @@ -24,9 +25,9 @@ val unfreeze : frozen_t -> unit (* Search and add a PP rule for an ast in the summary *) val find_syntax_entry : - string -> Coqast.t -> (syntax_entry * Ast.env) option -val add_rule : string -> syntax_entry -> unit -val add_ppobject : syntax_command -> unit + string -> Coqast.t -> (Ast.astpat syntax_entry * Ast.env) option +val add_rule : string -> Ast.astpat syntax_entry -> unit +val add_ppobject : Ast.astpat syntax_command -> unit val warning_verbose : bool ref (* Pretty-printing *) @@ -34,15 +35,17 @@ val warning_verbose : bool ref type std_printer = Coqast.t -> std_ppcmds type unparsing_subfunction = string -> tolerability option -> std_printer -(* Module of primitive printers *) +(* Module of constr primitive printers *) module Ppprim : sig type t = std_printer -> std_printer val add : string * t -> unit end +val declare_infix_symbol : Names.section_path -> string -> unit + (* Generic printing functions *) val token_printer: std_printer -> std_printer val print_syntax_entry : - string -> unparsing_subfunction -> Ast.env -> syntax_entry -> std_ppcmds + string -> unparsing_subfunction -> Ast.env -> Ast.astpat syntax_entry -> std_ppcmds val genprint : std_printer -> unparsing_subfunction diff --git a/parsing/extend.ml b/parsing/extend.ml new file mode 100644 index 000000000..bc4ad211b --- /dev/null +++ b/parsing/extend.ml @@ -0,0 +1,243 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* not (Lexer.is_keyword s) + | _ -> false + +let is_number s = + match s.[0] with + | '0'..'9' -> true + | _ -> false + +let strip s = + let len = + let rec loop i len = + if i = String.length s then len + else if s.[i] == ' ' then loop (i + 1) len + else loop (i + 1) (len + 1) + in + loop 0 0 + in + if len == String.length s then s + else + let s' = String.create len in + let rec loop i i' = + if i == String.length s then s' + else if s.[i] == ' ' then loop (i + 1) i' + else begin s'.[i'] <- s.[i]; loop (i + 1) (i' + 1) end + in + loop 0 0 + +let terminal s = + let s = strip s in + if s = "" then failwith "empty token"; + if is_ident_not_keyword s then ("IDENT", s) + else if is_number s then ("INT", s) + else ("", s) + +(*s Non-terminal symbols interpretation *) + +(* For compatibility *) +let warn nt nt' = + warning ("'"^nt^"' grammar entry is obsolete; use name '"^nt'^"' instead"); + nt' + +let rename_command nt = + if String.length nt >= 7 & String.sub nt 0 7 = "command" + then warn nt ("constr"^(String.sub nt 7 (String.length nt - 7))) + else if nt = "lcommand" then warn nt "lconstr" + else if nt = "lassoc_command4" then warn nt "lassoc_constr4" + else nt + +let qualified_nterm current_univ = function + | NtQual (univ, en) -> (rename_command univ, rename_command en) + | NtShort en -> (current_univ, rename_command en) + +let nterm loc (get_entry_type,univ) nont = + let nt = qualified_nterm univ nont in + try + let et = match get_entry_type nt with + | PureAstType -> ETast + | GenAstType Genarg.ConstrArgType -> ETast + | AstListType -> ETastl + | _ -> error "Cannot arbitrarily extend non ast entries" + in (nt,et) + with Not_found -> + let (s,n) = nt in + user_err_loc(loc,"Externd.nterm",str("unknown grammar entry: "^s^":"^n)) + +let prod_item univ env = function + | VTerm s -> env, Term (terminal s) + | VNonTerm (loc, nt, Some p) -> + let (nont, etyp) = nterm loc univ nt in + ((p, etyp) :: env, NonTerm (ProdPrimitive nont, Some (p,etyp))) + | VNonTerm (loc, nt, None) -> + let (nont, etyp) = nterm loc univ nt in + env, NonTerm (ProdPrimitive nont, None) + +let rec prod_item_list univ penv pil = + match pil with + | [] -> [], penv + | pi :: pitl -> + let (env, pic) = prod_item univ penv pi in + let (pictl, act_env) = prod_item_list univ env pitl in + (pic :: pictl, act_env) + +let gram_rule univ etyp (name,pil,act) = + let (pilc, act_env) = prod_item_list univ [] pil in + let a = Ast.to_act_check_vars act_env etyp act in + { gr_name = name; gr_production = pilc; gr_action = a } + +let gram_entry univ (nt, etyp, ass, rl) = + { ge_name = rename_command nt; + ge_type = etyp; + gl_assoc = ass; + gl_rules = List.map (gram_rule univ etyp) rl } + +let interp_grammar_command univ ge entryl = + let univ = rename_command univ in + { gc_univ = univ; + gc_entries = List.map (gram_entry (ge,univ)) entryl } + +(*s Pretty-print. *) + +(* Dealing with precedences *) + +type precedence = int * int * int + +type parenRelation = L | E | Any | Prec of precedence + +type ppbox = + | PpHB of int + | PpHOVB of int + | PpHVB of int + | PpVB of int + | PpTB + +(* unparsing objects *) + +type 'pat unparsing_hunk = + | PH of 'pat * string option * parenRelation + | RO of string + | UNP_BOX of ppbox * 'pat unparsing_hunk list + | UNP_BRK of int * int + | UNP_TBRK of int * int + | UNP_TAB + | UNP_FNL + | UNP_INFIX of Nametab.extended_global_reference * string * string * + (parenRelation * parenRelation) + +(* Checks if the precedence of the parent printer (None means the + highest precedence), and the child's one, follow the given + relation. *) + +type tolerability = (string * precedence) * parenRelation + +let compare_prec (a1,b1,c1) (a2,b2,c2) = + match (a1=a2),(b1=b2),(c1=c2) with + | true,true,true -> 0 + | true,true,false -> c1-c2 + | true,false,_ -> b1-b2 + | false,_,_ -> a1-a2 + +let tolerable_prec oparent_prec_reln (_,child_prec) = + match oparent_prec_reln with + | Some ((_,pprec), L) -> (compare_prec child_prec pprec) < 0 + | Some ((_,pprec), E) -> (compare_prec child_prec pprec) <= 0 + | Some (_, Prec level) -> (compare_prec child_prec level) <= 0 + | _ -> true + +let ppcmd_of_box = function + | PpHB n -> h n + | PpHOVB n -> hov n + | PpHVB n -> hv n + | PpVB n -> v n + | PpTB -> t + +type 'pat syntax_entry = { + syn_id : string; + syn_prec: precedence; + syn_astpat : 'pat; + syn_hunks : 'pat unparsing_hunk list } + +type 'pat syntax_command = { + sc_univ : string; + sc_entries : 'pat syntax_entry list } + +type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list +type syntax_entry_ast = precedence * syntax_rule list + +let rec interp_unparsing env = function + | PH (ast,ext,pr) -> PH (Ast.val_of_ast env ast,ext,pr) + | UNP_BOX (b,ul) -> UNP_BOX (b,List.map (interp_unparsing env) ul) + | UNP_BRK _ | RO _ | UNP_TBRK _ | UNP_TAB | UNP_FNL + | UNP_INFIX _ as x -> x + +let rule_of_ast univ prec (s,spat,unp) = + let (astpat,meta_env) = Ast.to_pat [] spat in + let hunks = List.map (interp_unparsing meta_env) unp in + { syn_id = s; + syn_prec = prec; + syn_astpat = astpat; + syn_hunks = hunks } + +let level_of_ast univ (prec,rl) = List.map (rule_of_ast univ prec) rl + +let interp_syntax_entry univ sel = + { sc_univ = univ; + sc_entries = List.flatten (List.map (level_of_ast univ) sel)} + + diff --git a/parsing/extend.mli b/parsing/extend.mli index da77f531d..e80f011d3 100644 --- a/parsing/extend.mli +++ b/parsing/extend.mli @@ -1,39 +1,25 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* nonterm -> - (string * gram_universe) * string +type nonterm_prod = + | ProdList0 of nonterm_prod + | ProdList1 of nonterm_prod + | ProdOpt of nonterm_prod + | ProdPrimitive of (string * string) type prod_item = | Term of Token.pattern - | NonTerm of nonterm * entry_type * string option + | NonTerm of nonterm_prod * (string * ast_action_type) option -type grammar_rule = { +type grammar_rule = { gr_name : string; gr_production : prod_item list; gr_action : Ast.act } type grammar_entry = { ge_name : string; - ge_type : entry_type; + ge_type : ast_action_type; gl_assoc : Gramext.g_assoc option; gl_rules : grammar_rule list } @@ -41,27 +27,28 @@ type grammar_command = { gc_univ : string; gc_entries : grammar_entry list } -val gram_assoc : Coqast.t -> Gramext.g_assoc option +type grammar_associativity = Gramext.g_assoc option +type nonterm = + | NtShort of string + | NtQual of string * string +type grammar_production = + | VTerm of string + | VNonTerm of loc * nonterm * string option +type raw_grammar_rule = string * grammar_production list * grammar_action +type raw_grammar_entry = + string * ast_action_type * grammar_associativity * raw_grammar_rule list -val interp_grammar_command : string -> Coqast.t list -> grammar_command +val terminal : string -> string * string +val rename_command : string -> string (*s Pretty-print. *) (* Dealing with precedences *) type precedence = int * int * int -type parenRelation = L | E | Any | Prec of precedence - -(* Checks if the precedence of the parent printer (None means the - highest precedence), and the child's one, follow the given - relation. *) - -type tolerability = (string * precedence) * parenRelation - -val tolerable_prec : tolerability option -> (string * precedence) -> bool -(* unparsing objects *) +type parenRelation = L | E | Any | Prec of precedence type ppbox = | PpHB of int @@ -70,29 +57,45 @@ type ppbox = | PpVB of int | PpTB -type unparsing_hunk = - | PH of Ast.pat * string option * parenRelation +(* unparsing objects *) + +type 'pat unparsing_hunk = + | PH of 'pat * string option * parenRelation | RO of string - | UNP_BOX of ppbox * unparsing_hunk list + | UNP_BOX of ppbox * 'pat unparsing_hunk list | UNP_BRK of int * int | UNP_TBRK of int * int | UNP_TAB | UNP_FNL + | UNP_INFIX of Nametab.extended_global_reference * string * string * + (parenRelation * parenRelation) -val ppcmd_of_box : ppbox -> std_ppcmds -> std_ppcmds +(* Checks if the precedence of the parent printer (None means the + highest precedence), and the child's one, follow the given + relation. *) + +type tolerability = (string * precedence) * parenRelation + +val tolerable_prec : tolerability option -> (string * precedence) -> bool -val unparsing_of_ast : Ast.entry_env -> Coqast.t -> unparsing_hunk list +val ppcmd_of_box : ppbox -> std_ppcmds -> std_ppcmds -type syntax_entry = { +type 'pat syntax_entry = { syn_id : string; syn_prec: precedence; - syn_astpat : Ast.pat; - syn_hunks : unparsing_hunk list } + syn_astpat : 'pat; + syn_hunks : 'pat unparsing_hunk list } -type syntax_command = { +type 'pat syntax_command = { sc_univ : string; - sc_entries : syntax_entry list } + sc_entries : 'pat syntax_entry list } -val interp_syntax_entry : string -> Coqast.t list -> syntax_command +type syntax_rule = string * Coqast.t * Coqast.t unparsing_hunk list +type syntax_entry_ast = precedence * syntax_rule list +val interp_grammar_command : + string -> (string * string -> entry_type) -> + raw_grammar_entry list -> grammar_command +val interp_syntax_entry : + string -> syntax_entry_ast list -> Ast.astpat syntax_command diff --git a/parsing/g_basevernac.ml4 b/parsing/g_basevernac.ml4 index 6dbaabc1c..6fdabef5a 100644 --- a/parsing/g_basevernac.ml4 +++ b/parsing/g_basevernac.ml4 @@ -11,360 +11,335 @@ (* $Id$ *) open Coqast +open Extend +open Vernacexpr open Pcoq - open Vernac_ +open Goptions +open Constr +open Prim GEXTEND Gram - GLOBAL: identarg ne_identarg_list numarg ne_numarg_list numarg_list - stringarg ne_stringarg_list constrarg sortarg tacarg - ne_qualidarg_list qualidarg commentarg - commentarg_list; + GLOBAL: class_rawexpr; - identarg: - [ [ id = Constr.ident -> id ] ] - ; - ne_identarg_list: - [ [ l = LIST1 identarg -> l ] ] - ; - qualidarg: - [ [ l = Constr.qualid -> <:ast< (QUALIDARG ($LIST $l)) >> ] ] - ; - ne_qualidarg_list: - [ [ q = qualidarg; l = ne_qualidarg_list -> q::l - | q = qualidarg -> [q] ] ] - ; - numarg: - [ [ n = Prim.number -> n - | "-"; n = Prim.number -> Num (loc, ( - Ast.num_of_ast n)) ] ] - ; - ne_numarg_list: - [ [ n = numarg; l = ne_numarg_list -> n::l - | n = numarg -> [n] ] ] - ; - numarg_list: - [ [ l = ne_numarg_list -> l - | -> [] ] ] - ; - stringarg: - [ [ s = Prim.string -> s ] ] - ; - ne_stringarg_list: - [ [ n = stringarg; l = ne_stringarg_list -> n::l - | n = stringarg -> [n] ] ] - ; - constrarg: - [ [ c = Constr.constr -> <:ast< (CONSTR $c) >> ] ] - ; - sortarg: - [ [ c = Constr.sort -> <:ast< (CONSTR $c) >> ] ] - ; - tacarg: - [ [ tcl = Tactic.tactic -> tcl ] ] - ; - commentarg: - [ [ c = constrarg -> c - | c = stringarg -> c - | c = numarg -> c ] ] - ; - commentarg_list: - [ [ c = commentarg; l = commentarg_list -> c::l - | -> [] ] ] + class_rawexpr: + [ [ IDENT "FUNCLASS" -> FunClass + | IDENT "SORTCLASS" -> SortClass + | qid = Prim.qualid -> RefClass qid ] ] ; END; GEXTEND Gram GLOBAL: command; + comment: + [ [ c = constr -> CommentConstr c + | s = STRING -> CommentString s + | n = natural -> CommentInt n ] ] + ; command: - [ [ IDENT "Comments"; args = commentarg_list -> - <:ast< (Comments ($LIST $args)) >> - | IDENT "Pwd" -> <:ast< (PWD) >> - | IDENT "Cd" -> <:ast< (CD) >> - | IDENT "Cd"; dir = stringarg -> <:ast< (CD $dir) >> - - | IDENT "Drop" -> <:ast< (DROP) >> - | IDENT "ProtectedLoop" -> <:ast< (PROTECTEDLOOP)>> - | "Quit" -> <:ast< (QUIT) >> - - | IDENT "Print"; IDENT "All" -> <:ast< (PrintAll) >> - | IDENT "Print" -> <:ast< (PRINT) >> - | IDENT "Print"; IDENT "Hint"; "*" - -> <:ast< (PrintHint) >> - | IDENT "Print"; IDENT "Hint"; s = qualidarg -> - <:ast< (PrintHintId $s) >> - | IDENT "Print"; IDENT "Hint" -> - <:ast< (PrintHintGoal) >> - | IDENT "Print"; IDENT "HintDb"; s = identarg -> - <:ast< (PrintHintDb $s) >> - | IDENT "Print"; IDENT "Section"; s = qualidarg -> - <:ast< (PrintSec $s) >> - (* This should be in "syntax" section but is here for factorization *) - | IDENT "Print"; "Grammar"; uni = identarg; ent = identarg -> - <:ast< (PrintGrammar $uni $ent) >> + [ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l + + (* System directory *) + | IDENT "Pwd" -> VernacChdir None + | IDENT "Cd" -> VernacChdir None + | IDENT "Cd"; dir = STRING -> VernacChdir (Some dir) + + (* Toplevel control *) + | IDENT "Drop" -> VernacToplevelControl Drop + | IDENT "ProtectedLoop" -> VernacToplevelControl ProtectedLoop + | "Quit" -> VernacToplevelControl Quit (* Dump of the universe graph - to file or to stdout *) - | IDENT "Dump"; IDENT "Universes"; f = stringarg -> - <:ast< (DumpUniverses $f) >> - | IDENT "Dump"; IDENT "Universes" -> - <:ast< (DumpUniverses) >> - - | IDENT "Locate"; IDENT "File"; f = stringarg -> - <:ast< (LocateFile $f) >> - | IDENT "Locate"; IDENT "Library"; id = qualidarg -> - <:ast< (LocateLibrary $id) >> - | IDENT "Locate"; id = qualidarg -> - <:ast< (Locate $id) >> + | IDENT "Dump"; IDENT "Universes"; fopt = OPT STRING -> + VernacPrint (PrintUniverses fopt) + + | IDENT "Locate"; qid = qualid -> VernacLocate (LocateTerm qid) + | IDENT "Locate"; IDENT "File"; f = STRING -> VernacLocate (LocateFile f) + | IDENT "Locate"; IDENT "Library"; qid = qualid -> + VernacLocate (LocateLibrary qid) (* Managing load paths *) - | IDENT "Add"; IDENT "LoadPath"; dir = stringarg; - "as"; alias = qualidarg -> <:ast< (ADDPATH $dir $alias) >> - | IDENT "Add"; IDENT "LoadPath"; dir = stringarg -> - <:ast< (ADDPATH $dir) >> - | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = stringarg; - "as"; alias=qualidarg -> <:ast< (RECADDPATH $dir $alias) >> - | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = stringarg -> - <:ast< (RECADDPATH $dir) >> - | IDENT "Remove"; IDENT "LoadPath"; dir = stringarg -> - <:ast< (DELPATH $dir) >> - | IDENT "Print"; IDENT "LoadPath" -> <:ast< (PrintPath) >> + | IDENT "Add"; IDENT "LoadPath"; dir = STRING; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = STRING; + alias = as_dirpath -> VernacAddLoadPath (true, dir, alias) + | IDENT "Remove"; IDENT "LoadPath"; dir = STRING -> + VernacRemoveLoadPath dir (* For compatibility *) - | IDENT "AddPath"; dir = stringarg; "as"; alias = qualidarg -> - <:ast< (ADDPATH $dir $alias) >> - | IDENT "AddPath"; dir = stringarg -> <:ast< (ADDPATH $dir) >> - | IDENT "AddRecPath"; dir = stringarg; "as"; alias=qualidarg -> - <:ast< (RECADDPATH $dir $alias) >> - | IDENT "AddRecPath"; dir = stringarg -> - <:ast< (RECADDPATH $dir) >> - | IDENT "DelPath"; dir = stringarg -> <:ast< (DELPATH $dir) >> - - | IDENT "Print"; IDENT "Modules" -> <:ast< (PrintModules) >> - | IDENT "Print"; "Proof"; id = qualidarg -> - <:ast< (PrintOpaqueId $id) >> -(* Pris en compte dans PrintOption ci-dessous (CADUC) *) - | IDENT "Print"; id = qualidarg -> <:ast< (PrintId $id) >> - | IDENT "Search"; id = qualidarg; l = in_or_out_modules -> - <:ast< (SEARCH $id ($LIST $l)) >> - | IDENT "Inspect"; n = numarg -> <:ast< (INSPECT $n) >> - | IDENT "SearchPattern"; c = constrarg; l = in_or_out_modules -> - <:ast< (SearchPattern $c ($LIST $l)) >> - | IDENT "SearchRewrite"; c = constrarg; l = in_or_out_modules -> - <:ast< (SearchRewrite $c ($LIST $l)) >> + | IDENT "AddPath"; dir = STRING; "as"; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "AddRecPath"; dir = STRING; "as"; alias = as_dirpath -> + VernacAddLoadPath (true, dir, alias) + | IDENT "DelPath"; dir = STRING -> + VernacRemoveLoadPath dir + + (* Printing (careful factorization of entries) *) + | IDENT "Print"; p = printable -> VernacPrint p + | IDENT "Print"; qid = qualid -> VernacPrint (PrintName qid) + | IDENT "Print" -> VernacPrint PrintLocalContext + | IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n) + + (* Searching the environment *) + | IDENT "Search"; qid = Prim.qualid; l = in_or_out_modules -> + VernacSearch (SearchHead qid, l) + | IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules -> + VernacSearch (SearchPattern c, l) + | IDENT "SearchRewrite"; c = constr_pattern; l = in_or_out_modules -> + VernacSearch (SearchRewrite c, l) + (* TODO: rapprocher Eval et Check *) - | IDENT "Eval"; r = Tactic.red_tactic; "in"; c = constrarg -> - <:ast< (Eval (TACTIC_ARG (REDEXP $r)) $c) >> - | IDENT "Eval"; g = numarg; r = Tactic.red_tactic; - "in"; c = constrarg -> - <:ast< (Eval (TACTIC_ARG (REDEXP $r)) $c $g) >> - | check = check_tok; c = constrarg -> <:ast< (Check $check $c) >> - | check = check_tok; g = numarg; c = constrarg -> - <:ast< (Check $check $c $g) >> - | IDENT "Print"; IDENT "ML"; IDENT "Path" -> - <:ast< (PrintMLPath) >> - | IDENT "Print"; IDENT "ML"; IDENT "Modules" -> - <:ast< (PrintMLModules) >> - | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = stringarg -> - <:ast< (AddMLPath $dir) >> - | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; - dir = stringarg -> - <:ast< (RecAddMLPath $dir) >> - | IDENT "Print"; IDENT "Graph" -> <:ast< (PrintGRAPH) >> - | IDENT "Print"; IDENT "Classes" -> <:ast< (PrintCLASSES) >> - | IDENT "Print"; IDENT "Coercions" -> <:ast< (PrintCOERCIONS) >> - | IDENT "Print"; IDENT "Coercion"; IDENT "Paths"; - c = qualidarg; d = qualidarg -> - <:ast< (PrintPATH $c $d) >> - - | IDENT "SearchIsos"; com = constrarg -> - <:ast< (Searchisos $com) >> - | "Set"; IDENT "Undo"; n = numarg -> - <:ast< (SETUNDO $n) >> - | IDENT "Unset"; IDENT "Undo" -> <:ast< (UNSETUNDO) >> - | "Set"; IDENT "Hyps_limit"; n = numarg -> - <:ast< (SETHYPSLIMIT $n) >> - | IDENT "Unset"; IDENT "Hyps_limit" -> - <:ast< (UNSETHYPSLIMIT) >> + | IDENT "Eval"; g = OPT natural; r = Tactic.red_tactic; "in"; + c = constr -> VernacCheckMayEval (Some r, g, c) + | IDENT "Check"; g = OPT natural; c = constr -> + VernacCheckMayEval (None, g, c) + | "Type"; c = constr -> VernacGlobalCheck c (* pas dans le RefMan *) + + | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = STRING -> + VernacAddMLPath (false, dir) + | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = STRING -> + VernacAddMLPath (true, dir) +(* + | IDENT "SearchIsos"; c = constr -> VernacSearch (SearchIsos c) +*) (* Pour intervenir sur les tables de paramtres *) - | "Set"; table = identarg; field = identarg; - value = option_value -> - <:ast< (SetTableField $table $field $value) >> - | "Set"; table = identarg; field = identarg -> - <:ast< (SetTableField $table $field) >> - | IDENT "Unset"; table = identarg; field = identarg -> - <:ast< (UnsetTableField $table $field) >> - | IDENT "Unset"; table = identarg; field = identarg; id = qualidarg -> - <:ast< (RemoveTableField $table $field $id) >> - | "Set"; table = identarg; value = option_value -> - <:ast< (SetTableField $table $value) >> - | "Set"; table = identarg -> - <:ast< (SetTableField $table) >> - | IDENT "Unset"; table = identarg -> - <:ast< (UnsetTableField $table) >> - | IDENT "Print"; IDENT "Table"; - table = identarg; field = identarg -> - <:ast< (PrintOption $table $field) >> - (* Le cas suivant inclut aussi le "Print id" standard *) - | IDENT "Print"; IDENT "Table"; table = identarg -> - <:ast< (PrintOption $table) >> - | IDENT "Add"; table = identarg; field = identarg; id = qualidarg - -> <:ast< (AddTableField $table $field $id) >> - | IDENT "Add"; table = identarg; field = identarg; id = stringarg - -> <:ast< (AddTableField $table $field $id) >> - | IDENT "Add"; table = identarg; id = qualidarg - -> <:ast< (AddTableField $table $id) >> - | IDENT "Add"; table = identarg; id = stringarg - -> <:ast< (AddTableField $table $id) >> - | IDENT "Test"; table = identarg; field = identarg; id = qualidarg - -> <:ast< (MemTableField $table $field $id) >> - | IDENT "Test"; table = identarg; field = identarg; id = stringarg - -> <:ast< (MemTableField $table $field $id) >> - | IDENT "Test"; table = identarg; id = identarg - -> <:ast< (MemTableField $table $id) >> - | IDENT "Test"; table = identarg; id = qualidarg - -> <:ast< (MemTableField $table $id) >> - | IDENT "Test"; table = identarg; id = stringarg - -> <:ast< (MemTableField $table $id) >> - | IDENT "Test"; table = identarg - -> <:ast< (MemTableField $table) >> - | IDENT "Remove"; table = identarg; field = identarg; id = qualidarg -> - <:ast< (RemoveTableField $table $field $id) >> - | IDENT "Remove"; table = identarg; field = identarg; id = stringarg -> - <:ast< (RemoveTableField $table $field $id) >> - | IDENT "Remove"; table = identarg; id = qualidarg -> - <:ast< (RemoveTableField $table $id) >> - | IDENT "Remove"; table = identarg; id = stringarg -> - <:ast< (RemoveTableField $table $id) >> ] ] + + | "Set"; table = IDENT; field = IDENT; v = option_value -> + VernacSetOption (SecondaryTable (table,field),v) + | "Set"; table = IDENT; field = IDENT; lv = LIST1 option_ref_value -> + VernacAddOption (SecondaryTable (table,field),lv) + | "Set"; table = IDENT; field = IDENT -> + VernacSetOption (SecondaryTable (table,field),BoolValue true) + | IDENT "Unset"; table = IDENT; field = IDENT -> + VernacUnsetOption (SecondaryTable (table,field)) + | IDENT "Unset"; table = IDENT; field = IDENT; lv = LIST1 option_ref_value -> + VernacRemoveOption (SecondaryTable (table,field),lv) + | "Set"; table = IDENT; value = option_value -> + VernacSetOption (PrimaryTable table, value) + | "Set"; table = IDENT -> + VernacSetOption (PrimaryTable table, BoolValue true) + | IDENT "Unset"; table = IDENT -> + VernacUnsetOption (PrimaryTable table) + + | IDENT "Print"; IDENT "Table"; table = IDENT; field = IDENT -> + VernacPrintOption (SecondaryTable (table,field)) + | IDENT "Print"; IDENT "Table"; table = IDENT -> + VernacPrintOption (PrimaryTable table) + + | IDENT "Add"; table = IDENT; field = IDENT; v = LIST1 option_ref_value + -> VernacAddOption (SecondaryTable (table,field), v) + + (* Un value qualid ci-dessous va tre cach par un field au dessus! *) + | IDENT "Add"; table = IDENT; v = LIST1 option_ref_value -> + VernacAddOption (PrimaryTable table, v) + + | IDENT "Test"; table = IDENT; field = IDENT; v = LIST1 option_ref_value + -> VernacMemOption (SecondaryTable (table,field), v) + | IDENT "Test"; table = IDENT; field = IDENT -> + VernacPrintOption (SecondaryTable (table,field)) + | IDENT "Test"; table = IDENT; v = LIST1 option_ref_value -> + VernacMemOption (PrimaryTable table, v) + | IDENT "Test"; table = IDENT -> + VernacPrintOption (PrimaryTable table) + + | IDENT "Remove"; table = IDENT; field = IDENT; v= LIST1 option_ref_value + -> VernacRemoveOption (SecondaryTable (table,field), v) + | IDENT "Remove"; table = IDENT; v = LIST1 option_ref_value -> + VernacRemoveOption (PrimaryTable table, v) ] ] + ; + printable: + [ [ IDENT "All" -> PrintFullContext + | IDENT "Section"; s = qualid -> PrintSectionContext s + | "Grammar"; uni = IDENT; ent = IDENT -> + (* This should be in "syntax" section but is here for factorization*) + PrintGrammar uni ent + | IDENT "LoadPath" -> PrintLoadPath + | IDENT "Modules" -> PrintModules + + | IDENT "ML"; IDENT "Path" -> PrintMLLoadPath + | IDENT "ML"; IDENT "Modules" -> PrintMLModules + | IDENT "Graph" -> PrintGraph + | IDENT "Classes" -> PrintClasses + | IDENT "Coercions" -> PrintCoercions + | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr + -> PrintCoercionPaths (s,t) + | IDENT "Tables" -> PrintTables + | "Proof"; qid = qualid -> PrintOpaqueName qid + | IDENT "Hint" -> PrintHintGoal + | IDENT "Hint"; qid = qualid -> PrintHint qid + | IDENT "Hint"; "*" -> PrintHintDb + | IDENT "HintDb"; s = IDENT -> PrintHintDbName s ] ] ; option_value: - [ [ id = qualidarg -> id - | n = numarg -> n - | s = stringarg -> s ] ] + [ [ n = integer -> IntValue n + | s = STRING -> StringValue s ] ] ; - check_tok: - [ [ IDENT "Check" -> <:ast< "CHECK" >> - | "Type" -> <:ast< "PRINTTYPE" >> ] ] (* pas dans le RefMan *) + option_ref_value: + [ [ id = qualid -> QualidRefValue id + | s = STRING -> StringRefValue s ] ] + ; + as_dirpath: + [ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ] ; in_or_out_modules: - [ [ IDENT "inside"; l = ne_qualidarg_list -> <:ast< "inside" >> :: l - | IDENT "outside"; l = ne_qualidarg_list -> <:ast< "outside" >> :: l - | -> [] ] ] + [ [ IDENT "inside"; l = LIST1 qualid -> SearchInside l + | IDENT "outside"; l = LIST1 qualid -> SearchOutside l + | -> SearchOutside [] ] ] ; END; (* Grammar extensions *) GEXTEND Gram - GLOBAL: syntax Prim.syntax_entry Prim.grammar_entry; + GLOBAL: syntax; univ: [ [ univ = IDENT -> - let _ = set_default_action_parser_by_name univ in univ ] ] + set_default_action_parser (parser_type_from_name univ); univ ] ] ; syntax: - [ [ IDENT "Token"; s = STRING -> <:ast< (TOKEN ($STR $s)) >> + [ [ IDENT "Token"; s = STRING -> + Pp.warning "Token declarations are now useless"; VernacNop + + | "Grammar"; IDENT "tactic"; IDENT "simple_tactic"; ":="; + OPT "|"; tl = LIST0 grammar_tactic_rule SEP "|" -> + VernacTacticGrammar tl - | "Grammar"; univ = univ; - tl = LIST1 Prim.grammar_entry SEP "with" -> - <:ast< (GRAMMAR { $univ } (ASTLIST ($LIST $tl))) >> + | "Grammar"; u = univ; + tl = LIST1 grammar_entry SEP "with" -> VernacGrammar (u,tl) - | "Syntax"; univ = univ; el=LIST1 Prim.syntax_entry SEP ";" -> - <:ast< (SYNTAX { $univ } (ASTLIST ($LIST $el))) >> + | "Syntax"; u = univ; el = LIST1 syntax_entry SEP ";" -> + VernacSyntax (u,el) (* Faudrait une version de qualidarg dans Prim pour respecter l'ordre *) - | IDENT "Infix"; as_ = entry_prec; n = numarg; op = Prim.string; - p = qualidarg -> <:ast< (INFIX (AST $as_) $n $op $p) >> - | IDENT "Distfix"; as_ = entry_prec; n = numarg; s = Prim.string; - p = qualidarg -> <:ast< (DISTFIX (AST $as_) $n $s $p) >> + | IDENT "Infix"; a = entry_prec; n = natural; op = STRING; p = qualid + -> VernacInfix (a,n,op,p) + | IDENT "Distfix"; a = entry_prec; n = natural; s = STRING; p = qualid -> VernacDistfix (a,n,s,p) + (* "Print" "Grammar" should be here but is in "command" entry in order to factorize with other "Print"-based vernac entries *) ] ] ; - (* Syntax entries for Grammar. Only grammar_entry is exported *) - Prim.grammar_entry: - [[ nont = Prim.ident; etyp = Prim.entry_type; ":="; + grammar_entry: + [[ nont = located_ident; etyp = set_entry_type; ":="; ep = entry_prec; OPT "|"; rl = LIST0 grammar_rule SEP "|" -> - <:ast< (GRAMMARENTRY $nont $etyp $ep ($LIST $rl)) >> ]] + (nont,etyp,ep,rl) ]] + ; + located_ident: + [[ id = IDENT -> (loc,id) ]] ; entry_prec: - [[ IDENT "LEFTA" -> <:ast< (LEFTA) >> - | IDENT "RIGHTA" -> <:ast< (RIGHTA) >> - | IDENT "NONA" -> <:ast< (NONA) >> - | -> <:ast< (NONE) >> ] ] + [[ IDENT "LEFTA" -> Some Gramext.LeftA + | IDENT "RIGHTA" -> Some Gramext.RightA + | IDENT "NONA" -> Some Gramext.NonA + | -> None ]] + ; + grammar_tactic_rule: + [[ name = rule_name; "["; s = STRING; pil = LIST0 production_item; "]"; + "->"; "["; t = Tactic.tactic; "]" -> (name, (s,pil), t) ]] ; grammar_rule: [[ name = rule_name; "["; pil = LIST0 production_item; "]"; "->"; - a = Prim.astact -> - <:ast< (GRAMMARRULE ($ID $name) $a ($LIST $pil)) >> ]] + a = action -> (name, pil, a) ]] ; rule_name: [[ name = IDENT -> name ]] ; production_item: - [[ s = STRING -> <:ast< ($STR $s) >> + [[ s = STRING -> VTerm s | nt = non_terminal; po = OPT [ "("; p = Prim.metaident; ")" -> p ] -> match po with - | Some p -> <:ast< (NT $nt $p) >> - | _ -> <:ast< (NT $nt) >> ]] + | Some p -> VNonTerm (loc,nt,Some (Ast.meta_of_ast p)) + | _ -> VNonTerm (loc,nt,None) ]] ; non_terminal: - [[ u = Prim.ident; ":"; nt = Prim.ident -> <:ast< (QUAL $u $nt) >> - | nt = Prim.ident -> <:ast< $nt >> ]] + [[ u = IDENT; ":"; nt = IDENT -> NtQual(u, nt) + | nt = IDENT -> NtShort nt ]] ; (* Syntax entries for Syntax. Only syntax_entry is exported *) - Prim.syntax_entry: + syntax_entry: [ [ IDENT "level"; p = precedence; ":"; - OPT "|"; rl = LIST1 syntax_rule SEP "|" -> - <:ast< (SYNTAXENTRY $p ($LIST $rl)) >> ] ] + OPT "|"; rl = LIST1 syntax_rule SEP "|" -> (p,rl) ] ] ; syntax_rule: - [ [ nm = IDENT; "["; s = Prim.astpat; "]"; "->"; u = unparsing -> - <:ast< (SYNTAXRULE ($ID $nm) $s $u) >> ] ] + [ [ nm = IDENT; "["; s = astpat; "]"; "->"; u = unparsing -> (nm,s,u) ] ] ; precedence: - [ [ a = Prim.number -> <:ast< (PREC $a 0 0) >> - | "["; a1 = Prim.number; a2 = Prim.number; a3 = Prim.number; "]" -> - <:ast< (PREC $a1 $a2 $a3) >> ] ] + [ [ a = natural -> (a,0,0) + | "["; a1 = natural; a2 = natural; a3 = natural; "]" -> (a1,a2,a3) ] ] ; unparsing: - [ [ "["; ll = LIST0 next_hunks; "]" -> <:ast< (UNPARSING ($LIST $ll))>> ] ] + [ [ "["; ll = LIST0 next_hunks; "]" -> ll ] ] ; next_hunks: - [ [ IDENT "FNL" -> <:ast< (UNP_FNL) >> - | IDENT "TAB" -> <:ast< (UNP_TAB) >> - | c = STRING -> <:ast< (RO ($STR $c)) >> + [ [ IDENT "FNL" -> UNP_FNL + | IDENT "TAB" -> UNP_TAB + | c = STRING -> RO c | "["; x = - [ b = box; ll = LIST0 next_hunks -> <:ast<(UNP_BOX $b ($LIST $ll))>> - | n = Prim.number; m = Prim.number -> <:ast< (UNP_BRK $n $m) >> - | IDENT "TBRK"; - n = Prim.number; m = Prim.number -> <:ast< (UNP_TBRK $n $m) >> ]; + [ b = box; ll = LIST0 next_hunks -> UNP_BOX (b,ll) + | n = natural; m = natural -> UNP_BRK (n, m) + | IDENT "TBRK"; n = natural; m = natural -> UNP_TBRK (n, m) ]; "]" -> x | e = Prim.ast; oprec = OPT [ ":"; pr = paren_reln_or_extern -> pr ] -> match oprec with - | Some pr -> <:ast< (PH $e $pr) >> - | None -> <:ast< (PH $e (Any)) >> ]] + | Some (ext,pr) -> PH (e,ext,pr) + | None -> PH (e,None,Any) ]] ; box: [ [ "<"; bk = box_kind; ">" -> bk ] ] ; box_kind: - [ [ IDENT "h"; n = Prim.number -> <:ast< (PpHB $n) >> - | IDENT "v"; n = Prim.number -> <:ast< (PpVB $n) >> - | IDENT "hv"; n = Prim.number -> <:ast< (PpHVB $n) >> - | IDENT "hov"; n = Prim.number -> <:ast< (PpHOVB $n) >> - | IDENT "t" -> <:ast< (PpTB) >> ] ] + [ [ IDENT "h"; n = natural -> PpHB n + | IDENT "v"; n = natural -> PpVB n + | IDENT "hv"; n = natural -> PpHVB n + | IDENT "hov"; n = natural -> PpHOVB n + | IDENT "t" -> PpTB ] ] ; paren_reln_or_extern: - [ [ IDENT "L" -> <:ast< (L) >> - | IDENT "E" -> <:ast< (E) >> + [ [ IDENT "L" -> None, L + | IDENT "E" -> None, E | pprim = STRING; precrec = OPT [ ":"; p = precedence -> p ] -> match precrec with - | Some p -> <:ast< (EXTERN ($STR $pprim) $p) >> - | None -> <:ast< (EXTERN ($STR $pprim)) >> ] ] + | Some p -> Some pprim, Prec p + | None -> Some pprim, Any ] ] + ; + (* meta-syntax entries *) + astpat: + [ [ "<<" ; a = Prim.ast; ">>" -> a + | a = default_action_parser -> + match a with + | Ast.PureAstNode a -> a + | _ -> failwith "Cannot deal with non pure ast expression" ] ] + ; + action: + [ [ IDENT "let"; p = Prim.astlist; et = set_internal_entry_type; + "="; e1 = action; "in"; e = action -> Ast.CaseAction (loc,e1,et,[p,e]) + | IDENT "case"; a = action; et = set_internal_entry_type; "of"; + cl = LIST1 case SEP "|"; IDENT "esac" -> Ast.CaseAction (loc,a,et,cl) + | "["; a = default_action_parser; "]" -> Ast.SimpleAction (loc,a) ] ] + ; + case: + [[ p = Prim.astlist; "->"; a = action -> (p,a) ]] + ; + set_internal_entry_type: + [[ ":"; IDENT "ast"; IDENT "list" -> Ast.ETastl + | [ ":"; IDENT "ast" -> () | -> () ] -> Ast.ETast ]] + ; + set_entry_type: + [[ ":"; et = entry_type -> set_default_action_parser et; entry_type_of_parser et + | -> None ]] + ; + entry_type: + [[ IDENT "ast"; IDENT "list" -> AstListParser + | IDENT "ast" -> AstParser + | IDENT "constr" -> ConstrParser + | IDENT "tactic" -> TacticParser + | IDENT "vernac" -> VernacParser ]] ; END diff --git a/parsing/g_cases.ml4 b/parsing/g_cases.ml4 index 48638a4fe..66536703a 100644 --- a/parsing/g_cases.ml4 +++ b/parsing/g_cases.ml4 @@ -23,7 +23,7 @@ GEXTEND Gram compound_pattern: [ [ p = pattern ; lp = ne_pattern_list -> <:ast< (PATTCONSTRUCT $p ($LIST $lp)) >> - | p = pattern; "as"; id = ident -> + | p = pattern; "as"; id = Prim.var -> <:ast< (PATTAS $id $p)>> | p1 = pattern; ","; p2 = pattern -> <:ast< (PATTCONSTRUCT Coq.Init.Datatypes.pair $p1 $p2) >> diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index c93acba58..cbe3a14c2 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -14,10 +14,10 @@ open Pcoq open Constr GEXTEND Gram - GLOBAL: ident constr0 constr1 constr2 constr3 lassoc_constr4 constr5 + GLOBAL: constr0 constr1 constr2 constr3 lassoc_constr4 constr5 constr6 constr7 constr8 constr9 constr10 lconstr constr ne_ident_comma_list ne_constr_list sort ne_binders_list qualid - global; + global constr_pattern ident; ident: [ [ id = Prim.var -> id @@ -25,14 +25,14 @@ GEXTEND Gram | id = Prim.metaident -> id ] ] ; global: - [ [ l = qualid -> <:ast< (QUALID ($LIST $l)) >> + [ [ l = qualid -> l (* This is used in quotations *) | id = Prim.metaident -> id ] ] ; qualid: - [ [ id = Prim.var; l = fields -> id :: l - | id = Prim.var -> [ id ] + [ [ id = Prim.var; l = fields -> <:ast< (QUALID $id ($LIST $l)) >> + | id = Prim.var -> <:ast< (QUALID $id) >> ] ] ; fields: @@ -40,30 +40,23 @@ GEXTEND Gram | id = FIELD -> [ <:ast< ($VAR $id) >> ] ] ] ; -(* - qualid: - [ [ id = IDENT; l = fields -> <:ast< ($VAR $id) >> :: l ] ] - ; - fields: - [ [ id = FIELD; l = fields -> <:ast< ($VAR $id) >> :: l - | -> [] - ] ] - ; -*) raw_constr: [ [ c = Prim.ast -> c ] ] ; constr: - [ [ c = constr8 -> c - | IDENT "Inst"; id = ident; "["; c = constr8; "]" -> - <:ast<(CONTEXT $id $c)>> - | IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = constr8 -> - <:ast<(EVAL $c (REDEXP $rtc))>> - | IDENT "Check"; c = constr8 -> <:ast<(CHECK $c)>> ] ] + [ [ c = constr8 -> (* Genarg.ConstrTerm *) c +(* | IDENT "Inst"; id = Prim.rawident; "["; c = constr; "]" -> + Genarg.ConstrContext (id, c) + | IDENT "Eval"; rtc = Tactic.raw_red_tactic; "in"; c = constr -> + Genarg.ConstrEval (rtc,c) + | IDENT "Check"; c = constr8 -> <:ast<(CHECK $c)>> *)] ] ; lconstr: [ [ c = constr10 -> c ] ] ; + constr_pattern: + [ [ c = constr -> c ] ] + ; ne_constr_list: [ [ c1 = constr; cl = ne_constr_list -> c1::cl | c1 = constr -> [c1] ] ] @@ -75,7 +68,8 @@ GEXTEND Gram ; constr0: [ [ "?" -> <:ast< (ISEVAR) >> - | "?"; n = Prim.number -> <:ast< (META $n) >> + | "?"; n = Prim.natural -> + let n = Coqast.Num (loc,n) in <:ast< (META $n) >> | bl = binders; c = constr -> <:ast< ($ABSTRACT "LAMBDALIST" $bl $c) >> | "("; lc1 = lconstr; ":"; c = constr; body = product_tail -> let id = Ast.coerce_to_var lc1 in @@ -122,12 +116,17 @@ GEXTEND Gram | IDENT "Match"; c = constr; "with"; cl = ne_constr_list; "end" -> <:ast< (MATCH "SYNTH" $c ($LIST $cl)) >> | IDENT "let"; "("; b = ne_ident_comma_list; ")"; "="; - c = constr; "in"; c1 = constr -> + c = constr; "in"; c1 = constr-> <:ast< (LET "SYNTH" $c ($ABSTRACT "LAMBDALIST" (BINDERS (BINDER (ISEVAR) ($LIST $b))) $c1)) >> + | IDENT "let"; id1 = ident ; "="; c = opt_casted_constr; + "in"; c1 = constr -> + <:ast< (LETIN $c [$id1]$c1) >> +(* | IDENT "let"; id1 = ident ; "="; c = opt_casted_constr; "in"; c1 = constr -> <:ast< (LETIN $c [$id1]$c1) >> +*) | IDENT "if"; c1 = constr; IDENT "then"; c2 = constr; IDENT "else"; c3 = constr -> <:ast< ( IF "SYNTH" $c1 $c2 $c3) >> @@ -170,9 +169,28 @@ GEXTEND Gram [ [ c1 = constr8 -> c1 | c1 = constr8; "::"; c2 = constr8 -> <:ast< (CAST $c1 $c2) >> ] ] ; + simple_binding: + [ [ id = ident; ":="; c = constr -> <:ast< (BINDING $id $c) >> + | n = Prim.numarg; ":="; c = constr -> <:ast< (BINDING $n $c) >> ] ] + ; + simple_binding_list: + [ [ b = simple_binding; l = simple_binding_list -> b :: l | -> [] ] ] + ; + binding_list: + [ [ c1 = constr; ":="; c2 = constr; bl = simple_binding_list -> + Coqast.Node + (loc, "EXPLICITBINDINGS", + (Coqast.Node (loc, "BINDING", [Ast.coerce_to_var c1; c2]) :: bl)) + | n = Prim.numarg; ":="; c = constr; bl = simple_binding_list -> + <:ast<(EXPLICITBINDINGS (BINDING $n $c) ($LIST $bl))>> + | c1 = constr; bl = LIST0 constr -> + <:ast<(IMPLICITBINDINGS $c1 ($LIST $bl))>> ] ] + ; constr10: [ [ "!"; f = global; args = ne_constr9_list -> <:ast< (APPLISTEXPL $f ($LIST $args)) >> + | "!"; f = global; "with"; b = binding_list -> + <:ast< (APPLISTWITH $f $b) >> | f = constr9; args = ne_constr91_list -> <:ast< (APPLIST $f ($LIST $args)) >> | f = constr9 -> f ] ] @@ -216,7 +234,8 @@ GEXTEND Gram | -> <:ast< (ISEVAR) >> ] ] ; constr91: - [ [ n = Prim.number; "!"; c1 = constr9 -> + [ [ n = Prim.natural; "!"; c1 = constr9 -> + let n = Coqast.Num (loc,n) in <:ast< (EXPL $n $c1) >> | c1 = constr9 -> c1 ] ] ; @@ -229,8 +248,10 @@ GEXTEND Gram | c1 = constr9 -> [c1] ] ] ; fixbinder: - [ [ id = ident; "/"; recarg = Prim.number; ":"; type_ = constr; - ":="; def = constr -> <:ast< (NUMFDECL $id $recarg $type_ $def) >> + [ [ id = ident; "/"; recarg = Prim.natural; ":"; type_ = constr; + ":="; def = constr -> + let recarg = Coqast.Num (loc,recarg) in + <:ast< (NUMFDECL $id $recarg $type_ $def) >> | id = ident; bl = ne_binders_list; ":"; type_ = constr; ":="; def = constr -> <:ast< (FDECL $id (BINDERS ($LIST $bl)) $type_ $def) >> ] ] diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4 index ec482d43d..0a8e2c5f8 100644 --- a/parsing/g_ltac.ml4 +++ b/parsing/g_ltac.ml4 @@ -11,167 +11,250 @@ (* $Id$ *) open Pp +open Util open Ast open Pcoq +open Coqast +open Rawterm +open Tacexpr +open Ast + +ifdef Quotify then +open Qast + +ifdef Quotify then +open Q + open Tactic -open Util -(* Tactics grammar rules *) +type let_clause_kind = + | LETTOPCLAUSE of Names.identifier * Genarg.constr_ast + | LETCLAUSE of + (Names.identifier Util.located * Genarg.constr_ast may_eval option * raw_tactic_arg) + +let fail_default_value = 0 + +let out_letin_clause loc = function + | LETTOPCLAUSE _ -> user_err_loc (loc, "", (str "Syntax Error")) + | LETCLAUSE (id,c,d) -> (id,c,d) + +let make_letin_clause _ = List.map (out_letin_clause dummy_loc) + +ifdef Quotify then +let fail_default_value = Qast.Int "0" + +ifdef Quotify then +let out_letin_clause loc = function + | Qast.Node ("LETTOPCLAUSE", _) -> user_err_loc (loc, "", (str "Syntax Error")) + | Qast.Node ("LETCLAUSE", [id;c;d]) -> + Qast.Tuple [id;c;d] + +ifdef Quotify then +let make_letin_clause loc = function + | Qast.List l -> Qast.List (List.map (out_letin_clause loc) l) (* Tactics grammar rules *) GEXTEND Gram + GLOBAL: tactic Vernac_.command tactic_arg; + +(* GLOBAL: tactic_atom tactic_atom0 tactic_expr input_fun; +*) input_fun: - [ [ l = identarg -> l - | "()" -> <:ast< (VOID) >> ] ] + [ [ l = Prim.ident -> Some l + | "()" -> None ] ] ; let_clause: - [ [ id = identarg; "="; te=tactic_atom0 -> <:ast< (LETCLAUSE $id $te) >> - | id = identarg; ":"; c = constrarg; ":="; "Proof" -> - (match c with - | Coqast.Node(_,"COMMAND",[csr]) -> - <:ast< (LETTOPCLAUSE $id (CONSTR $csr)) >> - | _ -> errorlabstrm "Gram.let_clause" (str "Not a COMMAND")) - | id = identarg; ":"; c = constrarg; ":="; te = tactic_expr -> - <:ast< (LETCUTCLAUSE $id $c $te) >> - | id = identarg; ":"; c = constrarg -> - (match c with - | Coqast.Node(_,"COMMAND",[csr]) -> - <:ast< (LETTOPCLAUSE $id (CONSTR $csr)) >> - | _ -> errorlabstrm "Gram.let_clause" (str "Not a COMMAND")) ] ] + [ [ id = Prim.rawident; "="; te = tactic_letarg -> LETCLAUSE (id, None, te) + | id = Prim.ident; ":"; c = Constr.constr; ":="; "Proof" -> + LETTOPCLAUSE (id, c) + | id = Prim.rawident; ":"; c = constrarg; ":="; te = tactic_letarg -> + LETCLAUSE (id, Some c, te) + | id = Prim.ident; ":"; c = Constr.constr -> + LETTOPCLAUSE (id, c) ] ] ; rec_clause: - [ [ name = identarg; it = LIST1 input_fun; "->"; body = tactic_atom -> - <:ast< (RECCLAUSE $name (FUNVAR ($LIST $it)) $body) >> ] ] + [ [ name = Prim.rawident; it = LIST1 input_fun; "->"; body = tactic_expr -> + (name,(it,body)) ] ] ; match_pattern: - [ [ id = constrarg; "["; pc = constrarg; "]" -> - (match id with - | Coqast.Node(_,"COMMAND", - [Coqast.Node(_,"QUALID",[Coqast.Nvar(_,_) as s])]) -> - <:ast< (SUBTERM $s $pc) >> - | _ -> - errorlabstrm "Gram.match_pattern" (str "Not a correct SUBTERM")) - | "["; pc = constrarg; "]" -> <:ast< (SUBTERM $pc) >> - | pc = constrarg -> <:ast< (TERM $pc) >> ] ] + [ [ id = Constr.constr_pattern; "["; pc = Constr.constr_pattern; "]" -> + let s = coerce_to_id id in Subterm (Some s, pc) + | "["; pc = Constr.constr_pattern; "]" -> Subterm (None,pc) + | pc = Constr.constr_pattern -> Term pc ] ] ; match_hyps: - [ [ id = identarg; ":"; mp = match_pattern -> - <:ast< (MATCHCONTEXTHYPS $id $mp) >> - | IDENT "_"; ":"; mp = match_pattern -> - <:ast< (MATCHCONTEXTHYPS $mp) >> ] ] + [ [ id = Prim.rawident; ":"; mp = match_pattern -> Hyp (id, mp) + | IDENT "_"; ":"; mp = match_pattern -> NoHypId mp ] ] ; match_context_rule: [ [ "["; largs = LIST0 match_hyps SEP ";"; "|-"; mp = match_pattern; "]"; - "->"; te = tactic_expr -> - <:ast< (MATCHCONTEXTRULE ($LIST $largs) $mp $te) >> - | IDENT "_"; "->"; te = tactic_expr -> <:ast< (MATCHCONTEXTRULE $te) >> - ] ] + "->"; te = tactic_expr -> Pat (largs, mp, te) + | IDENT "_"; "->"; 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 -> - <:ast<(MATCHRULE $mp $te)>> - | IDENT "_"; "->"; te = tactic_expr -> <:ast< (MATCHRULE $te) >> ] ] + [ [ "["; mp = match_pattern; "]"; "->"; te = tactic_expr -> Pat ([],mp,te) + | IDENT "_"; "->"; te = tactic_expr -> All te ] ] ; match_list: [ [ mrl = LIST1 match_rule SEP "|" -> mrl | "|"; mrl = LIST1 match_rule SEP "|" -> mrl ] ] ; tactic_expr: - [ [ ta0 = tactic_expr; ";"; ta1 = tactic_expr -> - <:ast< (TACTICLIST $ta0 $ta1) >> - | ta = tactic_expr; ";"; "["; lta = LIST0 tactic_expr SEP "|"; "]" -> - <:ast< (TACTICLIST $ta (TACLIST ($LIST $lta))) >> - | y = tactic_atom -> y ] ] + [ [ ta = tactic_expr5 -> ta ] ] + ; + tactic_expr5: + [ [ ta0 = tactic_expr5; ";"; ta1 = tactic_expr4 -> TacThen (ta0, ta1) + | ta = tactic_expr5; ";"; "["; lta = LIST0 tactic_expr SEP "|"; "]" -> + TacThens (ta, lta) + | y = tactic_expr4 -> y ] ] + ; + tactic_expr4: + [ [ ta = tactic_expr3 -> ta ] ] + ; + tactic_expr3: + [ [ IDENT "Try"; ta = tactic_expr3 -> TacTry ta + | IDENT "Do"; n = Prim.natural; ta = tactic_expr3 -> TacDo (n,ta) + | IDENT "Repeat"; ta = tactic_expr3 -> TacRepeat ta + | IDENT "Progress"; ta = tactic_expr3 -> TacProgress ta + | IDENT "Info"; tc = tactic_expr3 -> TacInfo tc + | ta = tactic_expr2 -> ta ] ] + ; + tactic_expr2: + [ [ ta0 = tactic_atom; "Orelse"; ta1 = tactic_expr3 -> TacOrelse (ta0,ta1) + | ta = tactic_atom -> ta ] ] ; tactic_atom: [ [ IDENT "Fun"; it = LIST1 input_fun ; "->"; body = tactic_expr -> - <:ast< (FUN (FUNVAR ($LIST $it)) $body) >> - | IDENT "Rec"; rc = rec_clause -> <:ast< (REC $rc) >> + TacFun (it,body) + | IDENT "Rec"; rc = rec_clause -> + TacFunRec rc | IDENT "Rec"; rc = rec_clause; IDENT "In"; body = tactic_expr -> - <:ast< (REC (RECDECL $rc) $body) >> + TacLetRecIn ([rc],body) | IDENT "Rec"; rc = rec_clause; "And"; rcl = LIST1 rec_clause SEP "And"; IDENT "In"; - body = tactic_expr -> - <:ast< (REC (RECDECL $rc ($LIST $rcl)) $body) >> + body = tactic_expr -> TacLetRecIn (rc::rcl,body) | IDENT "Let"; llc = LIST1 let_clause SEP "And"; IDENT "In"; - u = tactic_expr -> <:ast< (LET (LETDECL ($LIST $llc)) $u) >> - | IDENT "Let"; llc = LIST1 let_clause SEP "And" -> + u = tactic_expr -> TacLetIn (make_letin_clause loc llc,u) +(* Let cas LetCut est subsum par "Assert id := c" tandis que le cas + StartProof ne fait pas vraiment de sens en tant que sous-expression + d'une tactique complexe... + | IDENT "Let"; llc = LIST1 let_clause SEP "And" -> (match llc with - | [Coqast.Node(_,"LETTOPCLAUSE",[id;c])] -> - <:ast< (StartProof "LETTOP" $id $c) >> - | _ -> <:ast< (LETCUT (LETDECL ($LIST $llc))) >>) + | [LETTOPCLAUSE (id,c)] -> + VernacStartProof ((NeverDischarge,false),id,c,true,(fun _ _ -> ())) + | l -> + let l = List.map (function + | LETCLAUSE (id,Some a,t) -> (id,a,t) + | _ -> user_err_loc (loc, "", str "Syntax Error")) l in + TacLetCut (loc, l)) +*) +(* | IDENT "Let"; llc = LIST1 let_clause SEP "And"; tb = Vernac_.theorem_body; "Qed" -> - (match llc with - | [Coqast.Node(_,"LETTOPCLAUSE",[id;c])] -> - <:ast< (TheoremProof "LETTOP" $id $c $tb) >> - | _ -> errorlabstrm "Gram.tactic_atom" (str "Not a LETTOPCLAUSE")) - | IDENT "Match"; "Reverse"; IDENT "Context"; IDENT "With"; - mrl = match_context_list -> <:ast< (MATCHCONTEXT "LR" ($LIST $mrl)) >> - | IDENT "Match"; IDENT "Context"; IDENT "With"; - mrl = match_context_list -> <:ast< (MATCHCONTEXT "RL" ($LIST $mrl)) >> - | IDENT "Match"; com = constrarg; IDENT "With"; mrl = match_list -> - <:ast< (MATCH $com ($LIST $mrl)) >> -(* - | "("; te = tactic_expr; tel=LIST1 tactic_expr; ")" -> - <:ast< (APP $te ($LIST $tel)) >> + (match llc with + | [LETTOPCLAUSE (id,c)] -> + EscapeVernac <:ast< (TheoremProof "LETTOP" $id $c $tb) >> + | _ -> + errorlabstrm "Gram.tactic_atom" (str "Not a LETTOPCLAUSE")) *) + + | IDENT "Match"; IDENT "Context"; IDENT "With"; mrl = match_context_list + -> TacMatchContext (false,mrl) + | IDENT "Match"; IDENT "Reverse"; IDENT "Context"; IDENT "With"; mrl = match_context_list + -> TacMatchContext (true,mrl) + | IDENT "Match"; c = constrarg; IDENT "With"; mrl = match_list -> + TacMatch (c,mrl) +(*To do: put Abstract in Refiner*) + | IDENT "Abstract"; tc = tactic_expr -> TacAbstract (tc,None) + | IDENT "Abstract"; tc = tactic_expr; "using"; s = Prim.ident -> + TacAbstract (tc,Some s) +(*End of To do*) | IDENT "First" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - <:ast<(FIRST ($LIST $l))>> - | IDENT "Info"; tc = tactic_expr -> <:ast< (INFO $tc) >> + TacFirst l | IDENT "Solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - <:ast<(TCLSOLVE ($LIST $l))>> - | IDENT "Try"; ta0 = tactic_atom; "Orelse"; ta1 = tactic_atom -> - <:ast< (TRY (ORELSE $ta0 $ta1)) >> - | IDENT "Try"; ta = tactic_atom -> <:ast< (TRY $ta) >> - | IDENT "Do"; n = pure_numarg; ta = tactic_atom -> <:ast< (DO $n $ta) >> - | IDENT "Repeat"; ta0 = tactic_atom; "Orelse"; ta1 = tactic_atom -> - <:ast< (REPEAT (ORELSE $ta0 $ta1)) >> - | IDENT "Repeat"; ta = tactic_atom -> <:ast< (REPEAT $ta) >> - | IDENT "Idtac" -> <:ast< (IDTAC) >> - | IDENT "Fail" -> <:ast<(FAIL)>> - | IDENT "Fail"; n = pure_numarg -> <:ast<(FAIL $n)>> - | ta0 = tactic_atom; "Orelse"; ta1 = tactic_atom -> - <:ast< (ORELSE $ta0 $ta1) >> - | IDENT "Progress"; ta = tactic_atom -> <:ast< (PROGRESS $ta) >> - (* Copy here to factorize with pure ident *) - | IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = Constr.constr -> - <:ast< (COMMAND (EVAL $c (REDEXP $rtc))) >> - | IDENT "Inst"; id = identarg; "["; c = Constr.constr; "]" -> - <:ast< (COMMAND (CONTEXT $id $c)) >> - | IDENT "Check"; c = Constr.constr -> <:ast<(COMMAND (CHECK $c))>> - | st = simple_tactic -> st - | id = lqualid; la = LIST1 tactic_atom0 -> - <:ast< (APP $id ($LIST $la)) >> - | id = lqualid -> id - | ta = tactic_atom0 -> ta ] ] - ; - tactic_atom0: - (* Tactic arguments *) - [ [ "("; te = tactic_expr; ")" -> te - | id = lqualid -> id - | "()" -> <:ast< (VOID) >> - | n = pure_numarg -> n - | id = Prim.metaident -> id - | "?" -> <:ast< (COMMAND (ISEVAR)) >> - | "?"; n = Prim.number -> <:ast< (COMMAND (META $n)) >> - | IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = Constr.constr -> - <:ast< (COMMAND (EVAL $c (REDEXP $rtc))) >> - | IDENT "Inst"; id = identarg; "["; c = Constr.constr; "]" -> - <:ast< (COMMAND (CONTEXT $id $c)) >> - | IDENT "Check"; c = Constr.constr -> <:ast<(COMMAND (CHECK $c))>> - | "'"; c = constrarg -> c ] ] + TacSolve l + | IDENT "Idtac" -> TacId + | IDENT "Fail" -> TacFail fail_default_value + | IDENT "Fail"; n = Prim.natural -> TacFail n + | st = simple_tactic -> TacAtom (loc,st) + | "("; a = tactic_expr; ")" -> a + | a = tactic_arg -> TacArg a + ] ] + ; + (* Tactic arguments *) + tactic_arg: + [ [ ta = tactic_arg1 -> ta ] ] + ; + tactic_letarg: + (* Cannot be merged with tactic_arg1, since then "In"/"And" are + parsed as lqualid! *) + [ [ IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = Constr.constr -> + ConstrMayEval (ConstrEval (rtc,c)) + | IDENT "Inst"; id = Prim.rawident; "["; c = Constr.constr; "]" -> + ConstrMayEval (ConstrContext (id,c)) + | IDENT "Check"; c = Constr.constr -> + ConstrMayEval (ConstrTypeOf c) + | qid = lqualid -> qid + | ta = tactic_arg0 -> ta ] ] + ; + tactic_arg1: + [ [ IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = Constr.constr -> + ConstrMayEval (ConstrEval (rtc,c)) + | IDENT "Inst"; id = Prim.rawident; "["; c = Constr.constr; "]" -> + ConstrMayEval (ConstrContext (id,c)) + | IDENT "Check"; c = Constr.constr -> + ConstrMayEval (ConstrTypeOf c) + | qid = lqualid; la = LIST1 tactic_arg0 -> TacCall (loc,qid,la) + | qid = lqualid -> qid + | ta = tactic_arg0 -> ta ] ] + ; + tactic_arg0: + [ [ "("; a = tactic_expr; ")" -> Tacexp a + | "()" -> TacVoid + | qid = lqualid -> qid + | n = Prim.integer -> Integer n + | id = METAIDENT -> MetaIdArg (loc,id) + | "?" -> ConstrMayEval (ConstrTerm <:ast< (ISEVAR) >>) + | "?"; n = Prim.natural -> MetaNumArg (loc,n) + | "'"; c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ] ; lqualid: - [ [ l = Constr.qualid -> - (match l with - | [id] -> id - | _ -> <:ast< (QUALIDARG ($LIST $l)) >>) ] ] + [ [ qid = Prim.reference -> +(* match Nametab.repr_qualid qid with + | (dir,id) when dir = Nameops.empty_dirpath -> Identifier id + | _ -> Qualid *) Reference qid + ] ] + ; + + (* Definitions for tactics *) + deftok: + [ [ IDENT "Meta" + | IDENT "Tactic" ] ] + ; + vrec_clause: + [ [ name = Prim.rawident; it=LIST1 input_fun; ":="; body = tactic_expr -> + (name, TacFunRec (name, (it, body))) + | name = Prim.rawident; ":="; body = tactic_expr -> + (name, body) ] ] + ; + tactic: + [ [ tac = tactic_expr -> tac ] ] + ; + Vernac_.command: + [ [ deftok; "Definition"; name = Prim.rawident; ":="; body=Tactic.tactic -> + Vernacexpr.VernacDeclareTacticDefinition (loc, [name, body]) + | deftok; "Definition"; name = Prim.rawident; largs=LIST1 input_fun; + ":="; body=tactic_expr -> + Vernacexpr.VernacDeclareTacticDefinition + (loc, [name, TacFun (largs,body)]) + | IDENT "Recursive"; deftok; "Definition"; + vcl=LIST1 vrec_clause SEP "And" -> + Vernacexpr.VernacDeclareTacticDefinition (loc, vcl) ] ] ; END diff --git a/parsing/g_natsyntax.ml b/parsing/g_natsyntax.ml index 9a97f4b06..97ffba1d0 100644 --- a/parsing/g_natsyntax.ml +++ b/parsing/g_natsyntax.ml @@ -19,12 +19,12 @@ open Coqast open Ast open Coqlib open Termast +open Extend exception Non_closed_number let ast_O = ast_of_ref glob_O let ast_S = ast_of_ref glob_S -let ast_myvar = ast_of_ref glob_My_special_variable_nat (* For example, (nat_of_string "3") is <<(S (S (S O)))>> *) let nat_of_int n dloc = @@ -69,37 +69,29 @@ let int_of_nat p = with Non_closed_number -> None -let replace_S p = - let rec aux p = - match p with - | Node (l,"APPLIST", [b; a]) when b = ast_S -> - Node (l, "APPLIST", [ast_myvar; (aux a)]) - | _ -> p - in - ope("APPLIST", [ast_myvar; (aux p)]) +let pr_S a = hov 0 (str "S" ++ brk (1,1) ++ a) +let rec pr_external_S std_pr = function + | Node (l,"APPLIST", [b; a]) when b = ast_S -> + str"(" ++ pr_S (pr_external_S std_pr a) ++ str")" + | p -> std_pr p (* Declare the primitive printer *) (* Prints not p, but the SUCCESSOR of p !!!!! *) let nat_printer std_pr p = match (int_of_nat p) with - | Some i -> (str (string_of_int i)) - | None -> std_pr (replace_S p) + | Some i -> str (string_of_int i) + | None -> pr_S (pr_external_S std_pr p) let _ = Esyntax.Ppprim.add ("nat_printer", nat_printer) (* Declare the primitive parser *) -let number = - match create_entry (get_univ "nat") "number" ETast with - | Ast n -> n - | _ -> anomaly "G_natsyntax : create_entry number failed" +let unat = create_univ_if_new "nat" -let pat_number = - match create_entry (get_univ "nat") "pat_number" ETast with - | Ast n -> n - | _ -> anomaly "G_natsyntax : create_entry number failed" +let number = create_constr_entry unat "number" +let pat_number = create_constr_entry unat "pat_number" let _ = Gram.extend number None diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index e8cd55117..be51cc0bc 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -8,39 +8,133 @@ (*i $Id$ i*) +(* +camlp4o pa_ifdef.cmo pa_extend.cmo pr_o.cmo pr_extend.cmo -quotify -DQuotify -impl parsing/g_prim.ml4 +*) + open Coqast open Pcoq open Names + +ifdef Quotify then +open Qast + open Prim +let local_make_qualid id l id' = Nametab.make_qualid (make_dirpath (id::l))id' +let local_make_short_qualid id = Nametab.make_short_qualid id +let local_make_posint = int_of_string +let local_make_negint n = - int_of_string n +let local_make_path id l a = make_path (make_dirpath (l@[id_of_string id])) a +let local_make_binding loc a b = + match a with + | Nvar (_,id) -> Slam(loc,Some id,b) + | Nmeta (_,s) -> Smetalam(loc,s,b) + | _ -> failwith "Slam expects a var or a metavar" +let local_append l id = l@[id] + +(* camlp4o pa_extend.cmo pa_extend_m.cmo pr_o.cmo q_prim.ml *) + +ifdef Quotify then +let id_of_string s = Apply ("Names.id_of_string", [s]) +ifdef Quotify then +let make_dirpath l = Apply ("Names.make_dirpath", [l]) +ifdef Quotify then +let local_make_posint s = s +ifdef Quotify then +let local_make_negint s = Qast.Apply ("-", [s]) +ifdef Quotify then +let local_make_qualid id l id' = + Qast.Apply ("Nametab.make_qualid", + [Qast.Apply ("Nametab.make_dirpath", + [Qast.Cons (id, l)]); id']) +ifdef Quotify then +let local_make_short_qualid id = + Qast.Apply ("Nametab.make_short_qualid",[id]) +ifdef Quotify then +let local_make_path id l a = + Qast.Apply ("Names.make_path", + [Qast.Apply ("Names.make_dirpath", + [Qast.Apply ("List.append",[l; id_of_string id]); a])]) +ifdef Quotify then +let local_make_binding loc a b = + match a with + | Qast.Node ("Nvar", [_;id]) -> + Qast.Node ("Slam", [Qast.Loc; Qast.Option (Some id); b]) + | Qast.Node ("Nmeta", [_;s]) -> + Qast.Node ("Smetalam", [Qast.Loc;s;b]) + | _ -> + Qast.Apply ("failwith", [Qast.Str "Slam expects a var or a metavar"]) +ifdef Quotify then +let local_append l id = Qast.Apply ("@", [l; Qast.List [id]]) + +ifdef Quotify then +open Q + GEXTEND Gram - GLOBAL: var ident metaident number string (*path*) ast astpat - astact entry_type; + GLOBAL: var ident natural metaident integer string preident ast astpat + astact astlist qualid reference dirpath rawident numarg; metaident: - [ [ s = METAIDENT -> Nmeta(loc,s) ] ] + [ [ s = METAIDENT -> Nmeta (loc,s) ] ] ; var: [ [ s = IDENT -> Nvar(loc, id_of_string s) ] ] ; + preident: + [ [ s = IDENT -> s ] ] + ; ident: - [ [ s = IDENT -> Id(loc,s) ] ] + [ [ s = IDENT -> id_of_string s ] ] + ; + rawident: + [ [ id = ident -> (loc,id) ] ] + ; + natural: + [ [ i = INT -> local_make_posint i ] ] + ; + integer: + [ [ i = INT -> local_make_posint i + | "-"; i = INT -> local_make_negint i ] ] ; - number: - [ [ i = INT -> Num(loc, int_of_string i) ] ] + numarg: + [ [ n = integer -> Num (loc, n) ] ] + ; + field: + [ [ s = FIELD -> id_of_string s ] ] + ; + dirpath: + [ [ id = ident; l = LIST0 field -> make_dirpath (id::l) ] ] + ; + fields: + [ [ id = field; (l,id') = fields -> (id::l,id') + | id = field -> ([],id) + ] ] + ; + basequalid: + [ [ id = ident; (l,id') = fields -> local_make_qualid id l id' + | id = ident -> local_make_short_qualid id + ] ] + ; + qualid: + [ [ qid = basequalid -> loc, qid ] ] + ; + reference: + [ [ id = ident; (l,id') = fields -> + Tacexpr.RQualid (loc, local_make_qualid id l id') + | id = ident -> Tacexpr.RIdent (loc,id) + ] ] ; string: - [ [ s = STRING -> Str(loc,s) ] ] + [ [ s = STRING -> s ] ] ; astpath: - [ [ id = IDENT; (l,a) = astfields -> - let p = make_dirpath (List.rev (id_of_string id :: l)) in - Path(loc, make_path p a) + [ [ id = IDENT; (l,a) = astfields -> Path(loc, local_make_path id l a) | id = IDENT -> Nvar(loc, id_of_string id) ] ] ; astfields: - [ [ id = FIELD; (l,a) = astfields -> id_of_string id :: l, a + [ [ id = FIELD; (l,a) = astfields -> local_append l (id_of_string id), a | id = FIELD -> [], id_of_string id ] ] ; @@ -51,7 +145,7 @@ GEXTEND Gram ast: [ [ id = metaident -> id | p = astpath -> p - | s = INT -> Num(loc, int_of_string s) + | s = INT -> Num(loc, local_make_posint s) | s = STRING -> Str(loc, s) | "{"; s = METAIDENT; "}" -> Id(loc,s) | "("; nname = astident; l = LIST0 ast; ")" -> Node(loc,nname,l) @@ -63,11 +157,7 @@ GEXTEND Gram | "("; METAIDENT "$PATH"; id = metaident; ")" -> Node(loc,"$PATH",[id]) | "("; METAIDENT "$NUM"; id = metaident; ")" -> Node(loc,"$NUM",[id]) | "["; "<>"; "]"; b = ast -> Slam(loc,None,b) - | "["; a = ast; "]"; b = ast -> - (match a with - | Nvar (_,id) -> Slam(loc,Some id,b) - | Nmeta (_,s) -> Smetalam(loc,s,b) - | _ -> failwith "Slam expects a var or a metavar") + | "["; a = ast; "]"; b = ast -> local_make_binding loc a b (* | "["; ido = astidoption; "]"; b = ast -> Slam(loc,ido,b) @@ -75,45 +165,13 @@ GEXTEND Gram *) | "'"; a = ast -> Node(loc,"$QUOTE",[a]) ] ] ; + astlist: + [ [ l = LIST0 Prim.ast -> l ] ] + ; (* astidoption: [ [ "<>" -> None | id = IDENT -> Some (id_of_string id) ] ] ; *) - (* meta-syntax entries *) - astpat: - [ [ "<<" ; a = ast; ">>" -> Node loc "ASTPAT" [a] - | a = default_action_parser -> Node loc "ASTPAT" [a] ] ] - ; - astact: - [ [ a = action -> Node loc "ASTACT" [a] ] ] - ; - astlist: - [ [ l = LIST0 ast -> Node loc "ASTLIST" l ] ] - ; - action: - [ [ IDENT "let"; p = astlist; et = entry_type; "="; e1 = action; "in"; - e = action -> Node(loc,"$CASE",[e1; et; Node(loc,"CASE",[p;e])]) - | IDENT "case"; a = action; et = entry_type; "of"; - cl = LIST1 case SEP "|"; IDENT "esac" -> - Node(loc,"$CASE",a::et::cl) - | "["; al = default_action_parser; "]" -> al ] ] - ; - case: - [[ p = astlist; "->"; a = action -> Node(loc,"CASE",[p;a]) ]] - ; - entry_type: - [[ ":"; IDENT "ast"; IDENT "list" -> - let _ = set_default_action_parser astlist in Id(loc,"LIST") - | ":"; IDENT "ast" -> - let _ = set_default_action_parser ast in Id(loc,"AST") - | ":"; IDENT "constr" -> - let _ = set_default_action_parser Constr.constr in Id(loc,"AST") - | ":"; IDENT "tactic" -> - let _ = set_default_action_parser Tactic.tactic in Id(loc,"AST") - | ":"; IDENT "vernac" -> - let _ = set_default_action_parser Vernac_.vernac in Id(loc,"AST") - | -> Id(loc,"AST") ]] - ; END diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 5e8d853ee..78294f2a7 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -15,168 +15,110 @@ open Pp open Tactic open Util open Vernac_ +open Coqast +open Vernacexpr +open Prim (* Proof commands *) GEXTEND Gram - GLOBAL: command ne_constrarg_list; + GLOBAL: command; destruct_location : - [ [ IDENT "Conclusion" -> <:ast< (CONCL)>> - | IDENT "Discardable"; "Hypothesis" -> <:ast< (DiscardableHYP)>> - | "Hypothesis" -> <:ast< (PreciousHYP)>> ]] + [ [ IDENT "Conclusion" -> Tacexpr.ConclLocation () + | discard = [ IDENT "Discardable" -> true | -> false ]; "Hypothesis" + -> Tacexpr.HypLocation discard ] ] ; - ne_constrarg_list: - [ [ l = LIST1 constrarg -> l ] ] - ; - opt_identarg_list: + opt_hintbases: [ [ -> [] - | ":"; l = LIST1 identarg -> l ] ] - ; - deftok: - [ [ IDENT "Meta" - | IDENT "Tactic" ] ] - ; - vrec_clause: - [ [ name=identarg; it=LIST1 input_fun; ":="; body=tactic_expr -> - <:ast<(RECCLAUSE $name (FUNVAR ($LIST $it)) $body)>> - | name=identarg; ":="; body=tactic_expr -> - <:ast<(RECCLAUSE $name $body)>> ] ] + | ":"; l = LIST1 IDENT -> l ] ] ; command: - [ [ IDENT "Goal"; c = constrarg -> <:ast< (GOAL $c) >> - | IDENT "Goal" -> <:ast< (GOAL) >> - | "Proof" -> <:ast< (GOAL) >> - | IDENT "Begin" -> <:ast< (GOAL) >> - | IDENT "Abort" -> <:ast< (ABORT) >> - | "Qed" -> <:ast< (SaveNamed) >> - | IDENT "Save" -> <:ast< (SaveNamed) >> - | IDENT "Defined" -> <:ast< (DefinedNamed) >> - | IDENT "Defined"; id = identarg -> <:ast< (DefinedAnonymous $id) >> - | IDENT "Save"; tok = thm_tok; id = identarg -> - <:ast< (SaveAnonymous $tok $id) >> - | IDENT "Save"; id = identarg -> <:ast< (SaveAnonymous $id) >> - | IDENT "Suspend" -> <:ast< (SUSPEND) >> - | IDENT "Resume" -> <:ast< (RESUME) >> - | IDENT "Resume"; id = identarg -> <:ast< (RESUME $id) >> - | IDENT "Abort"; IDENT "All" -> <:ast< (ABORTALL) >> - | IDENT "Abort"; id = identarg -> <:ast< (ABORT $id) >> - | IDENT "Restart" -> <:ast< (RESTART) >> - | "Proof"; c = constrarg -> <:ast< (PROOF $c) >> - | IDENT "Undo" -> <:ast< (UNDO 1) >> - | IDENT "Undo"; n = numarg -> <:ast< (UNDO $n) >> - | IDENT "Show"; n = numarg -> <:ast< (SHOW $n) >> - | IDENT "Show"; IDENT "Implicits"; n = numarg -> - <:ast< (SHOWIMPL $n) >> - | IDENT "Focus" -> <:ast< (FOCUS) >> - | IDENT "Focus"; n = numarg -> <:ast< (FOCUS $n) >> - | IDENT "Unfocus" -> <:ast< (UNFOCUS) >> - | IDENT "Show" -> <:ast< (SHOW) >> - | IDENT "Show"; IDENT "Intro" -> <:ast< (SHOWINTRO) >> - | IDENT "Show"; IDENT "Intros" -> <:ast< (SHOWINTROS) >> - | IDENT "Show"; IDENT "Implicits" -> <:ast< (SHOWIMPL) >> - | IDENT "Show"; IDENT "Node" -> <:ast< (ShowNode) >> - | IDENT "Show"; IDENT "Script" -> <:ast< (ShowScript) >> - | IDENT "Show"; IDENT "Existentials" -> <:ast< (ShowEx) >> - | IDENT "Existential"; n = numarg; ":="; c = constrarg -> - <:ast< (EXISTENTIAL $n $c) >> - | IDENT "Existential"; n = numarg; ":="; c1 = Constr.constr; ":"; - c2 = Constr.constr -> - <:ast< (EXISTENTIAL $n (CONSTR (CAST $c1 $c2))) >> - | IDENT "Existential"; n = numarg; ":"; c2 = Constr.constr; ":="; - c1 = Constr.constr -> - <:ast< (EXISTENTIAL $n (CONSTR (CAST $c1 $c2))) >> - | IDENT "Explain"; "Proof"; l = numarg_list -> - <:ast< (ExplainProof ($LIST $l)) >> - | IDENT "Explain"; "Proof"; IDENT "Tree"; l = numarg_list -> - <:ast< (ExplainProofTree ($LIST $l)) >> - | IDENT "Go"; n = numarg -> <:ast< (Go $n) >> - | IDENT "Go"; IDENT "top" -> <:ast< (Go "top") >> - | IDENT "Go"; IDENT "prev" -> <:ast< (Go "prev") >> - | IDENT "Go"; IDENT "next" -> <:ast< (Go "next") >> - | IDENT "Show"; "Proof" -> <:ast< (ShowProof) >> - | IDENT "Guarded" -> <:ast< (CheckGuard) >> - | IDENT "Show"; IDENT "Tree" -> <:ast< (ShowTree) >> - | IDENT "Show"; IDENT "Conjectures" -> <:ast< (ShowProofs) >> - -(* Definitions for tactics *) - - | deftok; "Definition"; name=identarg; ":="; body=Tactic.tactic - -> <:ast<(TACDEF $name (AST $body))>> - | deftok; "Definition"; name=identarg; largs=LIST1 input_fun; - ":="; body=Tactic.tactic -> - <:ast<(TACDEF $name (AST (FUN (FUNVAR ($LIST $largs)) $body)))>> - | IDENT "Recursive"; deftok; "Definition"; vc=vrec_clause -> - (match vc with - | Coqast.Node(_,"RECCLAUSE",nme::[_;_]) -> - <:ast<(TACDEF $nme (AST (REC $vc)))>> - | Coqast.Node(_,"RECCLAUSE",nme::[bd]) -> - <:ast<(TACDEF $nme (AST $bd))>> - | _ -> - anomalylabstrm "Gram.vernac" (str "Not a correct RECCLAUSE")) - | IDENT "Recursive"; deftok; "Definition"; vc=vrec_clause; "And"; - vcl=LIST1 vrec_clause SEP "And" -> - let nvcl= - List.fold_right - (fun e b -> match e with - | Coqast.Node(_,"RECCLAUSE",nme::[_;_]) -> - nme::<:ast<(AST (REC $e))>>::b - | Coqast.Node(_,"RECCLAUSE",nme::[bd]) -> - nme::<:ast<(AST $bd)>>::b - | _ -> - anomalylabstrm "Gram.vernac" - (str "Not a correct RECCLAUSE")) - (vc::vcl) [] - in - <:ast<(TACDEF ($LIST $nvcl))>> - + [ [ IDENT "Goal"; c = Constr.constr -> + VernacStartProof(StartTheoremProof Theorem,None,c,false,(fun _ _ ->())) +(*VernacGoal c*) +(* | IDENT "Goal" -> VernacGoal None*) + | "Proof" -> VernacNop +(* Used ?? + | IDENT "Begin" -> VernacNop +*) + | IDENT "Abort" -> VernacAbort None + | IDENT "Abort"; IDENT "All" -> VernacAbortAll + | IDENT "Abort"; id = ident -> VernacAbort (Some id) + | "Qed" -> VernacEndProof (true,None) + | IDENT "Save" -> VernacEndProof (true,None) + | IDENT "Defined" -> VernacEndProof (false,None) + | IDENT "Defined"; id = ident -> VernacEndProof (false,Some (id,None)) + | IDENT "Save"; tok = thm_token; id = ident -> + VernacEndProof (true,Some (id,Some tok)) + | IDENT "Save"; id = ident -> VernacEndProof (true,Some (id,None)) + | IDENT "Suspend" -> VernacSuspend + | IDENT "Resume" -> VernacResume None + | IDENT "Resume"; id = ident -> VernacResume (Some id) + | IDENT "Restart" -> VernacRestart + | "Proof"; c = Constr.constr -> VernacExactProof c + | IDENT "Undo" -> VernacUndo 1 + | IDENT "Undo"; n = natural -> VernacUndo n + | IDENT "Focus" -> VernacFocus None + | IDENT "Focus"; n = natural -> VernacFocus (Some n) + | IDENT "Unfocus" -> VernacUnfocus + | IDENT "Show" -> VernacShow (ShowGoal None) + | IDENT "Show"; n = natural -> VernacShow (ShowGoal (Some n)) + | IDENT "Show"; IDENT "Implicits"; n = natural -> + VernacShow (ShowGoalImplicitly (Some n)) + | IDENT "Show"; IDENT "Implicits" -> VernacShow (ShowGoalImplicitly None) + | IDENT "Show"; IDENT "Node" -> VernacShow ShowNode + | IDENT "Show"; IDENT "Script" -> VernacShow ShowScript + | IDENT "Show"; IDENT "Existentials" -> VernacShow ShowExistentials + | IDENT "Show"; IDENT "Tree" -> VernacShow ShowTree + | IDENT "Show"; IDENT "Conjectures" -> VernacShow ShowProofNames + | IDENT "Show"; "Proof" -> VernacShow ShowProof + | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false) + | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true) + | IDENT "Explain"; "Proof"; l = LIST0 integer -> + VernacShow (ExplainProof l) + | IDENT "Explain"; "Proof"; IDENT "Tree"; l = LIST0 integer -> + VernacShow (ExplainTree l) + | IDENT "Go"; n = natural -> VernacGo (GoTo n) + | IDENT "Go"; IDENT "top" -> VernacGo GoTop + | IDENT "Go"; IDENT "prev" -> VernacGo GoPrev + | IDENT "Go"; IDENT "next" -> VernacGo GoNext + | IDENT "Guarded" -> VernacCheckGuard (* Hints for Auto and EAuto *) - | IDENT "Hint"; hintname = identarg; dbname = opt_identarg_list; ":="; - IDENT "Resolve"; c = constrarg -> - <:ast<(HintResolve $hintname (VERNACARGLIST ($LIST $dbname)) $c)>> - - | IDENT "Hint"; hintname = identarg; dbnames = opt_identarg_list; ":="; - IDENT "Immediate"; c = constrarg -> - <:ast<(HintImmediate $hintname (VERNACARGLIST ($LIST $dbnames)) $c)>> - - | IDENT "Hint"; hintname = identarg; dbnames = opt_identarg_list; ":="; - IDENT "Unfold"; c = qualidarg -> - <:ast<(HintUnfold $hintname (VERNACARGLIST ($LIST $dbnames)) $c)>> - - | IDENT "Hint"; hintname = identarg; dbnames = opt_identarg_list; ":="; - IDENT "Constructors"; c = qualidarg -> - <:ast<(HintConstructors $hintname (VERNACARGLIST ($LIST $dbnames)) $c)>> - - | IDENT "Hint"; hintname = identarg; dbnames = opt_identarg_list; ":="; - IDENT "Extern"; n = numarg; c = constrarg ; tac = tacarg -> - <:ast<(HintExtern $hintname (VERNACARGLIST ($LIST $dbnames)) - $n $c (TACTIC $tac))>> - - | IDENT "Hints"; IDENT "Resolve"; l = ne_qualidarg_list; - dbnames = opt_identarg_list -> - <:ast< (HintsResolve (VERNACARGLIST ($LIST $dbnames)) ($LIST $l)) >> - - | IDENT "Hints"; IDENT "Immediate"; l = ne_qualidarg_list; - dbnames = opt_identarg_list -> - <:ast< (HintsImmediate (VERNACARGLIST ($LIST $dbnames)) ($LIST $l)) >> - - | IDENT "Hints"; IDENT "Unfold"; l = ne_qualidarg_list; - dbnames = opt_identarg_list -> - <:ast< (HintsUnfold (VERNACARGLIST ($LIST $dbnames)) ($LIST $l)) >> - | IDENT "HintDestruct"; - dloc = destruct_location; - na = identarg; - hyptyp = constrarg; - pri = numarg; - tac = Prim.astact -> - <:ast< (HintDestruct $na (AST $dloc) $hyptyp $pri (AST $tac))>> + dloc = destruct_location; + id = ident; + hyptyp = Constr.constr_pattern; + pri = natural; + tac = tactic -> + VernacHintDestruct (id,dloc,hyptyp,pri,tac) - | n = numarg; ":"; tac = tacarg -> - <:ast< (SOLVE $n (TACTIC $tac)) >> + | IDENT "Hint"; hintname = ident; dbnames = opt_hintbases; ":="; h = hint + -> VernacHints (dbnames, h hintname) + + | IDENT "Hints"; (dbnames,h) = hints -> VernacHints (dbnames, h) + (*This entry is not commented, only for debug*) - | IDENT "PrintConstr"; com = constrarg -> - <:ast< (PrintConstr $com)>> + | IDENT "PrintConstr"; c = Constr.constr -> + VernacExtend ("PrintConstr", + [Genarg.in_gen Genarg.rawwit_constr c]) ] ]; + + hint: + [ [ IDENT "Resolve"; c = Constr.constr -> fun name -> HintsResolve [Some name, c] + | IDENT "Immediate"; c = Constr.constr -> fun name -> HintsImmediate [Some name, c] + | IDENT "Unfold"; qid = qualid -> fun name -> HintsUnfold [Some name,qid] + | IDENT "Constructors"; c = qualid -> fun n -> HintsConstructors (n,c) + | IDENT "Extern"; n = natural; c = Constr.constr8 ; tac = tactic -> + fun name -> HintsExtern (name,n,c,tac) ] ] + ; + hints: + [ [ IDENT "Resolve"; l = LIST1 Constr.qualid; dbnames = opt_hintbases -> + (dbnames, HintsResolve (List.map (fun qid -> (None, qid)) l)) + | IDENT "Immediate"; l = LIST1 Constr.qualid; dbnames = opt_hintbases -> + (dbnames, HintsImmediate (List.map (fun qid -> (None, qid)) l)) + | IDENT "Unfold"; l = LIST1 qualid; dbnames = opt_hintbases -> + (dbnames, HintsUnfold (List.map (fun qid -> (None,qid)) l)) ] ] + ; END diff --git a/parsing/g_rsyntax.ml b/parsing/g_rsyntax.ml index 41fe0df2a..bc2fb999f 100644 --- a/parsing/g_rsyntax.ml +++ b/parsing/g_rsyntax.ml @@ -5,12 +5,14 @@ (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) + open Coqast open Ast open Pp open Util open Names open Pcoq +open Extend exception Non_closed_number @@ -36,10 +38,7 @@ let r_of_int n dloc = let r_of_string s dloc = r_of_int (int_of_string s) dloc -let rnumber = - match create_entry (get_univ "rnatural") "rnumber" ETast with - | Ast n -> n - | _ -> anomaly "G_rsyntax : create_entry rnumber failed" +let rnumber = create_constr_entry (get_univ "rnatural") "rnumber" let _ = Gram.extend rnumber None @@ -71,13 +70,13 @@ let replace_plus p = let r_printer std_pr p = let (_,ast1,astp,_) = get_r_sign dummy_loc in match (int_of_r p) with - | Some i -> (str (string_of_int (i+1))) + | Some i -> str (string_of_int (i+1)) | None -> std_pr (replace_plus p) let r_printer_outside std_pr p = let (_,ast1,astp,_) = get_r_sign dummy_loc in match (int_of_r p) with - | Some i -> (str "``" ++ str (string_of_int (i+1)) ++ str "``") + | Some i -> str "``" ++ str (string_of_int (i+1)) ++ str "``" | None -> std_pr (replace_plus p) let _ = Esyntax.Ppprim.add ("r_printer", r_printer) diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4 index f3b4d2f17..77bbc03ee 100644 --- a/parsing/g_tactic.ml4 +++ b/parsing/g_tactic.ml4 @@ -13,368 +13,325 @@ open Pp open Ast open Pcoq -open Tactic open Util +open Tacexpr +open Rawterm + +(* open grammar entries, possibly in quotified form *) +ifdef Quotify then open Qast + +open Constr +open Prim +open Tactic + +(* Functions overloaded by quotifier *) + +let induction_arg_of_constr c = + try ElimOnIdent (Ast.loc c,coerce_to_id c) with _ -> ElimOnConstr c + +let local_compute = [FBeta;FIota;FDeltaBut [];FZeta] + +let error_oldelim _ = error "OldElim no longer supported" + +ifdef Quotify then + let induction_arg_of_constr = function + | Qast.Node ("Nvar", [_;id]) -> Qast.Node ("ElimOnIdent", [id]) + | c -> Qast.Node ("ElimOnConstr", [c]) + +ifdef Quotify then +let make_red_flag s = Qast.Apply ("make_red_flag", [s]) + +ifdef Quotify then +let local_compute = + Qast.List [ + Qast.Node ("FBeta", []); + Qast.Node ("FDeltaBut", [Qast.List []]); + Qast.Node ("FIota", []); + Qast.Node ("FZeta", [])] + +(* + let + {rBeta = b; rIota = i; rZeta = z; rDelta = d; rConst = l} = make_red_flag s + in + let quotify_ident id = + Qast.Apply ("Names.id_of_string", [Qast.Str (Names.string_of_id id)]) + in + let quotify_path sp = + let dir, id = Names.repr_path sp in + let l = Names.repr_dirpath dir in + Qast.Apply ("Names.make_path", + [Qast.Apply ("Names.make_dirpath", + [Qast.List (List.map quotify_ident l)]); + quotify_ident id]) in + Qast.Record + ["rBeta",Qast.Bool b; "rIota",Qast.Bool i; + "rZeta",Qast.Bool z; "rDelta",Qast.Bool d; + "rConst",Qast.List (List.map quotify_path l)] +*) +ifdef Quotify then open Q (* Please leave several GEXTEND for modular compilation under PowerPC *) (* Auxiliary grammar rules *) GEXTEND Gram - GLOBAL: autoargs; + GLOBAL: simple_tactic constrarg constr_with_bindings quantified_hypothesis + red_tactic int_or_var castedopenconstr; + int_or_var: + [ [ n = integer -> Genarg.ArgArg n + | id = ident -> Genarg.ArgVar (loc,id) ] ] + ; autoarg_depth: - [ [ n = pure_numarg -> <:ast< $n>> - | -> Coqast.Str(loc,"NoAutoArg") ] ] + [ [ n = OPT natural -> n ] ] ; autoarg_adding: - [ [ IDENT "Adding" ; "["; l = ne_qualidarg_list; "]" -> l - | -> [] ] ] + [ [ IDENT "Adding" ; "["; l = LIST1 qualid; "]" -> l | -> [] ] ] ; autoarg_destructing: - [ [ IDENT "Destructing" -> Coqast.Str(loc,"Destructing") - | -> Coqast.Str(loc,"NoAutoArg") ] ] + [ [ IDENT "Destructing" -> true | -> false ] ] ; autoarg_usingTDB: - [ [ "Using"; "TDB" -> Coqast.Str(loc,"UsingTDB") - | -> Coqast.Str(loc,"NoAutoArg") ] ] + [ [ "Using"; "TDB" -> true | -> false ] ] ; autoargs: [ [ a0 = autoarg_depth; l = autoarg_adding; - a2 = autoarg_destructing; a3 = autoarg_usingTDB -> a0::a2::a3::l ] ] + a2 = autoarg_destructing; a3 = autoarg_usingTDB -> (a0,l,a2,a3) ] ] ; - END - -GEXTEND Gram - - identarg: - [ [ id = Constr.ident -> id ] ] - ; - idmeta_arg: - [ [ id = Constr.ident -> id - | "?"; n = Prim.number -> <:ast< (COMMAND (META $n)) >> ] ] + (* Either an hypothesis or a ltac ref (variable or pattern metavariable) *) + id_or_ltac_ref: + [ [ id = ident -> AN (loc,id) + | "?"; n = natural -> MetaNum (loc,n) ] ] ; - idmetahyp: - [ [ id = idmeta_arg -> <:ast< (INHYP $id) >> ] ] + (* Either a global ref or a ltac ref (variable or pattern metavariable) *) + qualid_or_ltac_ref: + [ [ (loc,qid) = qualid -> AN (loc,qid) + | "?"; n = natural -> MetaNum (loc,n) ] ] ; - qualidarg: - [ [ l = Constr.qualid -> <:ast< (QUALIDARG ($LIST $l)) >> - | "?"; n = Prim.number -> <:ast< (QUALIDMETA $n) >> ] ] - ; - pure_numarg: - [ [ n = Prim.number -> n - | "-"; n = Prim.number -> Coqast.Num (Ast.loc n, ( - Ast.num_of_ast n)) - ] ] + (* An identifier or a quotation meta-variable *) + id_or_meta: + [ [ id = rawident -> AI id + + (* This is used in quotations *) + | id = METAIDENT -> MetaId (loc,id) ] ] ; - numarg: - [ [ n = Prim.number -> n - | "-"; n = Prim.number -> Coqast.Num (Ast.loc n, ( - Ast.num_of_ast n)) - | id = identarg -> id + (* A number or a quotation meta-variable *) + num_or_meta: + [ [ n = integer -> AI n + | id = METAIDENT -> MetaId (loc,id) ] ] ; constrarg: - [ [ c = Constr.constr -> <:ast< (COMMAND $c) >> ] ] - ; - castedopenconstrarg: - [ [ c = Constr.constr -> <:ast< (CASTEDOPENCOMMAND $c) >> ] ] + [ [ IDENT "Inst"; id = rawident; "["; c = constr; "]" -> + ConstrContext (id, c) + | IDENT "Eval"; rtc = Tactic.red_tactic; "in"; c = constr -> + ConstrEval (rtc,c) + | IDENT "Check"; c = constr -> ConstrTypeOf c + | c = constr -> ConstrTerm c ] ] ; - lconstrarg: - [ [ c = Constr.lconstr -> <:ast< (COMMAND $c) >> ] ] + castedopenconstr: + [ [ c = constr -> c ] ] ; - castedconstrarg: - [ [ c = Constr.constr -> <:ast< (CASTEDCOMMAND $c) >> ] ] + induction_arg: + [ [ n = natural -> ElimOnAnonHyp n + | c = constr -> induction_arg_of_constr c + ] ] ; - ident_or_numarg: - [ [ id = identarg -> id - | n = numarg -> n ] ] - ; - ident_or_constrarg: - [ [ c = Constr.constr -> - match c with - | Coqast.Nvar(_,s) -> c - | Coqast.Node(_,"QUALID",[Coqast.Nvar(_,s) as c]) -> c - | _ -> <:ast< (COMMAND $c) >> ] ] - ; - ne_identarg_list: - [ [ l = LIST1 identarg -> l ] ] - ; - ne_idmetahyp_list: - [ [ l = LIST1 idmetahyp -> l ] ] - ; - ne_qualidarg_list: - [ [ l = LIST1 qualidarg -> l ] ] + quantified_hypothesis: + [ [ id = ident -> NamedHyp id + | n = natural -> AnonHyp n ] ] ; pattern_occ: - [ [ nl = LIST1 pure_numarg; c = constrarg -> - <:ast< (PATTERN $c ($LIST $nl)) >> - | c = constrarg -> <:ast< (PATTERN $c) >> ] ] + [ [ nl = LIST0 integer; c = constr -> (nl,c) ] ] ; - ne_pattern_list: - [ [ l = LIST1 pattern_occ -> l ] ] - ; - pattern_occ_hyp: - [ [ nl = LIST1 pure_numarg; IDENT "Goal" -> - <:ast<(CCLPATTERN ($LIST $nl))>> - | nl = LIST1 pure_numarg; id = identarg -> - <:ast<(HYPPATTERN $id ($LIST $nl))>> - | IDENT "Goal" -> <:ast< (CCLPATTERN) >> - | id = identarg -> <:ast< (HYPPATTERN $id) >> ] ] + pattern_occ_hyp_list: + [ [ nl = LIST1 natural; IDENT "Goal" -> (Some nl,[]) + | nl = LIST1 natural; id = id_or_meta; (g,l) = pattern_occ_hyp_list -> + (g,(id,nl)::l) + | IDENT "Goal" -> (Some [],[]) + | id = id_or_meta; (g,l) = pattern_occ_hyp_list -> (g,(id,[])::l) ] ] ; clause_pattern: - [ [ "in"; l = LIST1 pattern_occ_hyp -> <:ast< (LETPATTERNS ($LIST $l)) >> - | -> <:ast< (LETPATTERNS) >> ] ] - ; - one_intropattern: - [ [ p= ne_intropattern -> <:ast< (INTROPATTERN $p)>> ]] + [ [ "in"; p = pattern_occ_hyp_list -> p | -> None, [] ] ] ; - ne_intropattern: - [ [ tc = LIST1 simple_intropattern -> - <:ast< (LISTPATTERN ($LIST $tc))>> ] ] + intropatterns: + [ [ l = LIST0 simple_intropattern -> l ]] ; simple_intropattern: - [ [ "["; tc = LIST1 intropattern SEP "|" ; "]" -> - <:ast< (DISJPATTERN ($LIST $tc)) >> - | "("; tc = LIST1 intropattern SEP "," ; ")" -> - <:ast< (CONJPATTERN ($LIST $tc)) >> - | IDENT "_" -> - <:ast< (WILDCAR) >> - | id = identarg -> - <:ast< (IDENTIFIER $id) >> + [ [ "["; tc = LIST1 intropatterns SEP "|" ; "]" -> IntroOrAndPattern tc + | "("; tc = LIST1 simple_intropattern SEP "," ; ")" -> IntroOrAndPattern [tc] + | IDENT "_" -> IntroWildcard + | id = ident -> IntroIdentifier id ] ] ; - intropattern: - [ [ tc = LIST1 simple_intropattern -> - <:ast< (LISTPATTERN ($LIST $tc))>> - | -> <:ast< (LISTPATTERN) >> ]] - ; simple_binding: - [ [ id = identarg; ":="; c = constrarg -> <:ast< (BINDING $id $c) >> - | n = pure_numarg; ":="; c = constrarg -> <:ast< (BINDING $n $c) >> ] ] - ; - simple_binding_list: - [ [ b = simple_binding; l = simple_binding_list -> b :: l | -> [] ] ] - ; - com_binding_list: - [ [ c = constrarg; bl = com_binding_list -> <:ast< (BINDING $c) >> :: bl - | -> [] ] ] + [ [ id = ident; ":="; c = constr -> (NamedHyp id, c) + | n = natural; ":="; c = constr -> (AnonHyp n, c) ] ] ; binding_list: - [ [ c1 = constrarg; ":="; c2 = constrarg; bl = simple_binding_list -> - let id = match c1 with - | Coqast.Node(_,"COMMAND",[c]) -> coerce_to_var c - | _ -> assert false - in <:ast<(BINDINGS (BINDING $id $c2) ($LIST $bl))>> - | n = pure_numarg; ":="; c = constrarg; bl = simple_binding_list -> - <:ast<(BINDINGS (BINDING $n $c) ($LIST $bl))>> - | c1 = constrarg; bl = com_binding_list -> - <:ast<(BINDINGS (BINDING $c1) ($LIST $bl))>> - | -> <:ast<(BINDINGS)>> ] ] - ; - with_binding_list: - [ [ "with"; l = binding_list -> l | -> <:ast<(BINDINGS)>> ] ] - ; - constrarg_binding_list: - [ [ c = constrarg; l = with_binding_list -> [c; l] ] ] - ; - numarg_binding_list: - [ [ n = numarg; l = with_binding_list -> [n; l] ] ] + [ [ c1 = constr; ":="; c2 = constr; bl = LIST0 simple_binding -> + ExplicitBindings ((NamedHyp (coerce_to_id c1), c2) :: bl) + | n = natural; ":="; c = constr; bl = LIST0 simple_binding -> + ExplicitBindings ((AnonHyp n, c) :: bl) + | c1 = constr; bl = LIST0 constr -> + ImplicitBindings (c1 :: bl) ] ] ; - constrarg_list: - [ [ c = constrarg; l = constrarg_list -> c :: l | -> [] ] ] + constr_with_bindings: + [ [ c = constr; l = with_binding_list -> (c, l) ] ] ; - ne_constrarg_list: - [ [ c = constrarg; l = constrarg_list -> c :: l ] ] + with_binding_list: + [ [ "with"; bl = binding_list -> bl | -> NoBindings ] ] ; unfold_occ: - [ [ nl = LIST1 pure_numarg; c = qualidarg -> - <:ast< (UNFOLD $c ($LIST $nl)) >> - | c = qualidarg -> <:ast< (UNFOLD $c) >> ] ] - ; - ne_unfold_occ_list: - [ [ p = unfold_occ; l = ne_unfold_occ_list -> p :: l - | p = unfold_occ -> [p] ] ] + [ [ nl = LIST0 integer; c = qualid_or_ltac_ref -> (nl,c) ] ] ; red_flag: - [ [ IDENT "Beta" -> <:ast< (Beta) >> - | IDENT "Delta" -> <:ast< (Delta) >> - | IDENT "Iota" -> <:ast< (Iota) >> - | IDENT "Zeta" -> <:ast< (Zeta) >> - | "["; idl = ne_qualidarg_list; "]" -> <:ast< (Unf ($LIST $idl)) >> - | "-"; "["; idl = ne_qualidarg_list; "]" -> - <:ast< (UnfBut ($LIST $idl)) >> ] ] + [ [ IDENT "Beta" -> FBeta + | IDENT "Delta" -> FDeltaBut [] + | IDENT "Iota" -> FIota + | IDENT "Zeta" -> FZeta + | IDENT "Delta"; "["; idl = LIST1 qualid_or_ltac_ref; "]" -> FConst idl + | IDENT "Delta"; "-"; "["; idl = LIST1 qualid_or_ltac_ref; "]" -> FDeltaBut idl + ] ] ; red_tactic: - [ [ IDENT "Red" -> <:ast< (Red) >> - | IDENT "Hnf" -> <:ast< (Hnf) >> - | IDENT "Simpl" -> <:ast< (Simpl) >> - | IDENT "Cbv"; s = LIST1 red_flag -> <:ast< (Cbv ($LIST $s)) >> - | IDENT "Lazy"; s = LIST1 red_flag -> <:ast< (Lazy ($LIST $s)) >> - | IDENT "Compute" -> <:ast< (Cbv (Beta) (Delta) (Iota) (Zeta)) >> - | IDENT "Unfold"; ul = ne_unfold_occ_list -> - <:ast< (Unfold ($LIST $ul)) >> - | IDENT "Fold"; cl = constrarg_list -> <:ast< (Fold ($LIST $cl)) >> - | IDENT "Pattern"; pl = ne_pattern_list -> - <:ast< (Pattern ($LIST $pl)) >> ] ] + [ [ IDENT "Red" -> Red false + | IDENT "Hnf" -> Hnf + | IDENT "Simpl" -> Simpl + | IDENT "Cbv"; s = LIST1 red_flag -> Cbv (make_red_flag s) + | IDENT "Lazy"; s = LIST1 red_flag -> Lazy (make_red_flag s) + | IDENT "Compute" -> Cbv (make_red_flag [FBeta;FIota;FDeltaBut [];FZeta]) + | IDENT "Unfold"; ul = LIST1 unfold_occ -> Unfold ul + | IDENT "Fold"; cl = LIST1 constr -> Fold cl + | IDENT "Pattern"; pl = LIST1 pattern_occ -> Pattern pl ] ] ; hypident: - [ [ id = identarg -> <:ast< (INHYP $id) >> - | "("; "Type"; "of"; id = identarg; ")" -> <:ast< (INHYPTYPE $id) >> ] ] + [ [ id = id_or_meta -> InHyp id + | "("; "Type"; "of"; id = id_or_meta; ")" -> InHypType id ] ] ; - ne_hyp_list: - [ [ l = LIST1 hypident -> l ] ] - ; - clausearg: - [ [ "in"; idl = ne_hyp_list -> <:ast< (CLAUSE ($LIST $idl)) >> - | -> <:ast< (CLAUSE) >> ] ] + clause: + [ [ "in"; idl = LIST1 hypident -> idl + | -> [] ] ] ; fixdecl: - [ [ id = identarg; n = pure_numarg; ":"; c = constrarg; fd = fixdecl -> - <:ast< (FIXEXP $id $n $c) >> :: fd - | -> [] ] ] + [ [ id = ident; "/"; n = natural; ":"; c = constr -> (id,n,c) ] ] ; cofixdecl: - [ [ id = identarg; ":"; c = constrarg; fd = cofixdecl -> - <:ast< (COFIXEXP $id $c) >> :: fd - | -> [] ] ] + [ [ id = ident; ":"; c = constr -> (id,c) ] ] + ; + hintbases: + [ [ "with"; "*" -> None + | "with"; l = LIST1 IDENT -> Some l + | -> Some [] ] ] ; simple_tactic: - [ [ IDENT "Fix"; n = pure_numarg -> <:ast< (Fix $n) >> - | IDENT "Fix"; id = identarg; n = pure_numarg -> <:ast< (Fix $id $n) >> - | IDENT "Fix"; id = identarg; n = pure_numarg; "with"; fd = fixdecl -> - <:ast< (Fix $id $n ($LIST $fd)) >> - | IDENT "Cofix" -> <:ast< (Cofix) >> - | IDENT "Cofix"; id = identarg -> <:ast< (Cofix $id) >> - | IDENT "Cofix"; id = identarg; "with"; fd = cofixdecl -> - <:ast< (Cofix $id ($LIST $fd)) >> - | IDENT "Induction"; s = identarg -> <:ast< (Induction $s) >> - | IDENT "Induction"; n = pure_numarg -> <:ast< (Induction $n) >> - | IDENT "NewInduction"; c=ident_or_constrarg -> <:ast<(NewInduction $c)>> - | IDENT "NewInduction"; n = pure_numarg -> <:ast< (NewInduction $n) >> - | IDENT "Double"; IDENT "Induction"; i = pure_numarg; j = pure_numarg -> - <:ast< (DoubleInd $i $j) >> - | IDENT "Trivial" -> <:ast<(Trivial)>> - | IDENT "Trivial"; "with"; lid = ne_identarg_list -> - <:ast<(Trivial ($LIST $lid))>> - | IDENT "Trivial"; "with"; "*" -> <:ast<(Trivial "*")>> - | IDENT "Auto"; n = pure_numarg -> <:ast< (Auto $n) >> - | IDENT "Auto" -> <:ast< (Auto) >> - | IDENT "Auto"; n = pure_numarg; "with"; - lid = ne_identarg_list -> <:ast< (Auto $n ($LIST $lid)) >> - | IDENT "Auto"; "with"; - lid = ne_identarg_list -> <:ast< (Auto ($LIST $lid)) >> - | IDENT "Auto"; n = pure_numarg; "with"; "*" -> - <:ast< (Auto $n "*") >> - | IDENT "Auto"; "with"; "*" -> <:ast< (Auto "*") >> - - | IDENT "AutoTDB"; n = pure_numarg -> <:ast< (AutoTDB $n) >> - | IDENT "AutoTDB" -> <:ast< (AutoTDB) >> - | IDENT "DHyp"; id=identarg -> <:ast< (DHyp $id)>> - | IDENT "CDHyp"; id=identarg -> <:ast< (CDHyp $id)>> - | IDENT "DConcl" -> <:ast< (DConcl)>> - | IDENT "SuperAuto"; l = autoargs -> - <:ast< (SuperAuto ($LIST $l)) >> - | IDENT "Auto"; n = pure_numarg; IDENT "Decomp" -> <:ast< (DAuto $n) >> - | IDENT "Auto"; IDENT "Decomp" -> <:ast< (DAuto) >> - | IDENT "Auto"; n = pure_numarg; IDENT "Decomp"; p = pure_numarg -> - <:ast< (DAuto $n $p) >> - | IDENT "Intros" -> <:ast< (Intros) >> - | IDENT "Intros"; IDENT "until"; id = identarg - -> <:ast< (IntrosUntil $id) >> - | IDENT "Intros"; IDENT "until"; n = numarg -> <:ast<(IntrosUntil $n)>> - | IDENT "Intros"; pl = one_intropattern -> <:ast< (Intros $pl) >> - | IDENT "Intro"; id = identarg; - IDENT "after"; id2 = identarg -> <:ast< (IntroMove $id $id2) >> - | IDENT "Intro"; - IDENT "after"; id2 = identarg -> <:ast< (IntroMove $id2) >> - | IDENT "Intro"; id = identarg -> <:ast< (Intro $id) >> - | IDENT "Intro" -> <:ast< (Intro) >> - | IDENT "Apply"; cl = constrarg_binding_list -> - <:ast< (Apply ($LIST $cl)) >> - | IDENT "Elim"; cl = constrarg_binding_list; "using"; - el = constrarg_binding_list -> - <:ast< (Elim ($LIST $cl) ($LIST $el)) >> - | IDENT "Elim"; cl = constrarg_binding_list -> - <:ast< (Elim ($LIST $cl)) >> - | IDENT "Assumption" -> <:ast< (Assumption) >> - | IDENT "Contradiction" -> <:ast< (Contradiction) >> - | IDENT "Exact"; c = castedconstrarg -> <:ast< (Exact $c) >> - | IDENT "OldElim"; c = constrarg -> <:ast< (OldElim $c) >> - | IDENT "ElimType"; c = constrarg -> <:ast< (ElimType $c) >> - | IDENT "Case"; cl = constrarg_binding_list -> - <:ast< (Case ($LIST $cl)) >> - | IDENT "CaseType"; c = constrarg -> <:ast< (CaseType $c) >> - | IDENT "Destruct"; s = ident_or_constrarg -> <:ast< (Destruct $s) >> - | IDENT "Destruct"; n = numarg -> <:ast< (Destruct $n) >> - | IDENT "NewDestruct"; s = ident_or_constrarg -> <:ast<(NewDestruct $s)>> - | IDENT "NewDestruct"; n = numarg -> <:ast< (NewDestruct $n) >> - | IDENT "Decompose"; IDENT "Record" ; c = constrarg -> - <:ast< (DecomposeAnd $c) >> - | IDENT "Decompose"; IDENT "Sum"; c = constrarg -> - <:ast< (DecomposeOr $c) >> - | IDENT "Decompose"; "["; l = ne_qualidarg_list; "]"; c = constrarg -> - <:ast< (DecomposeThese $c ($LIST $l)) >> - | IDENT "Cut"; c = constrarg -> <:ast< (Cut $c) >> - | IDENT "Assert"; c = constrarg -> <:ast< (TrueCut $c) >> - | IDENT "Assert"; c = constrarg; ":"; t = constrarg -> - let id = match c with - | Coqast.Node(_,"COMMAND",[c]) -> coerce_to_var c - | _ -> assert false in - <:ast< (TrueCut $t $id) >> - | IDENT "Assert"; c = constrarg; ":="; b = constrarg -> - let id = match c with - | Coqast.Node(_,"COMMAND",[c]) -> coerce_to_var c - | _ -> assert false in - <:ast< (Forward "HideBody" $b $id) >> - | IDENT "Pose"; c = constrarg; ":="; b = constrarg -> - let id = match c with - | Coqast.Node(_,"COMMAND",[c]) -> coerce_to_var c - | _ -> assert false in - <:ast< (Forward "KeepBody" $b $id) >> - | IDENT "Pose"; b = constrarg -> - <:ast< (Forward "KeepBody" $b) >> - | IDENT "Specialize"; n = pure_numarg; lcb = constrarg_binding_list -> - <:ast< (Specialize $n ($LIST $lcb))>> - | IDENT "Specialize"; lcb = constrarg_binding_list -> - <:ast< (Specialize ($LIST $lcb)) >> - | IDENT "Generalize"; lc = constrarg_list -> - <:ast< (Generalize ($LIST $lc)) >> - | IDENT "Generalize"; IDENT "Dependent"; c = constrarg -> - <:ast< (GeneralizeDep $c) >> - | IDENT "LetTac"; s = identarg; ":="; c = constrarg; p = clause_pattern-> - <:ast< (LetTac $s $c $p) >> - | IDENT "LApply"; c = constrarg -> <:ast< (CutAndApply $c) >> - | IDENT "Clear"; l = ne_idmetahyp_list -> - <:ast< (Clear (CLAUSE ($LIST $l))) >> - | IDENT "ClearBody"; l = ne_idmetahyp_list -> - <:ast< (ClearBody (CLAUSE ($LIST $l))) >> - | IDENT "Move"; id1 = identarg; IDENT "after"; id2 = identarg -> - <:ast< (MoveDep $id1 $id2) >> - | IDENT "Rename"; id1 = identarg; IDENT "into"; id2 = identarg -> - <:ast< (Rename $id1 $id2) >> -(*To do: put Abstract in Refiner*) - | IDENT "Abstract"; tc = tactic_expr -> <:ast< (ABSTRACT (TACTIC $tc)) >> - | IDENT "Abstract"; tc = tactic_expr; "using"; s=identarg -> - <:ast< (ABSTRACT $s (TACTIC $tc)) >> -(*End of To do*) - | IDENT "Left"; bl = with_binding_list -> <:ast< (Left $bl) >> - | IDENT "Right"; bl = with_binding_list -> <:ast< (Right $bl) >> - | IDENT "Split"; bl = with_binding_list -> <:ast< (Split $bl) >> - | IDENT "Exists"; bl = binding_list -> <:ast< (Split $bl) >> - | IDENT "Constructor"; nbl = numarg_binding_list -> - <:ast<(Constructor ($LIST $nbl)) >> - | IDENT "Constructor" ; tc = tactic_expr -> <:ast<(Constructor (TACTIC $tc)) >> - | IDENT "Constructor" -> <:ast<(Constructor) >> - | IDENT "Reflexivity" -> <:ast< (Reflexivity) >> - | IDENT "Symmetry" -> <:ast< (Symmetry) >> - | IDENT "Transitivity"; c = constrarg -> <:ast< (Transitivity $c) >> - | IDENT "Absurd"; c = constrarg -> <:ast< (Absurd $c) >> - | IDENT "Idtac" -> <:ast< (Idtac) >> - | IDENT "Fail" -> <:ast< (Fail) >> - | r = red_tactic; cl = clausearg -> <:ast< (Reduce (REDEXP $r) $cl) >> + [ [ + (* Basic tactics *) + IDENT "Intros"; IDENT "until"; id = quantified_hypothesis -> + TacIntrosUntil id + | IDENT "Intros"; pl = intropatterns -> TacIntroPattern pl + | IDENT "Intro"; id = ident; IDENT "after"; id2 = rawident -> + TacIntroMove (Some id, Some id2) + | IDENT "Intro"; IDENT "after"; id2 = rawident -> + TacIntroMove (None, Some id2) + | IDENT "Intro"; id = ident -> TacIntroMove (Some id, None) + | IDENT "Intro" -> TacIntroMove (None, None) + + | IDENT "Assumption" -> TacAssumption + | IDENT "Exact"; c = constr -> TacExact c + + | IDENT "Apply"; cl = constr_with_bindings -> TacApply cl + | IDENT "Elim"; cl = constr_with_bindings; "using"; + el = constr_with_bindings -> TacElim (cl,Some el) + | IDENT "Elim"; cl = constr_with_bindings -> TacElim (cl,None) + | IDENT "OldElim"; c = constr -> + (* TacOldElim c *) error_oldelim () + | IDENT "ElimType"; c = constr -> TacElimType c + | IDENT "Case"; cl = constr_with_bindings -> TacCase cl + | IDENT "CaseType"; c = constr -> TacCaseType c + | IDENT "Fix"; n = natural -> TacFix (None,n) + | IDENT "Fix"; id = ident; n = natural -> TacFix (Some id,n) + | IDENT "Fix"; id = ident; n = natural; "with"; fd = LIST0 fixdecl -> + TacMutualFix (id,n,fd) + | IDENT "Cofix" -> TacCofix None + | IDENT "Cofix"; id = ident -> TacCofix (Some id) + | IDENT "Cofix"; id = ident; "with"; fd = LIST0 cofixdecl -> + TacMutualCofix (id,fd) + + | IDENT "Cut"; c = constr -> TacCut c + | IDENT "Assert"; c = constr -> TacTrueCut (None,c) + | IDENT "Assert"; c = constr; ":"; t = constr -> + TacTrueCut (Some (coerce_to_id c),t) + | IDENT "Assert"; c = constr; ":="; b = constr -> + TacForward (false,Names.Name (coerce_to_id c),b) + | IDENT "Pose"; c = constr; ":="; b = constr -> + TacForward (true,Names.Name (coerce_to_id c),b) + | IDENT "Pose"; b = constr -> TacForward (true,Names.Anonymous,b) + | IDENT "Generalize"; lc = LIST1 constr -> TacGeneralize lc + | IDENT "Generalize"; IDENT "Dependent"; c = constr -> TacGeneralizeDep c + | IDENT "LetTac"; id = ident; ":="; c = constr; p = clause_pattern + -> TacLetTac (id,c,p) + | IDENT "Instantiate"; n = natural; c = constr -> TacInstantiate (n,c) + + | IDENT "Specialize"; n = OPT natural; lcb = constr_with_bindings -> + TacSpecialize (n,lcb) + | IDENT "LApply"; c = constr -> TacLApply c + + (* Derived basic tactics *) + | IDENT "Induction"; h = quantified_hypothesis -> TacOldInduction h + | IDENT "NewInduction"; c = induction_arg -> TacNewInduction c + | IDENT "Double"; IDENT "Induction"; h1 = quantified_hypothesis; + h2 = quantified_hypothesis -> TacDoubleInduction (h1,h2) + | IDENT "Destruct"; h = quantified_hypothesis -> TacOldDestruct h + | IDENT "NewDestruct"; c = induction_arg -> TacNewDestruct c + | IDENT "Decompose"; IDENT "Record" ; c = constr -> TacDecomposeAnd c + | IDENT "Decompose"; IDENT "Sum"; c = constr -> TacDecomposeOr c + | IDENT "Decompose"; "["; l = LIST1 qualid_or_ltac_ref; "]"; c = constr + -> TacDecompose (l,c) + + (* Automation tactic *) + | IDENT "Trivial"; db = hintbases -> TacTrivial db + | IDENT "Auto"; n = OPT natural; db = hintbases -> TacAuto (n, db) + + | IDENT "AutoTDB"; n = OPT natural -> TacAutoTDB n + | IDENT "CDHyp"; id = rawident -> TacDestructHyp (true,id) + | IDENT "DHyp"; id = rawident -> TacDestructHyp (false,id) + | IDENT "DConcl" -> TacDestructConcl + | IDENT "SuperAuto"; l = autoargs -> TacSuperAuto l + | IDENT "Auto"; n = OPT natural; IDENT "Decomp"; p = OPT natural -> + TacDAuto (n, p) + + (* Context management *) + | IDENT "Clear"; l = LIST1 id_or_ltac_ref -> TacClear l + | IDENT "ClearBody"; l = LIST1 id_or_ltac_ref -> TacClearBody l + | IDENT "Move"; id1 = rawident; IDENT "after"; id2 = rawident -> + TacMove (true,id1,id2) + | IDENT "Rename"; id1 = rawident; IDENT "into"; id2 = rawident -> + TacRename (id1,id2) + + (* Constructors *) + | IDENT "Left"; bl = with_binding_list -> TacLeft bl + | IDENT "Right"; bl = with_binding_list -> TacRight bl + | IDENT "Split"; bl = with_binding_list -> TacSplit bl + | IDENT "Exists"; bl = binding_list -> TacSplit bl + | IDENT "Exists" -> TacSplit NoBindings + | IDENT "Constructor"; n = num_or_meta; l = with_binding_list -> + TacConstructor (n,l) + | IDENT "Constructor"; t = OPT tactic -> TacAnyConstructor t + + (* Equivalence relations *) + | IDENT "Reflexivity" -> TacReflexivity + | IDENT "Symmetry" -> TacSymmetry + | IDENT "Transitivity"; c = constr -> TacTransitivity c + + (* Conversion *) + | r = red_tactic; cl = clause -> TacReduce (r, cl) (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) - | IDENT "Change"; c = constrarg; cl = clausearg -> - <:ast< (Change $c $cl) >> - | IDENT "ML"; s = Prim.string -> <:ast< (MLTACTIC $s) >> ] + | IDENT "Change"; c = constr; cl = clause -> TacChange (c,cl) + +(* Unused ?? + | IDENT "ML"; s = string -> ExtraTactic<:ast< (MLTACTIC $s) >> +*) (* | [ id = identarg; l = constrarg_list -> match (isMeta (nvar_of_ast id), l) with @@ -383,9 +340,6 @@ GEXTEND Gram | _ -> Util.user_err_loc (loc, "G_tactic.meta_tactic", (str"Cannot apply arguments to a meta-tactic.")) - ] *)] - ; - tactic: - [ [ tac = tactic_expr -> tac ] ] + ] *)] ] ; END;; diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index a864227a7..9397e7658 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -10,11 +10,35 @@ (* $Id$ *) +open Coqast +open Vernacexpr open Pcoq open Pp open Tactic open Util +open Constr open Vernac_ +open Prim + +(* Dans join_binders, s'il y a un "?", on perd l'info qu'il est partag *) +let join_binders (idl,c) = List.map (fun id -> (id,c)) idl + +open Genarg + +let cast_constr loc c = function + | None -> c + | Some t -> <:ast< (CAST $c $t) >> + +let abstract_constr bl c = + let loc = dummy_loc in + <:ast<($ABSTRACT "LAMBDALIST" $bl $c)>> + +let generalize_constr bl c = + let loc = dummy_loc in + <:ast< ($ABSTRACT "PRODLIST" $bl $c) >> + +let evar_constr = let loc = dummy_loc in <:ast< (ISEVAR) >> + (* Rem: do not join the different GEXTEND into one, it breaks native *) (* compilation on PowerPC and Sun architectures *) @@ -28,121 +52,123 @@ GEXTEND Gram | g = gallina_ext; "." -> g | c = command; "." -> c | c = syntax; "." -> c - | "["; l = vernac_list_tail -> <:ast< (VernacList ($LIST $l)) >> - + | n = Prim.natural; ":"; tac = Tactic.tactic; "." -> VernacSolve (n,tac) + | "["; l = vernac_list_tail -> VernacList l (* This is for "Grammar vernac" rules *) - | id = Prim.metaident -> id ] ] + | id = Prim.metaident -> VernacVar (Ast.nvar_of_ast id) ] ] ; vernac: FIRST - [ [ IDENT "Time"; v = vernac -> <:ast< (Time $v)>> ] ] + [ [ IDENT "Time"; v = vernac -> VernacTime v ] ] ; vernac: LAST - [ [ tac = tacarg; "." -> + [ [ tac = Tactic.tactic; "." -> (match tac with - | Coqast.Node(_,kind,_) - when kind = "StartProof" || kind = "TheoremProof" -> tac - | _ -> <:ast< (SOLVE 1 (TACTIC $tac)) >>) ] ] +(* Horrible hack pour LETTOP ! + | VernacSolve (Coqast.Node(_,kind,_)) + when kind = "StartProof" || kind = "TheoremProof" -> ?? +*) + | _ -> VernacSolve (1,tac)) + | IDENT "Existential"; n = natural; c = constr_body -> + VernacSolveExistential (n,c) + ] ] + ; + def_body: + [ [ ":="; c = constr; ":"; t = constr -> c, Some t + | ":"; t = constr; ":="; c = constr -> c, Some t + | ":="; c = constr -> c, None ] ] + ; + constr_body: + [ [ (c,t) = def_body -> cast_constr loc c t ] ] ; vernac_list_tail: - [ [ v = vernac; l = vernac_list_tail -> v :: l + [ [ v = located_vernac; l = vernac_list_tail -> v :: l | "]"; "." -> [] ] ] ; + located_vernac: + [ [ v = vernac -> loc, v ] ] + ; END +let test_plurial_form = function + | [_] -> + Options.if_verbose warning + "Keywords Variables/Hypotheses/Parameters expect more than one assumption" + | _ -> () + (* Gallina declarations *) GEXTEND Gram - GLOBAL: gallina gallina_ext reduce thm_tok theorem_body; + GLOBAL: gallina gallina_ext (* reduce *) thm_token; - theorem_body_line: - [ [ n = numarg; ":"; tac = tacarg; "." -> - <:ast< (VERNACCALL "SOLVE" $n (TACTIC $tac)) >> - | tac = tacarg; "." -> <:ast< (VERNACCALL "SOLVE" 1 (TACTIC $tac)) >> - ] ] - ; - theorem_body_line_list: - [ [ t = theorem_body_line -> [t] - | t = theorem_body_line; l = theorem_body_line_list -> t :: l ] ] - ; - theorem_body: - [ [ l = theorem_body_line_list -> - <:ast< (VERNACARGLIST ($LIST $l)) >> ] ] - ; - thm_tok: - [ [ "Theorem" -> <:ast< "THEOREM" >> - | IDENT "Lemma" -> <:ast< "LEMMA" >> - | IDENT "Fact" -> <:ast< "FACT" >> - | IDENT "Remark" -> <:ast< "REMARK" >> - | IDENT "Decl" -> <:ast< "DECL" >> ] ] - ; - def_tok: - [ [ "Definition" -> <:ast< "DEFINITION" >> - | IDENT "Local" -> <:ast< "LOCAL" >> - | IDENT "SubClass" -> <:ast< "SUBCLASS" >> - | IDENT "Local"; IDENT "SubClass" -> <:ast< "LSUBCLASS" >> ] ] - ; - hyp_tok: - [ [ "Hypothesis" -> <:ast< "HYPOTHESIS" >> - | "Variable" -> <:ast< "VARIABLE" >> ] ] - ; - hyps_tok: - [ [ IDENT "Hypotheses" -> <:ast< "HYPOTHESES" >> - | IDENT "Variables" -> <:ast< "VARIABLES" >> ] ] - ; - param_tok: - [ [ "Axiom" -> <:ast< "AXIOM" >> - | "Parameter" -> <:ast< "PARAMETER" >> ] ] - ; - params_tok: - [ [ IDENT "Parameters" -> <:ast< "PARAMETERS" >> ] ] + thm_token: + [ [ "Theorem" -> Theorem + | IDENT "Lemma" -> Lemma + | IDENT "Fact" -> Fact + | IDENT "Remark" -> Remark + | IDENT "Decl" -> Decl ] ] + ; + def_token: + [ [ "Definition" -> (fun _ _ -> ()), Definition + | IDENT "Local" -> (fun _ _ -> ()), LocalDefinition + | IDENT "SubClass" -> Class.add_subclass_hook, Definition + | IDENT "Local"; IDENT "SubClass" -> + Class.add_subclass_hook, LocalDefinition + ] ] + ; + assumption_token: + [ [ "Hypothesis" -> AssumptionHypothesis + | "Variable" -> AssumptionVariable + | "Axiom" -> AssumptionAxiom + | "Parameter" -> AssumptionParameter ] ] + ; + assumptions_token: + [ [ IDENT "Hypotheses" -> AssumptionHypothesis + | IDENT "Variables" -> AssumptionVariable + | IDENT "Parameters" -> AssumptionParameter ] ] ; params: - [ [ idl = Constr.ne_ident_comma_list; ":"; c = Constr.constr -> - <:ast< (BINDER $c ($LIST $idl)) >> ] ] + [ [ idl = LIST1 ident SEP ","; ":"; c = constr -> join_binders (idl,c) + ] ] ; ne_params_list: - [ [ id = params; ";"; idl = ne_params_list -> id :: idl - | id = params -> [id] ] ] + [ [ ll = LIST1 params SEP ";" -> List.flatten ll ] ] ; reduce: - [ [ IDENT "Eval"; r = Tactic.red_tactic; "in" -> - [ <:ast< (TACTIC_ARG (REDEXP $r)) >> ] - | -> [] ] ] + [ [ IDENT "Eval"; r = Tactic.red_tactic; "in" -> Some r + | -> None ] ] ; binders_list: - [ [ idl = Constr.ne_binders_list -> <:ast< (BINDERS ($LIST $idl)) >> + [ [ idl = ne_binders_list -> <:ast< (BINDERS ($LIST $idl)) >> | -> <:ast< (BINDERS) >> ] ] ; gallina: (* Definition, Goal *) - [ [ thm = thm_tok; id = identarg; ":"; c = constrarg -> - <:ast< (StartProof $thm $id $c) >> - | thm = thm_tok; id = identarg; ":"; c = constrarg; ":="; "Proof"; - tb = theorem_body; "Qed" -> - <:ast< (TheoremProof $thm $id $c $tb) >> - | def = def_tok; s = identarg; bl = binders_list; - ":"; t = Constr.constr -> - <:ast< (StartProof $def $s (CONSTR ($ABSTRACT "PRODLIST" $bl $t))) >> - | def = def_tok; s = identarg; bl = binders_list; - ":="; red = reduce; c = Constr.constr -> - <:ast< (DEFINITION $def $s (CONSTR ($ABSTRACT "LAMBDALIST" $bl $c)) - ($LIST $red)) >> - | def = def_tok; s = identarg; bl = binders_list; - ":="; red = reduce; c = Constr.constr; ":"; t = Constr.constr -> - <:ast< (DEFINITION $def $s (CONSTR ($ABSTRACT "LAMBDALIST" $bl $c)) - (CONSTR ($ABSTRACT "PRODLIST" $bl $t)) ($LIST $red)) >> - | def = def_tok; s = identarg; bl = binders_list; - ":"; t = Constr.constr; ":="; red = reduce; c = Constr.constr -> - <:ast< (DEFINITION $def $s (CONSTR ($ABSTRACT "LAMBDALIST" $bl $c)) - (CONSTR ($ABSTRACT "PRODLIST" $bl $t)) ($LIST $red)) >> - - | hyp = hyp_tok; bl = ne_params_list -> - <:ast< (VARIABLE $hyp (BINDERLIST ($LIST $bl))) >> - | hyp = hyps_tok; bl = ne_params_list -> - <:ast< (VARIABLE $hyp (BINDERLIST ($LIST $bl))) >> - | hyp = param_tok; bl = ne_params_list -> - <:ast< (PARAMETER $hyp (BINDERLIST ($LIST $bl))) >> - | hyp = params_tok; bl = ne_params_list -> - <:ast< (PARAMETER $hyp (BINDERLIST ($LIST $bl))) >> + [ [ thm = thm_token; id = ident; ":"; c = constr -> + VernacStartProof + (StartTheoremProof thm, Some id, c, false, (fun _ _ -> ())) + + | (f,d) = def_token; id = ident; bl = binders_list; + ":"; t = constr -> + VernacStartProof + (StartDefinitionBody d, Some id, generalize_constr bl t, false, f) + | (f,def) = def_token; s = Prim.ident; bl = binders_list; + ":="; red = reduce; c = constr -> + VernacDefinition + (def, s, red, abstract_constr bl c, None, f) + | (f,def) = def_token; s = Prim.ident; bl = binders_list; + ":="; red = reduce; c = constr; ":"; t = constr -> + VernacDefinition + (def, s, red, abstract_constr bl c, + Some (generalize_constr bl t), f) + | (f,def) = def_token; s = Prim.ident; bl = binders_list; + ":"; t = constr; ":="; red = reduce; c = constr -> + VernacDefinition + (def, s, red, abstract_constr bl c, + Some (generalize_constr bl t), f) + | stre = assumption_token; bl = ne_params_list -> + VernacAssumption (stre, bl) + | stre = assumptions_token; bl = ne_params_list -> + test_plurial_form bl; + VernacAssumption (stre, bl) ] ] ; END @@ -151,248 +177,202 @@ GEXTEND Gram GEXTEND Gram GLOBAL: gallina gallina_ext; - finite_tok: - [ [ "Inductive" -> <:ast< "Inductive" >> - | "CoInductive" -> <:ast< "CoInductive" >> ] ] + finite_token: + [ [ "Inductive" -> true + | "CoInductive" -> false ] ] ; - record_tok: - [ [ IDENT "Record" -> <:ast< "Record" >> - | IDENT "Structure" -> <:ast< "Structure" >> ] ] + record_token: + [ [ IDENT "Record" -> () | IDENT "Structure" -> () ] ] ; constructor: - [ [ id = identarg; ":"; c = Constr.constr -> - <:ast< (BINDER $c $id) >> ] ] + [ [ id = ident; coe = of_type_with_opt_coercion; c = constr -> + (id,coe,c) ] ] ; ne_constructor_list: [ [ idc = constructor; "|"; l = ne_constructor_list -> idc :: l | idc = constructor -> [idc] ] ] ; constructor_list: - [ [ "|"; l = ne_constructor_list -> <:ast< (BINDERLIST ($LIST $l)) >> - | l = ne_constructor_list -> <:ast< (BINDERLIST ($LIST $l)) >> - | -> <:ast< (BINDERLIST) >> ] ] + [ [ "|"; l = ne_constructor_list -> l + | l = ne_constructor_list -> l + | -> [] ] ] ; block_old_style: [ [ ind = oneind_old_style; "with"; indl = block_old_style -> ind :: indl | ind = oneind_old_style -> [ind] ] ] ; oneind_old_style: - [ [ id = identarg; ":"; c = constrarg; ":="; lc = constructor_list -> - <:ast< (VERNACARGLIST $id $c $lc) >> ] ] + [ [ id = ident; ":"; c = constr; ":="; lc = constructor_list -> + (id,c,lc) ] ] ; block: [ [ ind = oneind; "with"; indl = block -> ind :: indl | ind = oneind -> [ind] ] ] ; oneind: - [ [ id = identarg; indpar = indpar; ":"; c = constrarg; ":="; - lc = constructor_list - -> <:ast< (VERNACARGLIST $id $c $indpar $lc) >> ] ] + [ [ id = ident; indpar = indpar; ":"; c = constr; ":="; + lc = constructor_list -> (id,indpar,c,lc) ] ] ; indpar: - [ [ bl = ne_simple_binders_list -> - <:ast< (BINDERLIST ($LIST $bl)) >> - | -> <:ast< (BINDERLIST) >> ] ] + [ [ bl = ne_simple_binders_list -> bl + | -> [] ] ] ; opt_coercion: - [ [ ">" -> "COERCION" - | -> "" ] ] + [ [ ">" -> true + | -> false ] ] ; of_type_with_opt_coercion: - [ [ ":>" -> "COERCION" - | ":"; ">" -> "COERCION" - | ":" -> "" ] ] + [ [ ":>" -> true + | ":"; ">" -> true + | ":" -> false ] ] ; onescheme: - [ [ id = identarg; ":="; dep = dep; indid = qualidarg; IDENT "Sort"; - s = sortarg -> - <:ast< (VERNACARGLIST $id $dep $indid $s) >> ] ] + [ [ id = ident; ":="; dep = dep; ind = qualid; IDENT "Sort"; + s = sort -> (id,dep,ind,s) ] ] ; - specifscheme: - [ [ rec_ = onescheme; "with"; recl = specifscheme -> rec_ :: recl - | rec_ = onescheme -> [rec_] ] ] + schemes: + [ [ recl = LIST1 onescheme SEP "with" -> recl ] ] ; dep: - [ [ IDENT "Induction"; IDENT "for" -> <:ast< "DEP" >> - | IDENT "Minimality"; IDENT "for" -> <:ast< "NODEP" >> ] ] + [ [ IDENT "Induction"; IDENT "for" -> true + | IDENT "Minimality"; IDENT "for" -> false ] ] ; onerec: - [ [ id = identarg; idl = ne_simple_binders_list; ":"; c = constrarg; - ":="; def = constrarg -> - <:ast< (VERNACARGLIST $id (BINDERLIST ($LIST $idl)) $c $def) >> ] ] + [ [ id = ident; idl = ne_simple_binders_list; ":"; c = constr; + ":="; def = constr -> (id,idl,c,def) ] ] ; specifrec: - [ [ rec_ = onerec; "with"; recl = specifrec -> rec_ :: recl - | rec_ = onerec -> [rec_] ] ] + [ [ l = LIST1 onerec SEP "with" -> l ] ] ; onecorec: - [ [ id = identarg; ":"; c = constrarg; ":="; def = constrarg -> - <:ast< (VERNACARGLIST $id $c $def) >> ] ] + [ [ id = ident; ":"; c = constr; ":="; def = constr -> + (id,c,def) ] ] ; specifcorec: - [ [ corec = onecorec; "with"; corecl = specifcorec -> corec :: corecl - | corec = onecorec -> [corec] ] ] - ; - field: - [ [ id = identarg; oc = of_type_with_opt_coercion; c = constrarg -> - <:ast< (VERNACARGLIST ($STR $oc) "ASSUM" $id $c) >> - | id = identarg; oc = of_type_with_opt_coercion; t = Constr.constr; - ":="; b = Constr.constr -> - <:ast< (VERNACARGLIST "" "DEF" $id $b (COMMAND (CAST $b $t))) >> - | id = identarg; ":="; b = constrarg -> - <:ast< (VERNACARGLIST "" "DEF" $id $b) >> -(* | id = identarg; ":>"; c = constrarg -> - <:ast< (VERNACARGLIST "COERCION" $id $c) >> *)] ] - ; - nefields: - [ [ idc = field; ";"; fs = nefields -> idc :: fs - | idc = field -> [idc] ] ] + [ [ l = LIST1 onecorec SEP "with" -> l ] ] + ; + record_field: + [ [ id = ident; oc = of_type_with_opt_coercion; t = constr -> + (oc,AssumExpr (id,t)) + | id = ident; oc = of_type_with_opt_coercion; t = constr; + ":="; b = constr -> + (oc,DefExpr (id,b,Some t)) + | id = ident; ":="; b = constr -> + (false,DefExpr (id,b,None)) ] ] ; fields: - [ [ fs = nefields -> <:ast< (VERNACARGLIST ($LIST $fs)) >> - | -> <:ast< (VERNACARGLIST) >> ] ] + [ [ fs = LIST0 record_field SEP ";" -> fs ] ] ; simple_params: - [ [ idl = Constr.ne_ident_comma_list; ":"; c = Constr.constr -> - <:ast< (BINDER $c ($LIST $idl)) >> - | idl = Constr.ne_ident_comma_list -> - <:ast< (BINDER (ISEVAR) ($LIST $idl)) >> ] ] - ; - ne_simple_params_list: - [ [ id = simple_params; ";"; idl = ne_simple_params_list -> id :: idl - | id = simple_params -> [id] ] ] + [ [ idl = LIST1 ident SEP ","; ":"; c = constr -> join_binders (idl, c) + | idl = LIST1 ident SEP "," -> join_binders (idl, evar_constr) + ] ] ; simple_binders: - [ [ "["; bl = ne_simple_params_list; "]" -> bl ] ] + [ [ "["; bll = LIST1 simple_params SEP ";"; "]" -> List.flatten bll ] ] ; ne_simple_binders_list: - [ [ bl = simple_binders; bll = ne_simple_binders_list -> bl @ bll - | bl = simple_binders -> bl ] ] + [ [ bll = LIST1 simple_binders -> List.flatten bll ] ] ; - rec_constr: - [ [ c = Vernac_.identarg -> <:ast< (VERNACARGLIST $c) >> - | -> <:ast< (VERNACARGLIST) >> ] ] + rec_constructor: + [ [ c = ident -> Some c + | -> None ] ] ; gallina_ext: - [ [ IDENT "Mutual"; bl = ne_simple_binders_list ; f = finite_tok; + [ [ IDENT "Mutual"; bl = ne_simple_binders_list ; f = finite_token; indl = block_old_style -> - <:ast< (OLDMUTUALINDUCTIVE (BINDERLIST ($LIST $bl)) $f - (VERNACARGLIST ($LIST $indl))) >> - | record_tok; oc = opt_coercion; name = identarg; ps = indpar; ":"; - s = sortarg; ":="; c = rec_constr; "{"; fs = fields; "}" -> - <:ast< (RECORD ($STR $oc) $name $ps $s $c $fs) >> -(* | record_tok; ">"; name = identarg; ps = indpar; ":"; - s = sortarg; ":="; c = rec_constr; "{"; fs = fields; "}" -> - <:ast< (RECORD "COERCION" $name $ps $s $c $fs) >> -*) ] ] + let indl' = List.map (fun (id,ar,c) -> (id,bl,ar,c)) indl in + VernacInductive (f,indl') + | record_token; oc = opt_coercion; name = ident; ps = indpar; ":"; + s = sort; ":="; c = rec_constructor; "{"; fs = fields; "}" -> + VernacRecord (oc,name,ps,s,c,fs) + ] ] ; gallina: - [ [ IDENT "Mutual"; f = finite_tok; indl = block -> - <:ast< (MUTUALINDUCTIVE $f (VERNACARGLIST ($LIST $indl))) >> - | "Fixpoint"; recs = specifrec -> - <:ast< (MUTUALRECURSIVE (VERNACARGLIST ($LIST $recs))) >> - | "CoFixpoint"; corecs = specifcorec -> - <:ast< (MUTUALCORECURSIVE (VERNACARGLIST ($LIST $corecs))) >> - | IDENT "Scheme"; schemes = specifscheme -> - <:ast< (INDUCTIONSCHEME (VERNACARGLIST ($LIST $schemes))) >> - | f = finite_tok; s = sortarg; id = identarg; indpar = indpar; ":="; - lc = constructor_list -> - <:ast< (ONEINDUCTIVE $f $id $s $indpar $lc) >> - | f = finite_tok; indl = block -> - <:ast< (MUTUALINDUCTIVE $f (VERNACARGLIST ($LIST $indl))) >> - - | record_tok; oc = opt_coercion; name = identarg; ps = indpar; ":"; - s = sortarg; ":="; c = rec_constr; "{"; fs = fields; "}" -> - <:ast< (RECORD ($STR $oc) $name $ps $s $c $fs) >> -(* | record_tok; ">"; name = identarg; ps = indpar; ":"; - s = sortarg; ":="; c = rec_constr; "{"; fs = fields; "}" -> - <:ast< (RECORD "COERCION" $name $ps $s $c $fs) >> -*) ] ]; - + [ [ IDENT "Mutual"; f = finite_token; indl = block -> + VernacInductive (f,indl) + | "Fixpoint"; recs = specifrec -> VernacFixpoint recs + | "CoFixpoint"; corecs = specifcorec -> VernacCoFixpoint corecs + | IDENT "Scheme"; l = schemes -> VernacScheme l + | f = finite_token; s = sort; id = ident; indpar = indpar; ":="; + lc = constructor_list -> + VernacInductive (f,[id,indpar,s,lc]) + | f = finite_token; indl = block -> + VernacInductive (f,indl) ] ] + ; END GEXTEND Gram GLOBAL: gallina_ext; def_body: - [ [ - c1 = Constr.constr; ":"; c2 = Constr.constr -> - <:ast< (CONSTR (CAST $c1 $c2)) >> - | c = Constr.constr -> <:ast< (CONSTR $c) >> - ] ]; - + [ [ ":="; c = constr; ":"; t = constr -> c, Some t + | ":"; t = constr; ":="; c = constr -> c, Some t + | ":="; c = constr -> c, None ] ] + ; gallina_ext: [ [ (* Sections *) - IDENT "Section"; id = identarg -> <:ast< (BeginSection $id) >> - | IDENT "Chapter"; id = identarg -> <:ast< (BeginSection $id) >> - | IDENT "Module"; id = identarg -> <:ast< (BeginModule $id) >> - | IDENT "End"; id = identarg -> <:ast< (EndSection $id) >> + IDENT "Section"; id = ident -> VernacBeginSection id + | IDENT "Chapter"; id = ident -> VernacBeginSection id + | IDENT "Module"; id = ident -> + warning "Module is currently unsupported"; VernacNop + | IDENT "End"; id = ident -> VernacEndSection id (* Transparent and Opaque *) - | IDENT "Transparent"; l = ne_qualidarg_list -> - <:ast< (TRANSPARENT ($LIST $l)) >> - | IDENT "Opaque"; l = ne_qualidarg_list -> - <:ast< (OPAQUE ($LIST $l)) >> + | IDENT "Transparent"; l = LIST1 qualid -> VernacSetOpacity (false, l) + | IDENT "Opaque"; l = LIST1 qualid -> VernacSetOpacity (true, l) (* Canonical structure *) - | IDENT "Canonical"; IDENT "Structure"; qid = qualidarg -> - <:ast< (CANONICAL $qid) >> - | IDENT "Canonical"; IDENT "Structure"; qid = qualidarg; ":"; - t = Constr.constr; ":="; c = Constr.constr -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "OBJECT" $s (CONSTR $c) (CONSTR $t)) >> - | IDENT "Canonical"; IDENT "Structure"; qid = qualidarg; ":="; - c = Constr.constr; ":"; t = Constr.constr -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "OBJECT" $s (CONSTR $c) (CONSTR $t)) >> - | IDENT "Canonical"; IDENT "Structure"; qid = qualidarg; ":="; - c = Constr.constr -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "OBJECT" $s (CONSTR $c)) >> + | IDENT "Canonical"; IDENT "Structure"; qid = qualid -> + VernacCanonical qid + | IDENT "Canonical"; IDENT "Structure"; qid = qualid; (c,t) = def_body + -> + let s = Ast.coerce_qualid_to_id qid in + VernacDefinition (Definition,s,None,c,t,Recordobj.add_object_hook) (* Rem: LOBJECT, OBJCOERCION, LOBJCOERCION have been removed - (they were unused and undocumented *) + (they were unused and undocumented) *) (* Coercions *) - | IDENT "Coercion"; qid = qualidarg; ":="; c = def_body -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "COERCION" $s $c) >> - | IDENT "Coercion"; IDENT "Local"; qid = qualidarg; ":="; - c = constrarg -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "LCOERCION" $s $c) >> - | IDENT "Coercion"; IDENT "Local"; qid = qualidarg; ":="; - c1 = Constr.constr; ":"; c2 = Constr.constr -> - let s = Ast.coerce_to_var qid in - <:ast< (DEFINITION "LCOERCION" $s (CONSTR (CAST $c1 $c2))) >> - | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = qualidarg; - ":"; c = qualidarg; ">->"; d = qualidarg -> - <:ast< (COERCION "LOCAL" "IDENTITY" $f $c $d) >> - | IDENT "Identity"; IDENT "Coercion"; f = qualidarg; ":"; - c = qualidarg; ">->"; d = qualidarg -> - <:ast< (COERCION "" "IDENTITY" $f $c $d) >> - | IDENT "Coercion"; IDENT "Local"; qid = qualidarg; ":"; c = qualidarg; - ">->"; d = qualidarg -> - <:ast< (COERCION "LOCAL" "" $qid $c $d) >> - | IDENT "Coercion"; qid = qualidarg; ":"; c = qualidarg; ">->"; - d = qualidarg -> <:ast< (COERCION "" "" $qid $c $d) >> - | IDENT "Class"; IDENT "Local"; c = qualidarg -> - <:ast< (CLASS "LOCAL" $c) >> - | IDENT "Class"; c = qualidarg -> <:ast< (CLASS "" $c) >> + | IDENT "Coercion"; qid = qualid; (c,t) = def_body -> + let s = Ast.coerce_qualid_to_id qid in + VernacDefinition (Definition,s,None,c,t,Class.add_coercion_hook) + | IDENT "Coercion"; IDENT "Local"; qid = qualid; (c,t) = def_body -> + let s = Ast.coerce_qualid_to_id qid in + VernacDefinition (LocalDefinition,s,None,c,t,Class.add_coercion_hook) + | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = Prim.ident; + ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> + VernacIdentityCoercion (Declare.make_strength_0 (), f, s, t) + | IDENT "Identity"; IDENT "Coercion"; f = Prim.ident; ":"; + s = class_rawexpr; ">->"; t = class_rawexpr -> + VernacIdentityCoercion (Nametab.NeverDischarge, f, s, t) + | IDENT "Coercion"; IDENT "Local"; qid = qualid; ":"; + s = class_rawexpr; ">->"; t = class_rawexpr -> + VernacCoercion (Declare.make_strength_0 (), qid, s, t) + | IDENT "Coercion"; qid = qualid; ":"; s = class_rawexpr; ">->"; + t = class_rawexpr -> + VernacCoercion (Nametab.NeverDischarge, qid, s, t) + | IDENT "Class"; IDENT "Local"; c = qualid -> + Pp.warning "Class is obsolete"; VernacNop + | IDENT "Class"; c = qualid -> + Pp.warning "Class is obsolete"; VernacNop (* Implicit *) - | IDENT "Syntactic"; "Definition"; id = identarg; ":="; com = constrarg - -> <:ast< (SyntaxMacro $id $com) >> - | IDENT "Syntactic"; "Definition"; id = identarg; ":="; com = constrarg; - "|"; n = numarg -> <:ast< (SyntaxMacro $id $com $n) >> + | IDENT "Syntactic"; "Definition"; id = ident; ":="; c = constr; + n = OPT [ "|"; n = natural -> n ] -> + VernacSyntacticDefinition (id,c,n) + + | IDENT "Implicits"; qid = qualid; "["; l = LIST0 natural; "]" -> + VernacDeclareImplicits (qid,Some l) + | IDENT "Implicits"; qid = qualid -> VernacDeclareImplicits (qid,None) + + (* For compatibility *) | IDENT "Implicit"; IDENT "Arguments"; IDENT "On" -> - <:ast< (IMPLICIT_ARGS_ON) >> + VernacSetOption + (Goptions.SecondaryTable ("Implicit","Arguments"), BoolValue true) | IDENT "Implicit"; IDENT "Arguments"; IDENT "Off" -> - <:ast< (IMPLICIT_ARGS_OFF) >> - | IDENT "Implicits"; qid = qualidarg; "["; l = numarg_list; "]" -> - <:ast< (IMPLICITS "" $qid ($LIST $l)) >> - | IDENT "Implicits"; qid = qualidarg -> - <:ast< (IMPLICITS "Auto" $qid) >> + VernacSetOption + (Goptions.SecondaryTable ("Implicit","Arguments"), BoolValue false) ] ] ; END @@ -401,20 +381,20 @@ END GEXTEND Gram GLOBAL: command; - import_tok: - [ [ IDENT "Import" -> <:ast< "IMPORT" >> - | IDENT "Export" -> <:ast< "EXPORT" >> - | -> <:ast< "IMPORT" >> ] ] + export_token: + [ [ IDENT "Import" -> false + | IDENT "Export" -> true + | -> false ] ] ; - specif_tok: - [ [ IDENT "Implementation" -> <:ast< "IMPLEMENTATION" >> - | IDENT "Specification" -> <:ast< "SPECIFICATION" >> - | -> <:ast< "UNSPECIFIED" >> ] ] + specif_token: + [ [ IDENT "Implementation" -> Some false + | IDENT "Specification" -> Some true + | -> None ] ] ; command: - [ [ "Load"; verbosely = [ IDENT "Verbose" -> "Verbose" | -> "" ]; + [ [ "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ]; s = [ s = STRING -> s | s = IDENT -> s ] -> - <:ast< (LoadFile ($STR $verbosely) ($STR $s)) >> + VernacLoad (verbosely, s) (* | "Compile"; verbosely = [ IDENT "Verbose" -> "Verbose" @@ -424,29 +404,28 @@ GEXTEND Gram [ IDENT "Specification" -> "Specification" | -> "" ]; mname = [ s = STRING -> s | s = IDENT -> s ]; - fname = OPT [ s = STRING -> s | s = IDENT -> s ] -> + fname = OPT [ s = STRING -> s | s = IDENT -> s ] -> ExtraVernac let fname = match fname with Some s -> s | None -> mname in <:ast< (CompileFile ($STR $verbosely) ($STR $only_spec) ($STR $mname) ($STR $fname))>> *) - | IDENT "Read"; IDENT "Module"; qidl = LIST1 qualidarg -> - <:ast< (ReadModule ($LIST $qidl)) >> - | IDENT "Require"; import = import_tok; specif = specif_tok; - qidl = LIST1 qualidarg -> - <:ast< (Require $import $specif ($LIST $qidl)) >> - | IDENT "Require"; import = import_tok; specif = specif_tok; - qid = qualidarg; filename = stringarg -> - <:ast< (RequireFrom $import $specif $qid $filename) >> - | IDENT "Write"; IDENT "Module"; id = identarg -> + | IDENT "Read"; IDENT "Module"; qidl = LIST1 qualid -> + VernacRequire (None, None, qidl) + | IDENT "Require"; export = export_token; specif = specif_token; + qidl = LIST1 qualid -> VernacRequire (Some export, specif, qidl) + | IDENT "Require"; export = export_token; specif = specif_token; + id = Prim.ident; filename = STRING -> + VernacRequireFrom (export, specif, id, filename) +(* + | IDENT "Write"; IDENT "Module"; id = identarg -> ExtraVernac <:ast< (WriteModule $id) >> - | IDENT "Write"; IDENT "Module"; id = identarg; s = stringarg -> + | IDENT "Write"; IDENT "Module"; id = identarg; s = stringarg -> ExtraVernac <:ast< (WriteModule $id $s) >> - | IDENT "Declare"; IDENT "ML"; IDENT "Module"; - l = ne_stringarg_list -> <:ast< (DeclareMLModule ($LIST $l)) >> - | IDENT "Import"; qidl = ne_qualidarg_list -> - <:ast< (ImportModule ($LIST $qidl)) >> - | IDENT "Export"; qidl = ne_qualidarg_list -> - <:ast< (ExportModule ($LIST $qidl)) >> +*) + | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 STRING -> + VernacDeclareMLModule l + | IDENT "Import"; qidl = LIST1 qualid -> VernacImport (false,qidl) + | IDENT "Export"; qidl = LIST1 qualid -> VernacImport (true,qidl) ] ] ; @@ -459,24 +438,20 @@ GEXTEND Gram [ [ (* State management *) - IDENT "Write"; IDENT "State"; id = identarg -> - <:ast< (WriteState $id) >> - | IDENT "Write"; IDENT "State"; s = stringarg -> - <:ast< (WriteState $s) >> - | IDENT "Restore"; IDENT "State"; id = identarg -> - <:ast< (RestoreState $id) >> - | IDENT "Restore"; IDENT "State"; s = stringarg -> - <:ast< (RestoreState $s) >> - | IDENT "Reset"; id = identarg -> <:ast< (ResetName $id) >> - | IDENT "Reset"; IDENT "Initial" -> <:ast< (ResetInitial) >> - | IDENT "Reset"; IDENT "Section"; id = identarg -> - <:ast< (ResetSection $id) >> - | IDENT "Back" -> <:ast<(Back)>> - | IDENT "Back"; n = numarg -> <:ast<(Back $n)>> + IDENT "Write"; IDENT "State"; s = IDENT -> VernacWriteState s + | IDENT "Write"; IDENT "State"; s = STRING -> VernacWriteState s + | IDENT "Restore"; IDENT "State"; s = IDENT -> VernacRestoreState s + | IDENT "Restore"; IDENT "State"; s = STRING -> VernacRestoreState s + +(* Resetting *) + | IDENT "Reset"; id = Prim.ident -> VernacResetName id + | IDENT "Reset"; IDENT "Initial" -> VernacResetInitial + | IDENT "Back" -> VernacBack 1 + | IDENT "Back"; n = Prim.natural -> VernacBack n (* Tactic Debugger *) - | IDENT "Debug"; IDENT "On" -> <:ast< (DebugOn) >> - | IDENT "Debug"; IDENT "Off" -> <:ast< (DebugOff) >> + | IDENT "Debug"; IDENT "On" -> VernacDebug true + | IDENT "Debug"; IDENT "Off" -> VernacDebug false ] ]; END diff --git a/parsing/g_zsyntax.ml b/parsing/g_zsyntax.ml index f6a836942..c4e94695d 100644 --- a/parsing/g_zsyntax.ml +++ b/parsing/g_zsyntax.ml @@ -14,6 +14,7 @@ open Pp open Util open Names open Ast +open Extend let get_z_sign loc = let ast_of_id id = Astterm.globalize_constr (Nvar(loc,id)) in @@ -22,10 +23,8 @@ let get_z_sign loc = ast_of_id (id_of_string "xH")), (ast_of_id (id_of_string "ZERO"), ast_of_id (id_of_string "POS"), - ast_of_id (id_of_string "NEG")), - (ast_of_id (id_of_string "My_special_variable0"), - ast_of_id (id_of_string "My_special_variable1"))) - + ast_of_id (id_of_string "NEG"))) + let int_array_of_string s = let a = Array.create (String.length s) 0 in for i = 0 to String.length s - 1 do @@ -85,7 +84,7 @@ let pos_of_int_array astxI astxO astxH x = pos_of x let z_of_string pos_or_neg s dloc = - let ((astxI,astxO,astxH),(astZERO,astPOS,astNEG),_) = get_z_sign dloc in + let ((astxI,astxO,astxH),(astZERO,astPOS,astNEG)) = get_z_sign dloc in let v = int_array_of_string s in if is_nonzero v then if pos_or_neg then @@ -112,8 +111,11 @@ let int_array_option_of_pos astxI astxO astxH p = with Non_closed_number -> None -let inside_printer posneg std_pr p = - let ((astxI,astxO,astxH),_,(myvar0,myvar1)) = get_z_sign dummy_loc in +let pr_pos a = hov 0 (str "POS" ++ brk (1,1) ++ a) +let pr_neg a = hov 0 (str "NEG" ++ brk (1,1) ++ a) + +let inside_printer posneg std_pr p = + let ((astxI,astxO,astxH),_) = get_z_sign dummy_loc in match (int_array_option_of_pos astxI astxO astxH p) with | Some n -> if posneg then @@ -121,11 +123,11 @@ let inside_printer posneg std_pr p = else (str "(-" ++ str (string_of_int_array n) ++ str ")") | None -> - let c = if posneg then myvar0 else myvar1 in - std_pr (ope("ZEXPR",[ope("APPLIST",[c; p])])) + let pr = if posneg then pr_pos else pr_neg in + str "(" ++ pr (std_pr (ope("ZEXPR",[p]))) ++ str ")" -let outside_printer posneg std_pr p = - let ((astxI,astxO,astxH),_,(myvar0,myvar1)) = get_z_sign dummy_loc in +let outside_printer posneg std_pr p = + let ((astxI,astxO,astxH),_) = get_z_sign dummy_loc in match (int_array_option_of_pos astxI astxO astxH p) with | Some n -> if posneg then @@ -133,8 +135,8 @@ let outside_printer posneg std_pr p = else (str "`-" ++ str (string_of_int_array n) ++ str "`") | None -> - let c = if posneg then myvar0 else myvar1 in - (str "(" ++ std_pr (ope("APPLIST", [c; p])) ++ str ")") + let pr = if posneg then pr_pos else pr_neg in + str "(" ++ pr (std_pr p) ++ str ")" (* Declare pretty-printers for integers *) let _ = Esyntax.Ppprim.add ("positive_printer", (outside_printer true)) @@ -145,15 +147,9 @@ let _ = Esyntax.Ppprim.add ("negative_printer_inside", (inside_printer false)) (* Declare the primitive parser *) open Pcoq -let number = - match create_entry (get_univ "znatural") "number" ETast with - | Ast n -> n - | _ -> anomaly "G_zsyntax.number" +let number = create_constr_entry (get_univ "znatural") "number" -let negnumber = - match create_entry (get_univ "znatural") "negnumber" ETast with - | Ast n -> n - | _ -> anomaly "G_zsyntax.negnumber" +let negnumber = create_constr_entry (get_univ "znatural") "negnumber" let _ = Gram.extend number None diff --git a/parsing/genarg.ml b/parsing/genarg.ml new file mode 100644 index 000000000..e0d3b8019 --- /dev/null +++ b/parsing/genarg.ml @@ -0,0 +1,181 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l) + | _ -> failwith "Genarg: not a list0" + +let fold_list1 f = function + | (List1ArgType t as u, l) -> + List.fold_right (fun x -> f (in_gen t x)) (Obj.magic l) + | _ -> failwith "Genarg: not a list1" + +let fold_opt f a = function + | (OptArgType t as u, l) -> + (match Obj.magic l with + | None -> a + | Some x -> f (in_gen t x)) + | _ -> failwith "Genarg: not a opt" + +let fold_pair f = function + | (PairArgType (t1,t2) as u, l) -> + let (x1,x2) = Obj.magic l in + f (in_gen t1 x1) (in_gen t2 x2) + | _ -> failwith "Genarg: not a pair" + +let app_list0 f = function + | (List0ArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not a list0" + +let app_list1 f = function + | (List1ArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (List.map (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not a list1" + +let app_opt f = function + | (OptArgType t as u, l) -> + let o = Obj.magic l in + (u, Obj.repr (option_app (fun x -> out_gen t (f (in_gen t x))) o)) + | _ -> failwith "Genarg: not an opt" + +let app_pair f1 f2 = function + | (PairArgType (t1,t2) as u, l) -> + let (o1,o2) = Obj.magic l in + let o1 = out_gen t1 (f1 (in_gen t1 o1)) in + let o2 = out_gen t2 (f2 (in_gen t2 o2)) in + (u, Obj.repr (o1,o2)) + | _ -> failwith "Genarg: not a pair" + +let or_var_app f = function + | ArgArg x -> ArgArg (f x) + | ArgVar _ as x -> x + +let smash_var_app t f g = function + | ArgArg x -> f x + | ArgVar (_,id) -> + let (u, _ as x) = g id in + if t <> u then failwith "Genarg: a variable bound to a wrong type"; + x + +let unquote x = x + +type an_arg_of_this_type = Obj.t + +let in_generic t x = (t, Obj.repr x) diff --git a/parsing/genarg.mli b/parsing/genarg.mli new file mode 100644 index 000000000..70e67330e --- /dev/null +++ b/parsing/genarg.mli @@ -0,0 +1,208 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* rawtype ----> rawconstr generic_argument ----> + | + | interp + V + type <---- constr generic_argument <---- + out in + +To distinguish between the uninterpreted (raw) and the interpreted +worlds, we annotate the type generic_argument by a phantom argument +which is either constr_ast or constr (actually we add also a second +argument raw_tactic_expr and tactic, but this is only for technical +reasons, because these types are undefined at the type of compilation +of Genarg). + +Transformation for each type : +tag f raw open type cooked closed type + +BoolArgType bool bool +IntArgType int int +IntOrVarArgType int or_var int +StringArgType string (parsed w/ "") string +IdentArgType identifier identifier +PreIdentArgType string (parsed w/o "") string +QualidArgType qualid located global_reference +ConstrArgType constr_ast constr +ConstrMayEvalArgType constr_ast may_eval constr +QuantHypArgType quantified_hypothesis quantified_hypothesis +TacticArgType raw_tactic_expr tactic +CastedOpenConstrArgType constr_ast open_constr +ConstrWithBindingsArgType constr_ast with_bindings constr with_bindings +List0ArgType of argument_type +List1ArgType of argument_type +OptArgType of argument_type +ExtraArgType of string '_a '_b +*) + +type ('a,'co,'ta) abstract_argument_type + +val rawwit_bool : (bool,'co,'ta) abstract_argument_type +val wit_bool : (bool,'co,'ta) abstract_argument_type + +val rawwit_int : (int,'co,'ta) abstract_argument_type +val wit_int : (int,'co,'ta) abstract_argument_type + +val rawwit_int_or_var : (int or_var,'co,'ta) abstract_argument_type +val wit_int_or_var : (int or_var,'co,'ta) abstract_argument_type + +val rawwit_string : (string,'co,'ta) abstract_argument_type +val wit_string : (string,'co,'ta) abstract_argument_type + +val rawwit_ident : (identifier,'co,'ta) abstract_argument_type +val wit_ident : (identifier,'co,'ta) abstract_argument_type + +val rawwit_pre_ident : (string,'co,'ta) abstract_argument_type +val wit_pre_ident : (string,'co,'ta) abstract_argument_type + +val rawwit_qualid : (qualid located,constr_ast,'ta) abstract_argument_type +val wit_qualid : (Nametab.global_reference,constr,'ta) abstract_argument_type + +val rawwit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type +val wit_quant_hyp : (quantified_hypothesis,'co,'ta) abstract_argument_type + +val rawwit_constr : (constr_ast,constr_ast,'ta) abstract_argument_type +val wit_constr : (constr,constr,'ta) abstract_argument_type + +val rawwit_constr_may_eval : (constr_ast may_eval,constr_ast,'ta) abstract_argument_type +val wit_constr_may_eval : (constr,constr,'ta) abstract_argument_type + +val rawwit_casted_open_constr : (open_rawconstr,constr_ast,'ta) abstract_argument_type +val wit_casted_open_constr : (open_constr,constr,'ta) abstract_argument_type + +val rawwit_constr_with_bindings : (constr_ast with_bindings,constr_ast,'ta) abstract_argument_type +val wit_constr_with_bindings : (constr with_bindings,constr,'ta) abstract_argument_type + +val rawwit_red_expr : ((constr_ast,Nametab.qualid or_metanum) red_expr_gen,constr_ast,'ta) abstract_argument_type +val wit_red_expr : ((constr,Closure.evaluable_global_reference) red_expr_gen,constr,'ta) abstract_argument_type + +(* TODO: transformer tactic en extra arg *) +val rawwit_tactic : ('ta,constr_ast,'ta) abstract_argument_type +val wit_tactic : ('ta,constr,'ta) abstract_argument_type + +val wit_list0 : + ('a,'co,'ta) abstract_argument_type -> ('a list,'co,'ta) abstract_argument_type + +val wit_list1 : + ('a,'co,'ta) abstract_argument_type -> ('a list,'co,'ta) abstract_argument_type + +val wit_opt : + ('a,'co,'ta) abstract_argument_type -> ('a option,'co,'ta) abstract_argument_type + +val wit_pair : + ('a,'co,'ta) abstract_argument_type -> + ('b,'co,'ta) abstract_argument_type -> + ('a * 'b,'co,'ta) abstract_argument_type + +(* 'a generic_argument = (Sigma t:type. t[constr/'a]) *) +type ('a,'b) generic_argument + +val fold_list0 : + (('a,'b) generic_argument -> 'c -> 'c) -> ('a,'b) generic_argument -> 'c -> 'c + +val fold_list1 : + (('a,'b) generic_argument -> 'c -> 'c) -> ('a,'b) generic_argument -> 'c -> 'c + +val fold_opt : + (('a,'b) generic_argument -> 'c) -> 'c -> ('a,'b) generic_argument -> 'c + +val fold_pair : + (('a,'b) generic_argument -> ('a,'b) generic_argument -> 'c) -> + ('a,'b) generic_argument -> 'c + +(* [app_list0] fails if applied to an argument not of tag [List0 t] + for some [t]; it's the responsability of the caller to ensure it *) + +val app_list0 : (('a,'b) generic_argument -> ('c,'d) generic_argument) -> +('a,'b) generic_argument -> ('c,'d) generic_argument + +val app_list1 : (('a,'b) generic_argument -> ('c,'d) generic_argument) -> +('a,'b) generic_argument -> ('c,'d) generic_argument + +val app_opt : (('a,'b) generic_argument -> ('c,'d) generic_argument) -> +('a,'b) generic_argument -> ('c,'d) generic_argument + +val app_pair : + (('a,'b) generic_argument -> ('c,'d) generic_argument) -> + (('a,'b) generic_argument -> ('c,'d) generic_argument) + -> ('a,'b) generic_argument -> ('c,'d) generic_argument + +(* Manque l'ordre suprieur, on aimerait ('co,'ta) 'a; manque aussi le + polymorphism, on aimerait que 'b et 'c restent polymorphes l'appel + de create *) +val create_arg : string -> + ('rawa,'rawco,'rawta) abstract_argument_type + * ('a,'co,'ta) abstract_argument_type + +val exists_argtype : string -> bool + +type argument_type = + | BoolArgType + | IntArgType + | IntOrVarArgType + | StringArgType + | PreIdentArgType + | IdentArgType + | QualidArgType + | ConstrArgType + | ConstrMayEvalArgType + | QuantHypArgType + | TacticArgType + | CastedOpenConstrArgType + | ConstrWithBindingsArgType + | RedExprArgType + | List0ArgType of argument_type + | List1ArgType of argument_type + | OptArgType of argument_type + | PairArgType of argument_type * argument_type + | ExtraArgType of string + +val genarg_tag : ('a,'b) generic_argument -> argument_type + +val unquote : ('a,'co,'ta) abstract_argument_type -> argument_type + +(* We'd like + + [in_generic: !b:type, !a:argument_type -> (f a) -> b generic_argument] + + with f a = b if a is Constr, f a = c if a is Tactic, otherwise f a = |a| + + in_generic is not typable; we replace the second argument by an absurd + type (with no introduction rule) +*) +type an_arg_of_this_type + +val in_generic : + argument_type -> an_arg_of_this_type -> ('a,'b) generic_argument + +val in_gen : + ('a,'co,'ta) abstract_argument_type -> 'a -> ('co,'ta) generic_argument +val out_gen : + ('a,'co,'ta) abstract_argument_type -> ('co,'ta) generic_argument -> 'a + diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4 index 717d509e8..10f6d67b7 100644 --- a/parsing/pcoq.ml4 +++ b/parsing/pcoq.ml4 @@ -10,6 +10,10 @@ open Pp open Util +open Ast +open Genarg +open Tacexpr +open Vernacexpr (* The lexer of Coq *) @@ -43,9 +47,62 @@ let grammar_delete e rls = List.iter (fun (pil,_) -> G.delete_rule e pil) (List.rev lev)) (List.rev rls) -type typed_entry = - | Ast of Coqast.t G.Entry.e - | ListAst of Coqast.t list G.Entry.e +module type Gramobj = +sig + type grammar_object + type 'a entry + + val in_entry : 'a -> 'b G.Entry.e -> 'a entry + val object_of_entry : 'a entry -> grammar_object G.Entry.e + val type_of_entry : 'a entry -> 'a +end + +module Gramobj : Gramobj = +struct + type grammar_object = Obj.t + type 'a entry = 'a * grammar_object G.Entry.e + + let in_entry t e = (t,Obj.magic e) + let object_of_entry (t,e) = e + let type_of_entry (t,e) = t +end + +type grammar_object = Gramobj.grammar_object +let in_typed_entry = Gramobj.in_entry +let type_of_typed_entry = Gramobj.type_of_entry +let object_of_typed_entry = Gramobj.object_of_entry +type typed_entry = entry_type Gramobj.entry + +module type Gramtypes = +sig + val inAstListType : Coqast.t list G.Entry.e -> typed_entry + val inTacticAtomAstType : raw_atomic_tactic_expr G.Entry.e -> typed_entry + val inThmTokenAstType : theorem_kind G.Entry.e -> typed_entry + val inDynamicAstType : typed_ast G.Entry.e -> typed_entry + val inReferenceAstType : reference_expr G.Entry.e -> typed_entry + val inPureAstType : constr_ast G.Entry.e -> typed_entry + val inGenAstType : 'a raw_abstract_argument_type -> + 'a G.Entry.e -> typed_entry + val outGenAstType : 'a raw_abstract_argument_type -> typed_entry -> 'a G.Entry.e +end + +module Gramtypes : Gramtypes = +struct + let inAstListType = in_typed_entry AstListType + let inTacticAtomAstType = in_typed_entry TacticAtomAstType + let inThmTokenAstType = in_typed_entry ThmTokenAstType + let inDynamicAstType = in_typed_entry DynamicAstType + let inReferenceAstType = in_typed_entry ReferenceAstType + let inPureAstType = in_typed_entry (GenAstType ConstrArgType) + let inGenAstType rawwit = in_typed_entry (GenAstType (unquote rawwit)) + + let outGenAstType (a:'a raw_abstract_argument_type) o = + if type_of_typed_entry o <> GenAstType (unquote a) + then anomaly "outGenAstType: wrong type"; + Obj.magic (object_of_typed_entry o) +end + +open Gramtypes type ext_kind = | ByGrammar of @@ -81,22 +138,21 @@ module Gram = let grammar_extend te pos rls = camlp4_state := ByGrammar (te,pos,rls) :: !camlp4_state; - match te with - | Ast e -> G.extend e pos rls - | ListAst e -> G.extend e pos rls + let a = !Gramext.warning_verbose in + Gramext.warning_verbose := Options.is_verbose (); + G.extend (object_of_typed_entry te) pos rls; + Gramext.warning_verbose := a (* n is the number of extended entries (not the number of Grammar commands!) to remove. *) +let remove_grammar rls te = grammar_delete (object_of_typed_entry te) rls + let rec remove_grammars n = if n>0 then (match !camlp4_state with | [] -> anomaly "Pcoq.remove_grammars: too many rules to remove" - | ByGrammar(Ast e,_,rls)::t -> - grammar_delete e rls; - camlp4_state := t; - remove_grammars (n-1) - | ByGrammar(ListAst e,_,rls)::t -> - grammar_delete e rls; + | ByGrammar(g,_,rls)::t -> + remove_grammar rls g; camlp4_state := t; remove_grammars (n-1) | ByGEXTEND (undo,redo)::t -> @@ -129,44 +185,28 @@ let map_entry f en = let parse_string f x = let strm = Stream.of_string x in Gram.Entry.parse f (Gram.parsable strm) +(* let slam_ast (_,fin) id ast = match id with | Coqast.Nvar (loc, s) -> Coqast.Slam (loc, Some s, ast) | Coqast.Nmeta (loc, s) -> Coqast.Smetalam (loc, s, ast) | _ -> invalid_arg "Ast.slam_ast" +*) -(* This is to interpret the macro $ABSTRACT used in binders *) -(* $ABSTRACT should occur in this configuration : *) -(* ($ABSTRACT name (s1 a1 ($LIST l1)) ... (s2 an ($LIST ln)) b) *) -(* where li is id11::...::id1p1 and it produces the ast *) -(* (s1' a1 [id11]...[id1p1](... (sn' an [idn1]...[idnpn]b)...)) *) -(* where s1' is overwritten by name if s1 is $BINDER otherwise s1 *) - -let abstract_binder_ast (_,fin as loc) name a b = - match a with - | Coqast.Node((deb,_),s,d::l) -> - let s' = if s="BINDER" then name else s in - Coqast.Node((deb,fin),s', [d; List.fold_right (slam_ast loc) l b]) - | _ -> invalid_arg "Bad usage of $ABSTRACT macro" - -let abstract_binders_ast loc name a b = - match a with - | Coqast.Node(_,"BINDERS",l) -> - List.fold_right (abstract_binder_ast loc name) l b - | _ -> invalid_arg "Bad usage of $ABSTRACT macro" - -type entry_type = ETast | ETastl - +(* let entry_type ast = match ast with | Coqast.Id (_, "LIST") -> ETastl | Coqast.Id (_, "AST") -> ETast | _ -> invalid_arg "Ast.entry_type" +*) -let type_of_entry e = - match e with - | Ast _ -> ETast - | ListAst _ -> ETastl +(* +let entry_type ast = + match ast with + | AstListType -> ETastl + | _ -> ETast +*) type gram_universe = (string, typed_entry) Hashtbl.t @@ -177,35 +217,61 @@ let trace = ref false let univ_tab = Hashtbl.create 7 -let get_univ s = - try - Hashtbl.find univ_tab s - with Not_found -> +let create_univ s = + let u = s, Hashtbl.create 29 in Hashtbl.add univ_tab s u; u + +let uprim = create_univ "prim" +let uconstr = create_univ "constr" +let utactic = create_univ "tactic" +let uvernac = create_univ "vernac" + +let create_univ_if_new s = + (* compatibilite *) + let s = if s = "command" then (warning "'command' grammar universe is obsolete; use name 'constr' instead"; "constr") else s in + try + Hashtbl.find univ_tab s + with Not_found -> if !trace then begin Printf.eprintf "[Creating univ %s]\n" s; flush stderr; () end; let u = s, Hashtbl.create 29 in Hashtbl.add univ_tab s u; u - + +let get_univ = create_univ_if_new + let get_entry (u, utab) s = - try + try Hashtbl.find utab s with Not_found -> errorlabstrm "Pcoq.get_entry" (str "unknown grammar entry " ++ str u ++ str ":" ++ str s) - + let new_entry etyp (u, utab) s = let ename = u ^ ":" ^ s in - let e = - match etyp with - | ETast -> Ast (Gram.Entry.create ename) - | ETastl -> ListAst (Gram.Entry.create ename) - in + let e = in_typed_entry etyp (Gram.Entry.create ename) in Hashtbl.add utab s e; e - + +let entry_type (u, utab) s = + try + let e = Hashtbl.find utab s in + Some (type_of_typed_entry e) + with Not_found -> None + +let get_entry_type (u,n) = type_of_typed_entry (get_entry (get_univ u) n) + +let create_entry_if_new (u, utab) s etyp = + try + if type_of_typed_entry (Hashtbl.find utab s) <> etyp then + failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type") + with Not_found -> + if !trace then begin + Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; () + end; + let _ = new_entry etyp (u, utab) s in () + let create_entry (u, utab) s etyp = try let e = Hashtbl.find utab s in - if type_of_entry e <> etyp then + if type_of_typed_entry e <> etyp then failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type"); e with Not_found -> @@ -213,11 +279,46 @@ let create_entry (u, utab) s etyp = Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; () end; new_entry etyp (u, utab) s - + +let create_constr_entry u s = + outGenAstType rawwit_constr (create_entry u s (GenAstType ConstrArgType)) + +let create_generic_entry s wit = + let (u,utab) = utactic in + let etyp = unquote wit in + try + let e = Hashtbl.find utab s in + if type_of_typed_entry e <> GenAstType etyp then + failwith ("Entry " ^ u ^ ":" ^ s ^ " already exists with another type"); + outGenAstType wit e + with Not_found -> + if !trace then begin + Printf.eprintf "[Creating entry %s:%s]\n" u s; flush stderr; () + end; + let e = Gram.Entry.create s in + Hashtbl.add utab s (inGenAstType wit e); e + +let get_generic_entry s = + let (u,utab) = utactic in + try + object_of_typed_entry (Hashtbl.find utab s) + with Not_found -> + error ("unknown grammar entry "^u^":"^s) + +let get_generic_entry_type (u,utab) s = + try + match type_of_typed_entry (Hashtbl.find utab s) with + | GenAstType t -> t + | _ -> error "Not a generic type" + with Not_found -> + error ("unknown grammar entry "^u^":"^s) + let force_entry_type (u, utab) s etyp = try let entry = Hashtbl.find utab s in - let extyp = type_of_entry entry in + let extyp = type_of_typed_entry entry in + if etyp = PureAstType && extyp = GenAstType ConstrArgType then + entry else if etyp = extyp then entry else begin @@ -232,190 +333,110 @@ let force_entry_type (u, utab) s etyp = (* Grammar entries *) +let make_entry (u,univ) in_fun s = + let e = Gram.Entry.create (u ^ ":" ^ s) in + Hashtbl.add univ s (in_fun e); e + +let make_gen_entry u rawwit = make_entry u (inGenAstType rawwit) + +module Prim = + struct + let gec_gen x = make_gen_entry uprim x + let gec = make_entry uprim inPureAstType + let gec_list = make_entry uprim inAstListType + + let preident = gec_gen rawwit_pre_ident "preident" + let ident = gec_gen rawwit_ident "ident" + let rawident = Gram.Entry.create "Prim.rawident" + let natural = gec_gen rawwit_int "natural" + let integer = gec_gen rawwit_int "integer" + let string = gec_gen rawwit_string "string" + let qualid = gec_gen rawwit_qualid "qualid" + let reference = make_entry uprim inReferenceAstType "reference" + let dirpath = Gram.Entry.create "Prim.dirpath" + let astpat = make_entry uprim inDynamicAstType "astpat" + let ast = gec "ast" + let astlist = gec_list "astlist" + let ast_eoi = eoi_entry ast + let astact = gec "astact" + let metaident = gec "metaident" + let numarg = gec "numarg" + let var = gec "var" + end + + module Constr = struct - let uconstr = snd (get_univ "constr") - let gec s = - let e = Gram.Entry.create ("Constr." ^ s) in - Hashtbl.add uconstr s (Ast e); e - - let gec_list s = - let e = Gram.Entry.create ("Constr." ^ s) in - Hashtbl.add uconstr s (ListAst e); e - - let constr = gec "constr" - let constr0 = gec "constr0" - let constr1 = gec "constr1" - let constr2 = gec "constr2" - let constr3 = gec "constr3" - let lassoc_constr4 = gec "lassoc_constr4" - let constr5 = gec "constr5" - let constr6 = gec "constr6" - let constr7 = gec "constr7" - let constr8 = gec "constr8" - let constr9 = gec "constr9" - let constr91 = gec "constr91" - let constr10 = gec "constr10" + let gec = make_entry uconstr inPureAstType + let gec_constr = make_gen_entry uconstr rawwit_constr + let gec_list = make_entry uconstr inAstListType + + let constr = gec_constr "constr" + let constr0 = gec_constr "constr0" + let constr1 = gec_constr "constr1" + let constr2 = gec_constr "constr2" + let constr3 = gec_constr "constr3" + let lassoc_constr4 = gec_constr "lassoc_constr4" + let constr5 = gec_constr "constr5" + let constr6 = gec_constr "constr6" + let constr7 = gec_constr "constr7" + let constr8 = gec_constr "constr8" + let constr9 = gec_constr "constr9" + let constr91 = gec_constr "constr91" + let constr10 = gec_constr "constr10" let constr_eoi = eoi_entry constr - let lconstr = gec "lconstr" + let lconstr = gec_constr "lconstr" let ident = gec "ident" - let qualid = gec_list "qualid" + let qualid = gec "qualid" let global = gec "global" let ne_ident_comma_list = gec_list "ne_ident_comma_list" let ne_constr_list = gec_list "ne_constr_list" - let sort = gec "sort" + let sort = gec_constr "sort" let pattern = gec "pattern" + let constr_pattern = gec "constr_pattern" let ne_binders_list = gec_list "ne_binders_list" - - let uconstr = snd (get_univ "constr") end module Tactic = struct - let utactic = snd (get_univ "tactic") - let gec s = - let e = Gram.Entry.create ("Tactic." ^ s) in - Hashtbl.add utactic s (Ast e); e - - let gec_list s = - let e = Gram.Entry.create ("Tactic." ^ s) in - Hashtbl.add utactic s (ListAst e); e - - let autoargs = gec_list "autoargs" - let binding_list = gec "binding_list" - let castedopenconstrarg = gec "castedopenconstrarg" - let castedconstrarg = gec "castedconstrarg" - let com_binding_list = gec_list "com_binding_list" - let constrarg = gec "constrarg" - let constrarg_binding_list = gec_list "constrarg_binding_list" - let numarg_binding_list = gec_list "numarg_binding_list" - let lconstrarg_binding_list = gec_list "lconstrarg_binding_list" - let constrarg_list = gec_list "constrarg_list" - let ne_constrarg_list = gec_list "ne_constrarg_list" - let ident_or_numarg = gec "ident_or_numarg" - let ident_or_constrarg = gec "ident_or_constrarg" - let identarg = gec "identarg" - let hypident = gec "hypident" - let idmeta_arg = gec "idmeta_arg" - let idmetahyp = gec "idmetahyp" - let qualidarg = gec "qualidarg" - let input_fun = gec "input_fun" - let lconstrarg = gec "lconstrarg" - let let_clause = gec "let_clause" - let letcut_clause = gec "letcut_clause" - let clausearg = gec "clausearg" - let match_context_rule = gec "match_context_rule" - let match_hyps = gec "match_hyps" - let match_pattern = gec "match_pattern" - let match_context_list = gec_list "match_context_list" - let match_rule = gec "match_rule" - let match_list = gec_list "match_list" - let ne_identarg_list = gec_list "ne_identarg_list" - let ne_hyp_list = gec_list "ne_hyp_list" - let ne_idmetahyp_list = gec_list "ne_idmetahyp_list" - let ne_qualidarg_list = gec_list "ne_qualidarg_list" - let ne_pattern_list = gec_list "ne_pattern_list" - let clause_pattern = gec "clause_pattern" - let one_intropattern = gec "one_intropattern" - let intropattern = gec "intropattern" - let ne_intropattern = gec "ne_intropattern" - let simple_intropattern = gec "simple_intropattern" - let ne_unfold_occ_list = gec_list "ne_unfold_occ_list" - let rec_clause = gec "rec_clause" - let red_tactic = gec "red_tactic" - let red_flag = gec "red_flag" - let numarg = gec "numarg" - let pattern_occ = gec "pattern_occ" - let pattern_occ_hyp = gec "pattern_occ_hyp" - let pure_numarg = gec "pure_numarg" - let simple_binding = gec "simple_binding" - let simple_binding_list = gec_list "simple_binding_list" - let simple_tactic = gec "simple_tactic" - let tactic = gec "tactic" - let tactic_arg = gec "tactic_arg" - let tactic_atom0 = gec "tactic_atom0" - let tactic_atom = gec "tactic_atom" - let tactic_expr = gec "tactic_expr" - let tactic_expr_par = gec "tactic_expr_par" - let unfold_occ = gec "unfold_occ" - let with_binding_list = gec "with_binding_list" - let fixdecl = gec_list "fixdecl" - let cofixdecl = gec_list "cofixdecl" - let tacarg_list = gec_list "tacarg_list" + let gec = make_entry utactic inPureAstType + let gec_list = make_entry utactic inAstListType + let castedopenconstr = + make_gen_entry utactic rawwit_casted_open_constr "castedopenconstr" + let constr_with_bindings = + make_gen_entry utactic rawwit_constr_with_bindings "constr_with_bindings" + let constrarg = make_gen_entry utactic rawwit_constr_may_eval "constrarg" + let quantified_hypothesis = + make_gen_entry utactic rawwit_quant_hyp "quantified_hypothesis" + let int_or_var = make_gen_entry utactic rawwit_int_or_var "int_or_var" + let red_tactic = make_gen_entry utactic rawwit_red_expr "red_tactic" + let simple_tactic = make_entry utactic inTacticAtomAstType "simple_tactic" + let tactic_arg = Gram.Entry.create "tactic:tactic_arg" + let tactic = make_gen_entry utactic rawwit_tactic "tactic" let tactic_eoi = eoi_entry tactic end module Vernac_ = struct - let uvernac = snd (get_univ "vernac") - let gec s = - let e = Gram.Entry.create ("Vernac." ^ s) in - Hashtbl.add uvernac s (Ast e); e - - let gec_list s = - let e = Gram.Entry.create ("Vernac." ^ s) in - Hashtbl.add uvernac s (ListAst e); e - - let identarg = gec "identarg" - let ne_identarg_list = gec_list "ne_identarg_list" - let qualidarg = gec "qualidarg" - let commentarg = gec "commentarg" - let commentarg_list = gec_list "commentarg_list" - let ne_qualidarg_list = gec_list "ne_qualidarg_list" - let numarg = gec "numarg" - let numarg_list = gec_list "numarg_list" - let ne_numarg_list = gec_list "ne_numarg_list" - let stringarg = gec "stringarg" - let ne_stringarg_list = gec_list "ne_stringarg_list" - let constrarg = gec "constrarg" - let ne_constrarg_list = gec_list "ne_constrarg_list" - let tacarg = gec "tacarg" - let reduce = gec_list "reduce" - let sortarg = gec "sortarg" - let theorem_body = gec "theorem_body" - let thm_tok = gec "thm_tok" - - let gallina = gec "gallina" - let gallina_ext = gec "gallina_ext" - let command = gec "command" - let syntax = gec "syntax_command" - let vernac = gec "vernac" + let thm_token = make_entry uvernac inThmTokenAstType "thm_token" + let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" + let gec_vernac s = Gram.Entry.create ("vernac:" ^ s) + let gallina = gec_vernac "gallina" + let gallina_ext = gec_vernac "gallina_ext" + let command = gec_vernac "command" + let syntax = gec_vernac "syntax_command" + let vernac = gec_vernac "Vernac_.vernac" let vernac_eoi = eoi_entry vernac end -module Prim = - struct - let uprim = snd (get_univ "prim") - let gec s = - let e = Gram.Entry.create ("Prim." ^ s) in - Hashtbl.add uprim s (Ast e); e - let ast = gec "ast" - let ast_eoi = eoi_entry ast - let astact = gec "astact" - let astpat = gec "astpat" - let entry_type = gec "entry_type" - let grammar_entry = gec "grammar_entry" - let grammar_entry_eoi = eoi_entry grammar_entry - let ident = gec "ident" - let metaident = gec "metaident" - let number = gec "number" - let path = gec "path" - let string = gec "string" - let syntax_entry = gec "syntax_entry" - let syntax_entry_eoi = eoi_entry syntax_entry - let uprim = snd (get_univ "prim") - let var = gec "var" - end - - let main_entry = Gram.Entry.create "vernac" GEXTEND Gram main_entry: - [ [ a = Vernac_.vernac -> Some a | EOI -> None ] ] + [ [ a = Vernac_.vernac -> Some (loc,a) | EOI -> None ] ] ; END @@ -439,13 +460,13 @@ let define_quotation default s e = *) END); (GEXTEND Gram - GLOBAL: ast constr vernac tactic; + GLOBAL: ast constr command tactic; ast: [ [ "<:"; IDENT $s$; ":<"; c = e; ">>" -> c ] ]; (* Uncomment this to keep compatibility with old grammar syntax constr: [ [ "<:"; IDENT $s$; ":<"; c = e; ">>" -> c ] ]; - vernac: + command: [ [ "<:"; IDENT $s$; ":<"; c = e; ">>" -> c ] ]; tactic: [ [ "<:"; IDENT $s$; ":<"; c = e; ">>" -> c ] ]; @@ -454,31 +475,67 @@ let define_quotation default s e = let _ = define_quotation false "ast" ast in () -let constr_parser = ref Constr.constr -let tactic_parser = ref Tactic.tactic -let vernac_parser = ref Vernac_.vernac +let gecdyn s = + let e = Gram.Entry.create ("Dyn." ^ s) in + Hashtbl.add (snd uconstr) s (inDynamicAstType e); e + +let dynconstr = gecdyn "dynconstr" +let dyntactic = gecdyn "dyntactic" +let dynvernac = gecdyn "dynvernac" +let dynastlist = gecdyn "dynastlist" +let dynast = gecdyn "dynast" + +let globalizer = ref (fun x -> x) +let set_globalizer f = globalizer := f -let update_constr_parser p = constr_parser := p -let update_tactic_parser p = tactic_parser := p -let update_vernac_parser p = vernac_parser := p +GEXTEND Gram + dynconstr: [ [ a = Constr.constr -> !globalizer (PureAstNode a) ] ]; +(* + dyntactic: [ [ a = Tactic.tactic -> !globalizer (TacticAstNode a) ] ]; + dynvernac: [ [ a = Vernac_.vernac -> !globalizer(VernacAstNode a) ] ]; +*) + dynastlist: [ [ a = Prim.astlist -> AstListNode a ] ]; + dynast: [ [ a = ast -> PureAstNode a ] ]; +END (**********************************************************************) (* The following is to dynamically set the parser in Grammar actions *) (* and Syntax pattern, according to the universe of the rule defined *) -let default_action_parser_ref = ref ast +type parser_type = + | AstListParser + | AstParser + | ConstrParser + | TacticParser + | VernacParser -let get_default_action_parser () = !default_action_parser_ref +let default_action_parser_ref = ref dynast -let set_default_action_parser f = (default_action_parser_ref := f) +let get_default_action_parser () = !default_action_parser_ref -let set_default_action_parser_by_name = function - | "constr" -> set_default_action_parser !constr_parser - | "tactic" -> set_default_action_parser !tactic_parser - | "vernac" -> set_default_action_parser !vernac_parser - | _ -> set_default_action_parser ast +let entry_type_from_name = function + | "constr" -> GenAstType ConstrArgType + | "tactic" -> failwith "Not supported" + | "vernac" -> failwith "Not supported" + | s -> GenAstType ConstrArgType + +let entry_type_of_parser = function + | AstListParser -> Some AstListType + | _ -> None + +let parser_type_from_name = function + | "constr" -> ConstrParser + | "tactic" -> TacticParser + | "vernac" -> VernacParser + | s -> AstParser + +let set_default_action_parser = function + | AstListParser -> default_action_parser_ref := dynastlist + | AstParser -> default_action_parser_ref := dynast + | ConstrParser -> default_action_parser_ref := dynconstr + | TacticParser -> default_action_parser_ref := dyntactic + | VernacParser -> default_action_parser_ref := dynvernac let default_action_parser = Gram.Entry.of_parser "default_action_parser" (fun strm -> Gram.Entry.parse_token (get_default_action_parser ()) strm) - diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index b9317f581..b3c3db414 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -8,15 +8,24 @@ (*i $Id$ i*) +open Names +open Tacexpr +open Ast +open Genarg +open Tacexpr +open Vernacexpr + (* The lexer and parser of Coq. *) val lexer : Token.lexer module Gram : Grammar.S -type typed_entry = - | Ast of Coqast.t Gram.Entry.e - | ListAst of Coqast.t list Gram.Entry.e +type grammar_object +type typed_entry + +val type_of_typed_entry : typed_entry -> entry_type +val object_of_typed_entry : typed_entry -> grammar_object Gram.Entry.e val grammar_extend : typed_entry -> Gramext.position option -> @@ -32,48 +41,87 @@ val parse_string : 'a Gram.Entry.e -> string -> 'a val eoi_entry : 'a Gram.Entry.e -> 'a Gram.Entry.e val map_entry : ('a -> 'b) -> 'a Gram.Entry.e -> 'b Gram.Entry.e +(* val slam_ast : Coqast.loc -> Coqast.t -> Coqast.t -> Coqast.t val abstract_binders_ast : Coqast.loc -> string -> Coqast.t -> Coqast.t -> Coqast.t +*) (* Entry types *) -type entry_type = ETast | ETastl - -val entry_type : Coqast.t -> entry_type -val type_of_entry : typed_entry -> entry_type - (* Table of Coq's grammar entries *) type gram_universe +val create_univ_if_new : string -> string * gram_universe val get_univ : string -> string * gram_universe -val get_entry : string * gram_universe -> string -> typed_entry - -val create_entry : string * gram_universe -> string -> entry_type -> - typed_entry -val force_entry_type : string * gram_universe -> string -> - entry_type -> typed_entry +val get_entry : string * gram_universe -> string -> typed_entry + +val entry_type : string * gram_universe -> string -> entry_type option +val get_entry_type : string * string -> entry_type +val create_entry_if_new : + string * gram_universe -> string -> entry_type -> unit +val create_entry : + string * gram_universe -> string -> entry_type -> typed_entry +val force_entry_type : + string * gram_universe -> string -> entry_type -> typed_entry + +val create_constr_entry : + string * gram_universe -> string -> Coqast.t Gram.Entry.e +val create_generic_entry : string -> ('a, constr_ast,raw_tactic_expr) abstract_argument_type -> 'a Gram.Entry.e +val get_generic_entry : string -> grammar_object Gram.Entry.e +val get_generic_entry_type : string * gram_universe -> string -> Genarg.argument_type + +type parser_type = + | AstListParser + | AstParser + | ConstrParser + | TacticParser + | VernacParser + +val entry_type_from_name : string -> entry_type +val entry_type_of_parser : parser_type -> entry_type option +val parser_type_from_name : string -> parser_type (* Quotations *) - val define_quotation : bool -> string -> (Coqast.t Gram.Entry.e) -> unit -val update_constr_parser : Coqast.t Gram.Entry.e -> unit -val update_tactic_parser : Coqast.t Gram.Entry.e -> unit -val update_vernac_parser : Coqast.t Gram.Entry.e -> unit +val set_globalizer : (typed_ast -> typed_ast) -> unit (* The default parser for actions in grammar rules *) -val default_action_parser : Coqast.t Gram.Entry.e -val set_default_action_parser : Coqast.t Gram.Entry.e -> unit -val set_default_action_parser_by_name : string -> unit +val default_action_parser : typed_ast Gram.Entry.e +val set_default_action_parser : parser_type -> unit (* The main entry: reads an optional vernac command *) -val main_entry : Coqast.t option Gram.Entry.e +val main_entry : (Coqast.loc * vernac_expr) option Gram.Entry.e (* Initial state of the grammar *) +module Prim : + sig + open Util + open Names + open Nametab + val preident : string Gram.Entry.e + val ident : identifier Gram.Entry.e + val rawident : identifier located Gram.Entry.e + val natural : int Gram.Entry.e + val integer : int Gram.Entry.e + val string : string Gram.Entry.e + val qualid : qualid located Gram.Entry.e + val reference : reference_expr Gram.Entry.e + val dirpath : dir_path Gram.Entry.e + val astpat: typed_ast Gram.Entry.e + val ast : Coqast.t Gram.Entry.e + val astlist : Coqast.t list Gram.Entry.e + val ast_eoi : Coqast.t Gram.Entry.e + val astact : Coqast.t Gram.Entry.e + val metaident : Coqast.t Gram.Entry.e + val numarg : Coqast.t Gram.Entry.e + val var : Coqast.t Gram.Entry.e + end + module Constr : sig val constr : Coqast.t Gram.Entry.e @@ -89,127 +137,47 @@ module Constr : val constr9 : Coqast.t Gram.Entry.e val constr91 : Coqast.t Gram.Entry.e val constr10 : Coqast.t Gram.Entry.e - val constr_eoi : Coqast.t Gram.Entry.e + val constr_eoi : constr_ast Gram.Entry.e val lconstr : Coqast.t Gram.Entry.e val ident : Coqast.t Gram.Entry.e - val qualid : Coqast.t list Gram.Entry.e + val qualid : Coqast.t Gram.Entry.e val global : Coqast.t Gram.Entry.e val ne_ident_comma_list : Coqast.t list Gram.Entry.e val ne_constr_list : Coqast.t list Gram.Entry.e val sort : Coqast.t Gram.Entry.e val pattern : Coqast.t Gram.Entry.e + val constr_pattern : Coqast.t Gram.Entry.e val ne_binders_list : Coqast.t list Gram.Entry.e end module Tactic : sig - val autoargs : Coqast.t list Gram.Entry.e - val binding_list : Coqast.t Gram.Entry.e - val castedopenconstrarg : Coqast.t Gram.Entry.e - val castedconstrarg : Coqast.t Gram.Entry.e - val clausearg : Coqast.t Gram.Entry.e - val cofixdecl : Coqast.t list Gram.Entry.e - val com_binding_list : Coqast.t list Gram.Entry.e - val constrarg : Coqast.t Gram.Entry.e - val constrarg_binding_list : Coqast.t list Gram.Entry.e - val constrarg_list : Coqast.t list Gram.Entry.e - val ne_constrarg_list : Coqast.t list Gram.Entry.e - val fixdecl : Coqast.t list Gram.Entry.e - val ident_or_numarg : Coqast.t Gram.Entry.e - val ident_or_constrarg : Coqast.t Gram.Entry.e - val identarg : Coqast.t Gram.Entry.e - val hypident : Coqast.t Gram.Entry.e - val idmeta_arg : Coqast.t Gram.Entry.e - val idmetahyp : Coqast.t Gram.Entry.e - val qualidarg : Coqast.t Gram.Entry.e - val input_fun : Coqast.t Gram.Entry.e - val intropattern : Coqast.t Gram.Entry.e - val lconstrarg : Coqast.t Gram.Entry.e - val lconstrarg_binding_list : Coqast.t list Gram.Entry.e - val let_clause : Coqast.t Gram.Entry.e - val letcut_clause : Coqast.t Gram.Entry.e - val match_context_rule : Coqast.t Gram.Entry.e - val match_hyps : Coqast.t Gram.Entry.e - val match_pattern : Coqast.t Gram.Entry.e - val match_context_list : Coqast.t list Gram.Entry.e - val match_rule : Coqast.t Gram.Entry.e - val match_list : Coqast.t list Gram.Entry.e - val ne_identarg_list : Coqast.t list Gram.Entry.e - val ne_hyp_list : Coqast.t list Gram.Entry.e - val ne_idmetahyp_list : Coqast.t list Gram.Entry.e - val ne_qualidarg_list : Coqast.t list Gram.Entry.e - val ne_intropattern : Coqast.t Gram.Entry.e - val ne_pattern_list : Coqast.t list Gram.Entry.e - val clause_pattern : Coqast.t Gram.Entry.e - val ne_unfold_occ_list : Coqast.t list Gram.Entry.e - val numarg : Coqast.t Gram.Entry.e - val numarg_binding_list : Coqast.t list Gram.Entry.e - val one_intropattern : Coqast.t Gram.Entry.e - val pattern_occ : Coqast.t Gram.Entry.e - val pattern_occ_hyp : Coqast.t Gram.Entry.e - val pure_numarg : Coqast.t Gram.Entry.e - val rec_clause : Coqast.t Gram.Entry.e - val red_flag : Coqast.t Gram.Entry.e - val red_tactic : Coqast.t Gram.Entry.e - val simple_binding : Coqast.t Gram.Entry.e - val simple_binding_list : Coqast.t list Gram.Entry.e - val simple_intropattern : Coqast.t Gram.Entry.e - val simple_tactic : Coqast.t Gram.Entry.e - val tactic : Coqast.t Gram.Entry.e - val tactic_arg : Coqast.t Gram.Entry.e - val tactic_atom0 : Coqast.t Gram.Entry.e - val tactic_atom : Coqast.t Gram.Entry.e - val tactic_eoi : Coqast.t Gram.Entry.e - val tactic_expr : Coqast.t Gram.Entry.e - val tactic_expr_par : Coqast.t Gram.Entry.e - val unfold_occ : Coqast.t Gram.Entry.e - val with_binding_list : Coqast.t Gram.Entry.e + open Rawterm + val castedopenconstr : constr_ast Gram.Entry.e + val constr_with_bindings : constr_ast with_bindings Gram.Entry.e + val constrarg : constr_ast may_eval Gram.Entry.e + val quantified_hypothesis : quantified_hypothesis Gram.Entry.e + val int_or_var : int or_var Gram.Entry.e + val red_tactic : raw_red_expr Gram.Entry.e + val simple_tactic : raw_atomic_tactic_expr Gram.Entry.e + val tactic_arg : raw_tactic_arg Gram.Entry.e + val tactic : raw_tactic_expr Gram.Entry.e + val tactic_eoi : raw_tactic_expr Gram.Entry.e end module Vernac_ : sig - val identarg : Coqast.t Gram.Entry.e - val ne_identarg_list : Coqast.t list Gram.Entry.e - val qualidarg : Coqast.t Gram.Entry.e - val commentarg : Coqast.t Gram.Entry.e - val commentarg_list : Coqast.t list Gram.Entry.e - val ne_qualidarg_list : Coqast.t list Gram.Entry.e - val numarg : Coqast.t Gram.Entry.e - val numarg_list : Coqast.t list Gram.Entry.e - val ne_numarg_list : Coqast.t list Gram.Entry.e - val stringarg : Coqast.t Gram.Entry.e - val ne_stringarg_list : Coqast.t list Gram.Entry.e - val constrarg : Coqast.t Gram.Entry.e - val ne_constrarg_list : Coqast.t list Gram.Entry.e - val tacarg : Coqast.t Gram.Entry.e + open Util + open Nametab + val thm_token : theorem_kind Gram.Entry.e + val class_rawexpr : class_rawexpr Gram.Entry.e + val gallina : vernac_expr Gram.Entry.e + val gallina_ext : vernac_expr Gram.Entry.e + val command : vernac_expr Gram.Entry.e + val syntax : vernac_expr Gram.Entry.e + val vernac : vernac_expr Gram.Entry.e + val vernac_eoi : vernac_expr Gram.Entry.e +(* val reduce : Coqast.t list Gram.Entry.e - val sortarg : Coqast.t Gram.Entry.e - val theorem_body : Coqast.t Gram.Entry.e - val thm_tok : Coqast.t Gram.Entry.e - - val gallina : Coqast.t Gram.Entry.e - val gallina_ext : Coqast.t Gram.Entry.e - val command : Coqast.t Gram.Entry.e - val syntax : Coqast.t Gram.Entry.e - val vernac : Coqast.t Gram.Entry.e - val vernac_eoi : Coqast.t Gram.Entry.e - end - -module Prim : - sig - val ast : Coqast.t Gram.Entry.e - val ast_eoi : Coqast.t Gram.Entry.e - val astact : Coqast.t Gram.Entry.e - val astpat: Coqast.t Gram.Entry.e - val entry_type : Coqast.t Gram.Entry.e - val grammar_entry : Coqast.t Gram.Entry.e - val grammar_entry_eoi : Coqast.t Gram.Entry.e - val ident : Coqast.t Gram.Entry.e - val metaident : Coqast.t Gram.Entry.e - val number : Coqast.t Gram.Entry.e - val path : Coqast.t Gram.Entry.e - val string : Coqast.t Gram.Entry.e - val syntax_entry : Coqast.t Gram.Entry.e - val syntax_entry_eoi : Coqast.t Gram.Entry.e - val var : Coqast.t Gram.Entry.e +*) end diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml new file mode 100644 index 000000000..d73c69ee6 --- /dev/null +++ b/parsing/ppconstr.ml @@ -0,0 +1,148 @@ +(****************************************************************************) +(* *) +(* The Coq Proof Assistant *) +(* *) +(* Projet Coq *) +(* *) +(* INRIA LRI-CNRS ENS-CNRS *) +(* Rocquencourt Orsay Lyon *) +(* *) +(****************************************************************************) + +(* $:Id$ *) + +open Ast +open Pp +open Nametab +open Names +open Nameops +open Coqast +open Extend +open Util + +let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; + +let pr_global ref = pr_global_env (Global.env()) ref + +let global_const_name sp = + try pr_global (ConstRef sp) + with Not_found -> (* May happen in debug *) + (str ("CONST("^(string_of_path sp)^")")) + +let global_var_name id = + try pr_global (VarRef id) + with Not_found -> (* May happen in debug *) + (str ("SECVAR("^(string_of_id id)^")")) + +let global_ind_name (sp,tyi) = + try pr_global (IndRef (sp,tyi)) + with Not_found -> (* May happen in debug *) + (str ("IND("^(string_of_path sp)^","^(string_of_int tyi)^")")) + +let global_constr_name ((sp,tyi),i) = + try pr_global (ConstructRef ((sp,tyi),i)) + with Not_found -> (* May happen in debug *) + (str ("CONSTRUCT("^(string_of_path sp)^","^(string_of_int tyi) + ^","^(string_of_int i)^")")) + +let globpr gt = match gt with + | Nvar(_,s) -> (pr_id s) + | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) + | Node(_,"CONST",[Path(_,sl)]) -> + global_const_name (section_path sl) + | Node(_,"SECVAR",[Nvar(_,s)]) -> + global_var_name s + | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> + global_ind_name (section_path sl, tyi) + | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> + global_constr_name ((section_path sl, tyi), i) + | Dynamic(_,d) -> + if (Dyn.tag d) = "constr" then (str"") + else dfltpr gt + | gt -> dfltpr gt + +let wrap_exception = function + Anomaly (s1,s2) -> + warning ("Anomaly ("^s1^")"); pp s2; + str"" + | Failure _ | UserError _ | Not_found -> + str"" + | s -> raise s + +let constr_syntax_universe = "constr" +(* This is starting precedence for printing constructions or tactics *) +(* Level 9 means no parentheses except for applicative terms (at level 10) *) +let constr_initial_prec = Some ((constr_syntax_universe,(9,0,0)),Extend.L) + +let gentermpr_fail gt = + Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt + +let gentermpr gt = + try gentermpr_fail gt + with s -> wrap_exception s + +(* [at_top] means ids of env must be avoided in bound variables *) +let gentermpr_core at_top env t = + gentermpr (Termast.ast_of_constr at_top env t) + +let pr_constr = gentermpr + +let pr_pattern = gentermpr + +let pr_qualid qid = str (Nametab.string_of_qualid qid) + +open Rawterm + +let pr_arg pr x = spc () ++ pr x + +let pr_red_flag pr r = + (if r.rBeta then pr_arg str "Beta" else mt ()) ++ + (if r.rIota then pr_arg str "Iota" else mt ()) ++ + (if r.rZeta then pr_arg str "Zeta" else mt ()) ++ + (if r.rConst = [] then + if r.rDelta then pr_arg str "Delta" + else mt () + else + pr_arg str "Delta" ++ (if r.rDelta then str "-" else mt ()) ++ + hov 0 (str "[" ++ prlist_with_sep spc pr r.rConst ++ str "]")) + +open Genarg + +let pr_metanum pr = function + | AN (_,x) -> pr x + | MetaNum (_,n) -> str "?" ++ int n + +let pr_red_expr (pr_constr,pr_ref) = function + | Red false -> str "Red" + | Hnf -> str "Hnf" + | Simpl -> str "Simpl" + | Cbv f -> + if f = {rBeta=true;rIota=true;rZeta=true;rDelta=true;rConst=[]} then + str "Compute" + else + hov 1 (str "Cbv" ++ spc () ++ pr_red_flag pr_ref f) + | Lazy f -> + hov 1 (str "Lazy" ++ spc () ++ pr_red_flag pr_ref f) + | Unfold l -> + hov 1 (str "Unfold" ++ + prlist (fun (nl,qid) -> + prlist (pr_arg int) nl ++ spc () ++ pr_ref qid) l) + | Fold l -> hov 1 (str "Fold" ++ prlist (pr_arg pr_constr) l) + | Pattern l -> + hov 1 (str "Pattern" ++ + prlist(fun (nl,c) -> prlist (pr_arg int) nl ++ (pr_arg pr_constr) c) l) + | (Red true | Cbv _ | Lazy _) -> error "Shouldn't be accessible from user" + | ExtraRedExpr (s,l) -> + hov 1 (str s ++ prlist (pr_arg pr_constr) l) + +let rec pr_may_eval pr = function + | ConstrEval (r,c) -> + hov 0 + (str "Eval" ++ brk (1,1) ++ pr_red_expr (pr,pr_metanum pr_qualid) r ++ + spc () ++ str "in" ++ brk (1,1) ++ pr c) + | ConstrContext ((_,id),c) -> + hov 0 + (str "Inst " ++ brk (1,1) ++ pr_id id ++ spc () ++ + str "[" ++ pr c ++ str "]") + | ConstrTypeOf c -> hov 0 (str "Check " ++ brk (1,1) ++ pr c) + | ConstrTerm c -> pr c diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli new file mode 100644 index 000000000..d9b8b6ea5 --- /dev/null +++ b/parsing/ppconstr.mli @@ -0,0 +1,33 @@ +(****************************************************************************) +(* *) +(* The Coq Proof Assistant *) +(* *) +(* Projet Coq *) +(* *) +(* INRIA LRI-CNRS ENS-CNRS *) +(* Rocquencourt Orsay Lyon *) +(* *) +(****************************************************************************) + +(* $Id$ *) + +open Pp +open Environ +open Term +open Nametab +open Pcoq +open Rawterm +open Extend + +val pr_global : global_reference -> std_ppcmds +val gentermpr : Coqast.t -> std_ppcmds +val gentermpr_core : bool -> env -> constr -> std_ppcmds + +val pr_qualid : qualid -> std_ppcmds +val pr_red_expr : + ('a -> std_ppcmds) * ('b -> std_ppcmds) -> + ('a,'b) red_expr_gen -> std_ppcmds + +val pr_pattern : Tacexpr.pattern_ast -> std_ppcmds +val pr_constr : Genarg.constr_ast -> std_ppcmds +val pr_may_eval : ('a -> std_ppcmds) -> 'a may_eval -> std_ppcmds diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml new file mode 100644 index 000000000..9dc447f7d --- /dev/null +++ b/parsing/pptactic.ml @@ -0,0 +1,607 @@ +(****************************************************************************) +(* *) +(* The Coq Proof Assistant *) +(* *) +(* Projet Coq *) +(* *) +(* INRIA LRI-CNRS ENS-CNRS *) +(* Rocquencourt Orsay Lyon *) +(* *) +(****************************************************************************) + +(* $Id$ *) + +open Pp +open Names +open Nameops +open Util +open Extend +open Ppconstr +open Tacexpr +open Rawterm +open Genarg + + (* Extensions *) +let prtac_tab = Hashtbl.create 17 + +let declare_extra_tactic_pprule s f g = + Hashtbl.add prtac_tab s (f,g) + +let genarg_pprule = ref Stringmap.empty + +let declare_extra_genarg_pprule (rawwit, f) (wit, g) = + let s = match unquote wit with + | ExtraArgType s -> s + | _ -> error + "Can declare a pretty-printing rule only for extra argument types" + in + let f x = f (out_gen rawwit x) in + let g x = g (out_gen wit x) in + genarg_pprule := Stringmap.add s (f,g) !genarg_pprule + +(* [pr_rawtac] is here to cheat with ML typing system, + gen_tactic_expr is polymorphic but with some occurrences of its + instance raw_tactic_expr in it; then pr_tactic should be + polymorphic but with some calls to instance of itself, what ML does + not accept; pr_rawtac0 denotes this instance of pr_tactic on + raw_tactic_expr *) + +let pr_rawtac = + ref (fun _ -> failwith "Printer for raw tactic expr is not defined" + : raw_tactic_expr -> std_ppcmds) +let pr_rawtac0 = + ref (fun _ -> failwith "Printer for raw tactic expr is not defined" + : raw_tactic_expr -> std_ppcmds) + +let pr_arg pr x = spc () ++ pr x + +let pr_name = function + | Anonymous -> mt () + | Name id -> spc () ++ pr_id id + +let pr_metanum pr = function + | AN (_,x) -> pr x + | MetaNum (_,n) -> str "?" ++ int n + +let pr_or_var pr = function + | ArgArg x -> pr x + | ArgVar (_,s) -> pr_id s + +let pr_opt pr = function + | None -> mt () + | Some x -> spc () ++ pr x + +let pr_or_meta pr = function + | AI x -> pr x + | _ -> failwith "pr_hyp_location: unexpected quotation meta-variable" + +let pr_casted_open_constr = pr_constr + +let pr_quantified_hypothesis = function + | AnonHyp n -> int n + | NamedHyp id -> pr_id id + +let pr_quantified_hypothesis_arg h = spc () ++ pr_quantified_hypothesis h + +let pr_binding prc = function + | NamedHyp id, c -> hov 1 (pr_id id ++ str ":=" ++ cut () ++ prc c) + | AnonHyp n, c -> hov 1 (int n ++ str ":=" ++ cut () ++ prc c) + +let pr_bindings prc = function + | ImplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + prlist_with_sep spc prc l + | ExplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + prlist_with_sep spc (pr_binding prc) l + | NoBindings -> mt () + +let pr_with_bindings prc (c,bl) = prc c ++ hv 0 (pr_bindings prc bl) + +let rec pr_intro_pattern = function + | IntroOrAndPattern pll -> + str "[" ++ + hv 0 (prlist_with_sep pr_coma (prlist_with_sep spc pr_intro_pattern) pll) + ++ str "]" +(* + | IntroAndPattern pl -> + str "(" ++ hov 0 (prlist_with_sep pr_coma pr_intro_pattern pl) ++ str ")" +*) + | IntroWildcard -> str "_" + | IntroIdentifier id -> pr_id id + +let pr_hyp_location pr_id = function + | InHyp id -> spc () ++ pr_id id + | InHypType id -> spc () ++ str "(Type of " ++ pr_id id ++ str ")" + +let pr_clause pr_id = function + | [] -> mt () + | l -> spc () ++ hov 0 (str "in" ++ prlist (pr_hyp_location pr_id) l) + +let pr_clause_pattern pr_id = function (* To check *) + | (None, []) -> mt () + | (glopt,l) -> + str "in" ++ + prlist + (fun (id,nl) -> spc () ++ prlist_with_sep spc int nl + ++ spc () ++ pr_id id) l ++ + pr_opt (prlist_with_sep spc int) glopt + +let pr_induction_arg prc = function + | ElimOnConstr c -> prc c + | ElimOnIdent (_,id) -> pr_id id + | ElimOnAnonHyp n -> int n + +let pr_match_pattern = function + | Term a -> pr_pattern a + | Subterm (None,a) -> str "[" ++ pr_pattern a ++ str "]" + | Subterm (Some id,a) -> pr_id id ++ str "[" ++ pr_pattern a ++ str "]" + +let pr_match_hyps = function + | NoHypId mp -> str "_:" ++ pr_match_pattern mp + | Hyp ((_,id),mp) -> pr_id id ++ str ":" ++ pr_match_pattern mp + +let pr_match_rule m pr = function + | Pat ([],mp,t) when m -> + str "[" ++ pr_match_pattern mp ++ str "]" + ++ spc () ++ str "->" ++ brk (1,2) ++ pr t + | Pat (rl,mp,t) -> + str "[" ++ prlist_with_sep pr_semicolon pr_match_hyps rl ++ spc () ++ + str "|-" ++ spc () ++ pr_match_pattern mp ++ spc () ++ str "]" ++ + spc () ++ str "->" ++ brk (1,2) ++ pr t + | All t -> str "_" ++ spc () ++ str "->" ++ brk (1,2) ++ pr t + +let pr_funvar = function + | None -> spc () ++ str "()" + | Some id -> spc () ++ pr_id id + +let pr_let_clause k pr = function + | ((_,id),None,t) -> hv 0(str k ++ pr_id id ++ str " =" ++ brk (1,1) ++ pr t) + | ((_,id),Some c,t) -> str "TODO(LETCLAUSE)" + +let pr_let_clauses pr = function + | hd::tl -> + hv 0 + (pr_let_clause "Let " pr hd ++ spc () ++ + prlist_with_sep spc (pr_let_clause "And " pr) tl) + | [] -> anomaly "LetIn must declare at least one binding" + +let pr_rec_clause pr ((_,id),(l,t)) = + pr_id id ++ prlist pr_funvar l ++ str "->" ++ spc () ++ pr t + +let pr_rec_clauses pr l = + prlist_with_sep (fun () -> fnl () ++ str "And ") (pr_rec_clause pr) l + +let pr_hintbases = function + | None -> spc () ++ str "with *" + | Some [] -> mt () + | Some l -> + spc () ++ str "with" ++ hv 0 (prlist (fun s -> spc () ++ str s) l) + +let pr_autoarg_adding = function + | [] -> mt () + | l -> + spc () ++ str "Adding [" ++ + hv 0 (prlist_with_sep spc (fun (_,x) -> pr_qualid x) l) ++ str "]" + +let pr_autoarg_destructing = function + | true -> spc () ++ str "Destructing" + | false -> mt () + +let pr_autoarg_usingTDB = function + | true -> spc () ++ str "Using TDB" + | false -> mt () + +let rec pr_rawgen prtac x = + match Genarg.genarg_tag x with + | BoolArgType -> pr_arg str (if out_gen rawwit_bool x then "true" else "false") + | IntArgType -> pr_arg int (out_gen rawwit_int x) + | IntOrVarArgType -> pr_arg (pr_or_var pr_int) (out_gen rawwit_int_or_var x) + | StringArgType -> spc () ++ str "\"" ++ str (out_gen rawwit_string x) ++ str "\"" + | PreIdentArgType -> pr_arg str (out_gen rawwit_pre_ident x) + | IdentArgType -> pr_arg pr_id (out_gen rawwit_ident x) + | QualidArgType -> pr_arg pr_qualid (snd (out_gen rawwit_qualid x)) + | ConstrArgType -> pr_arg pr_constr (out_gen rawwit_constr x) + | ConstrMayEvalArgType -> + pr_arg (pr_may_eval pr_constr) (out_gen rawwit_constr_may_eval x) + | QuantHypArgType -> + pr_arg pr_quantified_hypothesis (out_gen rawwit_quant_hyp x) + | RedExprArgType -> + pr_arg (pr_red_expr (pr_constr,pr_metanum pr_qualid)) (out_gen rawwit_red_expr x) + | TacticArgType -> pr_arg prtac (out_gen rawwit_tactic x) + | CastedOpenConstrArgType -> + pr_arg pr_casted_open_constr (out_gen rawwit_casted_open_constr x) + | ConstrWithBindingsArgType -> + pr_arg (pr_with_bindings pr_constr) + (out_gen rawwit_constr_with_bindings x) + | List0ArgType _ -> + hov 0 (fold_list0 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) + | List1ArgType _ -> + hov 0 (fold_list1 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) + | OptArgType _ -> hov 0 (fold_opt (pr_rawgen prtac) (mt()) x) + | PairArgType _ -> + hov 0 + (fold_pair + (fun a b -> pr_rawgen prtac a ++ spc () ++ pr_rawgen prtac b) + x) + | ExtraArgType s -> + try fst (Stringmap.find s !genarg_pprule) x + with Not_found -> str " [no printer for " ++ str s ++ str "] " + +let rec pr_raw_tacarg_using_rule pr_gen = function + | Egrammar.TacTerm s :: l, al -> str s ++ pr_raw_tacarg_using_rule pr_gen (l,al) + | Egrammar.TacNonTerm _ :: l, a :: al -> pr_gen a ++ pr_raw_tacarg_using_rule pr_gen (l,al) + | [], [] -> mt () + | _ -> failwith "Inconsistent arguments of extended tactic" + +let pr_raw_extend prt s l = + try + let (s,pl) = fst (Hashtbl.find prtac_tab s) l in + str s ++ pr_raw_tacarg_using_rule (pr_rawgen prt) (pl,l) + with Not_found -> + str "TODO(" ++ str s ++ prlist (pr_rawgen prt) l ++ str ")" + +open Closure + +let pr_evaluable_reference = function + | EvalVarRef id -> pr_id id + | EvalConstRef sp -> pr_global (Nametab.ConstRef sp) + +let pr_inductive ind = pr_global (Nametab.IndRef ind) + +let rec pr_generic prtac x = + match Genarg.genarg_tag x with + | BoolArgType -> pr_arg str (if out_gen wit_bool x then "true" else "false") + | IntArgType -> pr_arg int (out_gen wit_int x) + | IntOrVarArgType -> pr_arg (pr_or_var pr_int) (out_gen wit_int_or_var x) + | StringArgType -> spc () ++ str "\"" ++ str (out_gen wit_string x) ++ str "\"" + | PreIdentArgType -> pr_arg str (out_gen wit_pre_ident x) + | IdentArgType -> pr_arg pr_id (out_gen wit_ident x) + | QualidArgType -> pr_arg pr_global (out_gen wit_qualid x) + | ConstrArgType -> pr_arg Printer.prterm (out_gen wit_constr x) + | ConstrMayEvalArgType -> + pr_arg Printer.prterm (out_gen wit_constr_may_eval x) + | QuantHypArgType -> + pr_arg pr_quantified_hypothesis (out_gen wit_quant_hyp x) + | RedExprArgType -> + pr_arg (pr_red_expr (Printer.prterm,pr_evaluable_reference)) (out_gen wit_red_expr x) + | TacticArgType -> pr_arg prtac (out_gen wit_tactic x) + | CastedOpenConstrArgType -> + pr_arg Printer.prterm (snd (out_gen wit_casted_open_constr x)) + | ConstrWithBindingsArgType -> + pr_arg (pr_with_bindings Printer.prterm) + (out_gen wit_constr_with_bindings x) + | List0ArgType _ -> + hov 0 (fold_list0 (fun x a -> pr_generic prtac x ++ a) x (mt())) + | List1ArgType _ -> + hov 0 (fold_list1 (fun x a -> pr_generic prtac x ++ a) x (mt())) + | OptArgType _ -> hov 0 (fold_opt (pr_generic prtac) (mt()) x) + | PairArgType _ -> + hov 0 + (fold_pair + (fun a b -> pr_generic prtac a ++ spc () ++ pr_generic prtac b) + x) + | ExtraArgType s -> + try snd (Stringmap.find s !genarg_pprule) x + with Not_found -> str "[no printer for " ++ str s ++ str "]" + +let rec pr_tacarg_using_rule prt = function + | Egrammar.TacTerm s :: l, al -> str s ++ pr_tacarg_using_rule prt (l,al) + | Egrammar.TacNonTerm _ :: l, a :: al -> pr_generic prt a ++ pr_tacarg_using_rule prt (l,al) + | [], [] -> mt () + | _ -> failwith "Inconsistent arguments of extended tactic" + +let pr_extend prt s l = + try + let (s,pl) = snd (Hashtbl.find prtac_tab s) l in + str s ++ pr_tacarg_using_rule prt (pl,l) + with Not_found -> + str s ++ prlist (pr_generic prt) l + +let make_pr_tac (pr_constr,pr_cst,pr_ind,pr_ident,pr_extend) = + +let pr_bindings = pr_bindings pr_constr in +let pr_with_bindings = pr_with_bindings pr_constr in +let pr_constrarg c = spc () ++ pr_constr c in +let pr_intarg n = spc () ++ int n in + + (* Printing tactics as arguments *) +let rec pr_atom0 = function + | TacIntroPattern [] -> str "Intros" + | TacIntroMove (None,None) -> str "Intro" + | TacAssumption -> str "Assumption" + | TacAnyConstructor None -> str "Constructor" + | TacTrivial (Some []) -> str "Trivial" + | TacAuto (None,Some []) -> str "Auto" + | TacAutoTDB None -> str "AutoTDB" + | TacDestructConcl -> str "DConcl" + | TacReflexivity -> str "Reflexivity" + | TacSymmetry -> str "Symmetry" + | t -> str "(" ++ pr_atom1 t ++ str ")" + + (* Main tactic printer *) +and pr_atom1 = function + | TacExtend (s,l) -> pr_extend !pr_rawtac s l + | TacAlias (s,l,_) -> pr_extend !pr_rawtac s (List.map snd l) + + (* Basic tactics *) + | TacIntroPattern [] as t -> pr_atom0 t + | TacIntroPattern (_::_ as p) -> + hov 1 (str "Intro" ++ spc () ++ prlist_with_sep spc pr_intro_pattern p) + | TacIntrosUntil h -> + hv 1 (str "Intros until" ++ pr_arg pr_quantified_hypothesis h) + | TacIntroMove (None,None) as t -> pr_atom0 t + | TacIntroMove (Some id1,None) -> str "Intro " ++ pr_id id1 + | TacIntroMove (ido1,Some (_,id2)) -> + hov 1 + (str "Intro" ++ pr_opt pr_id ido1 ++ spc () ++ str "after " ++ pr_id id2) + | TacAssumption as t -> pr_atom0 t + | TacExact c -> hov 1 (str "Exact" ++ pr_arg pr_constr c) + | TacApply cb -> hov 1 (str "Apply" ++ spc () ++ pr_with_bindings cb) + | TacElim (cb,None) -> hov 1 (str "Elim" ++ spc () ++ pr_with_bindings cb) + | TacElim (cb,Some cb') -> + hov 1 (str "Elim" ++ pr_arg pr_with_bindings cb ++ + spc () ++ str "using" ++ pr_arg pr_with_bindings cb') + | TacElimType c -> hov 1 (str "ElimType" ++ pr_arg pr_constr c) + | TacCase cb -> hov 1 (str "Case" ++ spc () ++ pr_with_bindings cb) + | TacCaseType c -> hov 1 (str "CaseType" ++ pr_arg pr_constr c) + | TacFix (ido,n) -> hov 1 (str "Fix" ++ pr_opt pr_id ido ++ pr_intarg n) + | TacMutualFix (id,n,l) -> + hov 1 (str "Cofix" ++ spc () ++ pr_id id ++ pr_intarg n ++ spc () ++ + hov 0 (str "with" ++ brk (1,1) ++ + prlist_with_sep spc + (fun (id,n,c) -> + spc () ++ pr_id id ++ pr_intarg n ++ pr_arg pr_constr c) + l)) + | TacCofix ido -> hov 1 (str "Cofix" ++ pr_opt pr_id ido) + | TacMutualCofix (id,l) -> + hov 1 (str "Cofix" ++ spc () ++ pr_id id ++ spc () ++ + hov 0 (str "with" ++ brk (1,1) ++ + prlist (fun (id,c) -> spc () ++ pr_id id ++ pr_arg pr_constr c) + l)) + | TacCut c -> hov 1 (str "Cut" ++ pr_arg pr_constr c) + | TacTrueCut (None,c) -> + hov 1 (str "Assert" ++ pr_arg pr_constr c) + | TacTrueCut (Some id,c) -> + hov 1 (str "Assert" ++ spc () ++ pr_id id ++ str ":" ++ pr_constr c) + | TacForward (false,na,c) -> + hov 1 (str "Assert" ++ pr_name na ++ str ":=" ++ pr_constr c) + | TacForward (true,na,c) -> + hov 1 (str "Pose" ++ pr_name na ++ str ":=" ++ pr_constr c) + | TacGeneralize l -> + hov 0 (str "Generalize" ++ brk (1,1) ++ prlist_with_sep spc pr_constr l) + | TacGeneralizeDep c -> + hov 1 (str "Generalize" ++ spc () ++ str "Dependent" ++ spc () ++ + pr_constr c) + | TacLetTac (id,c,cl) -> + hov 1 (str "LetTac" ++ spc () ++ pr_id id ++ str ":=" ++ + pr_constr c ++ pr_clause_pattern pr_ident cl) + | TacInstantiate (n,c) -> + hov 1 (str "Instantiate" ++ pr_arg int n ++ pr_arg pr_constr c) + + (* Derived basic tactics *) + | TacOldInduction h -> + hov 1 (str "Induction" ++ pr_arg pr_quantified_hypothesis h) + | TacNewInduction h -> + hov 1 (str "NewInduction" ++ spc () ++ pr_induction_arg pr_constr h) + | TacOldDestruct h -> + hov 1 (str "Destruct" ++ pr_arg pr_quantified_hypothesis h) + | TacNewDestruct h -> + hov 1 (str "NewDestruct" ++ spc () ++ pr_induction_arg pr_constr h) + | TacDoubleInduction (h1,h2) -> + hov 1 + (str "Double Induction" ++ + pr_arg pr_quantified_hypothesis h1 ++ + pr_arg pr_quantified_hypothesis h2) + | TacDecomposeAnd c -> + hov 1 (str "Decompose Record" ++ pr_arg pr_constr c) + | TacDecomposeOr c -> + hov 1 (str "Decompose Sum" ++ pr_arg pr_constr c) + | TacDecompose (l,c) -> + hov 1 (str "Decompose" ++ spc () ++ + hov 0 (str "[" ++ prlist_with_sep spc (pr_metanum pr_ind) l + ++ str "]")) + | TacSpecialize (n,c) -> + hov 1 (str "Specialize" ++ pr_opt int n ++ pr_with_bindings c) + | TacLApply c -> + hov 1 (str "LApply" ++ pr_constr c) + + (* Automation tactics *) + | TacTrivial (Some []) as x -> pr_atom0 x + | TacTrivial db -> hov 0 (str "Trivial" ++ pr_hintbases db) + | TacAuto (None,Some []) as x -> pr_atom0 x + | TacAuto (n,db) -> hov 0 (str "Auto" ++ pr_opt int n ++ pr_hintbases db) + | TacAutoTDB None as x -> pr_atom0 x + | TacAutoTDB (Some n) -> hov 0 (str "AutoTDB" ++ spc () ++ int n) + | TacDestructHyp (true,(_,id)) -> hov 0 (str "CDHyp" ++ spc () ++ pr_id id) + | TacDestructHyp (false,(_,id)) -> hov 0 (str "DHyp" ++ spc () ++ pr_id id) + | TacDestructConcl as x -> pr_atom0 x + | TacSuperAuto (n,l,b1,b2) -> + hov 0 (str "SuperAuto" ++ pr_opt int n ++ pr_autoarg_adding l ++ + pr_autoarg_destructing b1 ++ pr_autoarg_usingTDB b2) + | TacDAuto (n,p) -> + hov 0 (str "Auto" ++ pr_opt int n ++ str "Decomp" ++ pr_opt int p) + + (* Context management *) + | TacClear l -> + hov 0 + (str "Clear" ++ brk (1,1) ++ prlist_with_sep spc (pr_metanum pr_id) l) + | TacClearBody l -> + hov 0 + (str "Clear" ++ brk (1,1) ++ prlist_with_sep spc (pr_metanum pr_id) l) + | TacMove (b,(_,id1),(_,id2)) -> + (* Rem: only b = true is available for users *) + assert b; + hov 1 + (str "Move" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ + str "after" ++ brk (1,1) ++ pr_id id2) + | TacRename ((_,id1),(_,id2)) -> + hov 1 + (str "Rename" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ + str "into" ++ brk (1,1) ++ pr_id id2) + + (* Constructors *) + | TacLeft l -> hov 1 (str "Left" ++ pr_bindings l) + | TacRight l -> hov 1 (str "Right" ++ pr_bindings l) + | TacSplit l -> hov 1 (str "Split" ++ pr_bindings l) + | TacAnyConstructor (Some t) -> + hov 1 (str "Constructor" ++ pr_arg !pr_rawtac0 t) + | TacAnyConstructor None as t -> pr_atom0 t + | TacConstructor (n,l) -> + hov 1 (str "Constructor" ++ pr_or_meta pr_intarg n ++ pr_bindings l) + + (* Conversion *) + | TacReduce (r,h) -> + hov 1 (pr_red_expr (pr_constr,pr_cst) r ++ pr_clause pr_ident h) + | TacChange (c,h) -> + hov 1 (str "Change" ++ brk (1,1) ++ pr_constr c ++ pr_clause pr_ident h) + + (* Equivalence relations *) + | (TacReflexivity | TacSymmetry) as x -> pr_atom0 x + | TacTransitivity c -> str "Transitivity" ++ pr_arg pr_constr c + +and pr_tactic_seq_body tl = + hv 0 (str "[ " ++ + prlist_with_sep (fun () -> spc () ++ str "| ") prtac tl ++ str " ]") + + (* Strictly closed atomic tactic expressions *) +and pr0 = function + | TacFirst tl -> str "First" ++ spc () ++ pr_tactic_seq_body tl + | TacSolve tl -> str "Solve" ++ spc () ++ pr_tactic_seq_body tl + | TacId -> str "Idtac" + | TacFail 0 -> str "Fail" + | TacAtom (_,t) -> pr_atom0 t + | TacArg c -> pr_tacarg c + | t -> str "(" ++ prtac t ++ str ")" + + (* Semi-closed atomic tactic expressions *) +and pr1 = function + | TacAtom (_,t) -> pr_atom1 t + | TacFail n -> str "Fail " ++ int n + | t -> pr0 t + + (* Orelse tactic expressions (printed as if parsed associating on the right + though the semantics is purely associative) *) +and pr2 = function + | TacOrelse (t1,t2) -> + hov 1 (pr1 t1 ++ str " Orelse" ++ brk (1,1) ++ pr3 t2) + | t -> pr1 t + + (* Non closed prefix tactic expressions *) +and pr3 = function + | TacTry t -> hov 1 (str "Try" ++ spc () ++ pr3 t) + | TacDo (n,t) -> hov 1 (str "Do " ++ int n ++ spc () ++ pr3 t) + | TacRepeat t -> hov 1 (str "Repeat" ++ spc () ++ pr3 t) + | TacProgress t -> hov 1 (str "Progress" ++ spc () ++ pr3 t) + | TacInfo t -> hov 1 (str "Info" ++ spc () ++ pr3 t) + | t -> pr2 t + +and pr4 = function + | t -> pr3 t + + (* THEN and THENS tactic expressions (printed as if parsed + associating on the left though the semantics is purely associative) *) +and pr5 = function + | TacThens (t,tl) -> + hov 1 (pr5 t ++ str ";" ++ spc () ++ pr_tactic_seq_body tl) + | TacThen (t1,t2) -> + hov 1 (pr5 t1 ++ str ";" ++ spc () ++ pr4 t2) + | t -> pr4 t + + (* Ltac tactic expressions *) +and pr6 = function + |(TacAtom _ + | TacThen _ + | TacThens _ + | TacFirst _ + | TacSolve _ + | TacTry _ + | TacOrelse _ + | TacDo _ + | TacRepeat _ + | TacProgress _ + | TacId + | TacFail _ + | TacInfo _) as t -> pr5 t + + | TacAbstract (t,None) -> str "Abstract " ++ pr6 t + | TacAbstract (t,Some s) -> + hov 0 + (str "Abstract " ++ pr6 t ++ spc () ++ str "using" ++ spc () ++ pr_id s) + | TacLetRecIn (l,t) -> + hv 0 + (str "Rec " ++ pr_rec_clauses prtac l ++ + spc () ++ str "In" ++ fnl () ++ prtac t) + | TacLetIn (llc,u) -> + v 0 + (hv 0 (pr_let_clauses pr_tacarg0 llc ++ spc () ++ str "In") ++ fnl () ++ prtac u) + | TacLetCut llc -> + pr_let_clauses pr_tacarg0 + (List.map (fun (id,c,t) -> ((dummy_loc,id),Some c,t)) llc) + ++ fnl () + | TacMatch (t,lrul) -> + hov 0 (str "Match" ++ spc () ++ pr_may_eval Ppconstr.pr_constr t ++ spc() + ++ str "With" + ++ prlist + (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule true prtac r) + lrul) + | TacMatchContext (lr,lrul) -> + hov 0 ( + str (if lr then "Match Reverse Context With" else "Match Context With") + ++ prlist + (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule false prtac r) + lrul) + | TacFun (lvar,body) -> + hov 0 (str "Fun" ++ + prlist pr_funvar lvar ++ spc () ++ str "->" ++ spc () ++ prtac body) + | TacFunRec t -> + hov 0 (str "Rec " ++ pr_rec_clause prtac t) + + | TacArg c -> pr_tacarg c + +and pr_tacarg0 = function + | TacDynamic (_,t) -> str ("") + | MetaNumArg (_,n) -> str ("?" ^ string_of_int n ) + | MetaIdArg (_,s) -> str ("$" ^ s) + | TacVoid -> str "()" + | Reference (RQualid (_,qid)) -> pr_qualid qid + | Reference (RIdent (_,id)) -> pr_id id + | ConstrMayEval (ConstrTerm c) -> str "'" ++ pr_constr c + | ConstrMayEval c -> pr_may_eval pr_constr c + | Integer n -> int n +(* + | Tac of + 'c * (Coqast.t,qualid or_metanum, identifier or_metaid) gen_tactic_expr +*) + | (TacCall _ | Tacexp _) as t -> str "(" ++ pr_tacarg1 t ++ str ")" + +and pr_tacarg1 = function + | TacCall (_,f,l) -> + hov 0 (pr_tacarg0 f ++ spc () ++ prlist_with_sep spc pr_tacarg0 l) + | Tacexp t -> !pr_rawtac t + | t -> pr_tacarg0 t + +and pr_tacarg x = pr_tacarg1 x + +and prtac x = pr6 x + +in (prtac,pr0,pr_match_rule) + +let (pr_raw_tactic,pr_raw_tactic0,pr_match_rule) = + make_pr_tac + (Ppconstr.pr_constr, + pr_metanum pr_qualid, + pr_qualid, + pr_or_meta (fun (loc,id) -> pr_id id), + pr_raw_extend) + +let _ = pr_rawtac := pr_raw_tactic +let _ = pr_rawtac0 := pr_raw_tactic0 + +let (pr_tactic,_,_) = + make_pr_tac + (Printer.prterm, + pr_evaluable_reference, + pr_inductive, + pr_id, + pr_extend) diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli new file mode 100644 index 000000000..b049a6c47 --- /dev/null +++ b/parsing/pptactic.mli @@ -0,0 +1,35 @@ +(****************************************************************************) +(* *) +(* The Coq Proof Assistant *) +(* *) +(* Projet Coq *) +(* *) +(* INRIA LRI-CNRS ENS-CNRS *) +(* Rocquencourt Orsay Lyon *) +(* *) +(****************************************************************************) + +(* $Id$ *) + +open Pp +open Genarg +open Tacexpr +open Proof_type + +val declare_extra_genarg_pprule : + ('a raw_abstract_argument_type * ('a -> std_ppcmds)) -> + ('b closed_abstract_argument_type * ('b -> std_ppcmds)) -> unit + +val declare_extra_tactic_pprule : + string -> + (raw_generic_argument list -> + string * Egrammar.grammar_tactic_production list) + -> (closed_generic_argument list -> + string * Egrammar.grammar_tactic_production list) + -> unit + +val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) -> (pattern_ast,raw_tactic_expr) match_rule -> std_ppcmds + +val pr_raw_tactic : raw_tactic_expr -> std_ppcmds + +val pr_tactic : Proof_type.tactic_expr -> std_ppcmds diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml index 0adecd707..c7da75c36 100644 --- a/parsing/prettyp.ml +++ b/parsing/prettyp.ml @@ -379,10 +379,11 @@ let list_filter_vec f vec = frec (Array.length vec -1) [] (* This is designed to print the contents of an opened section *) -let read_sec_context qid = +let read_sec_context (loc,qid) = let dir = try Nametab.locate_section qid - with Not_found -> error "Unknown section" in + with Not_found -> + user_err_loc (loc,"read_sec_context", str "Unknown section") in let rec get_cxt in_cxt = function | ((sp,Lib.OpenedSection (dir',_)) as hd)::rest -> if dir = dir' then (hd::in_cxt) else get_cxt (hd::in_cxt) rest @@ -410,7 +411,7 @@ let print_eval red_fun env {uj_val=trm;uj_type=typ} = let ntrm = red_fun env Evd.empty trm in (str " = " ++ print_judgment env {uj_val = ntrm; uj_type = typ}) -let print_name qid = +let print_name (loc,qid) = try let sp = Nametab.locate_obj qid in let (sp,lobj) = @@ -440,34 +441,28 @@ let print_name qid = let sp = Syntax_def.locate_syntactic_definition qid in print_syntactic_def " = " sp with Not_found -> - errorlabstrm "print_name" - (pr_qualid qid ++ spc () ++ str "not a defined object") + user_err_loc + (loc,"print_name",pr_qualid qid ++ spc () ++ str "not a defined object") let print_opaque_name qid = let sigma = Evd.empty in let env = Global.env () in let sign = Global.named_context () in - try - let x = global_qualified_reference qid in - match kind_of_term x with - | Const cst -> - let cb = Global.lookup_constant cst in - if cb.const_body <> None then - print_constant true " = " cst - else - error "not a defined constant" - | Ind (sp,_) -> - print_mutual sp - | Construct cstr -> - let ty = Inductive.type_of_constructor env cstr in - print_typed_value (x, ty) - | Var id -> - let (_,c,ty) = lookup_named id env in - print_named_decl (id,c,ty) - | _ -> - assert false - with Not_found -> - errorlabstrm "print_opaque" (pr_qualid qid ++ spc () ++ str "not declared") + match global qid with + | ConstRef cst -> + let cb = Global.lookup_constant cst in + if cb.const_body <> None then + print_constant true " = " cst + else + error "not a defined constant" + | IndRef (sp,_) -> + print_mutual sp + | ConstructRef cstr -> + let ty = Inductive.type_of_constructor env cstr in + print_typed_value (mkConstruct cstr, ty) + | VarRef id -> + let (_,c,ty) = lookup_named id env in + print_named_decl (id,c,ty) let print_local_context () = let env = Lib.contents_after None in diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli index c7cfc7deb..4a65492ae 100644 --- a/parsing/prettyp.mli +++ b/parsing/prettyp.mli @@ -10,6 +10,7 @@ (*i*) open Pp +open Util open Names open Sign open Term @@ -28,8 +29,8 @@ val print_context : bool -> Lib.library_segment -> std_ppcmds val print_library_entry : bool -> (section_path * Lib.node) -> std_ppcmds val print_full_context : unit -> std_ppcmds val print_full_context_typ : unit -> std_ppcmds -val print_sec_context : Nametab.qualid -> std_ppcmds -val print_sec_context_typ : Nametab.qualid -> std_ppcmds +val print_sec_context : Nametab.qualid located -> std_ppcmds +val print_sec_context_typ : Nametab.qualid located -> std_ppcmds val print_judgment : env -> unsafe_judgment -> std_ppcmds val print_safe_judgment : env -> Safe_typing.judgment -> std_ppcmds val print_eval : @@ -38,8 +39,8 @@ val print_eval : val build_inductive : section_path -> int -> global_reference * rel_context * types * identifier array * types array val print_mutual : section_path -> std_ppcmds -val print_name : Nametab.qualid -> std_ppcmds -val print_opaque_name : Nametab.qualid -> std_ppcmds +val print_name : Nametab.qualid located -> std_ppcmds +val print_opaque_name : Nametab.qualid located -> std_ppcmds val print_local_context : unit -> std_ppcmds (*i diff --git a/parsing/printer.ml b/parsing/printer.ml index dae3423be..a3ce32047 100644 --- a/parsing/printer.ml +++ b/parsing/printer.ml @@ -21,81 +21,21 @@ open Declare open Coqast open Ast open Termast +open Extend open Nametab +open Ppconstr let emacs_str s = if !Options.print_emacs then s else "" -let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; - -let pr_global ref = pr_global_env (Global.env()) ref - -let global_const_name sp = - try pr_global (ConstRef sp) - with Not_found -> (* May happen in debug *) - (str ("CONST("^(string_of_path sp)^")")) - -let global_var_name id = - try pr_global (VarRef id) - with Not_found -> (* May happen in debug *) - (str ("SECVAR("^(string_of_id id)^")")) - -let global_ind_name (sp,tyi) = - try pr_global (IndRef (sp,tyi)) - with Not_found -> (* May happen in debug *) - (str ("IND("^(string_of_path sp)^","^(string_of_int tyi)^")")) - -let global_constr_name ((sp,tyi),i) = - try pr_global (ConstructRef ((sp,tyi),i)) - with Not_found -> (* May happen in debug *) - (str ("CONSTRUCT("^(string_of_path sp)^","^(string_of_int tyi) - ^","^(string_of_int i)^")")) - -let globpr gt = match gt with - | Nvar(_,s) -> (pr_id s) - | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) - | Node(_,"CONST",[Path(_,sl)]) -> - global_const_name (section_path sl) - | Node(_,"SECVAR",[Nvar(_,s)]) -> - global_var_name s - | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> - global_ind_name (section_path sl, tyi) - | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> - global_constr_name ((section_path sl, tyi), i) - | Dynamic(_,d) -> - if (Dyn.tag d) = "constr" then (str"") - else dfltpr gt - | gt -> dfltpr gt - -let wrap_exception = function - Anomaly (s1,s2) -> - warning ("Anomaly ("^s1^")"); pp s2; - str"" - | Failure _ | UserError _ | Not_found -> - str"" - | s -> raise s - (* These are the names of the universes where the pp rules for constr and tactic are put (must be consistent with files PPConstr.v and PPTactic.v *) -let constr_syntax_universe = "constr" let tactic_syntax_universe = "tactic" (* This is starting precedence for printing constructions or tactics *) (* Level 9 means no parentheses except for applicative terms (at level 10) *) -let constr_initial_prec = Some ((constr_syntax_universe,(9,0,0)),Extend.L) let tactic_initial_prec = Some ((tactic_syntax_universe,(9,0,0)),Extend.L) -let gentermpr_fail gt = - Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt - -let gentermpr gt = - try gentermpr_fail gt - with s -> wrap_exception s - -(* [at_top] means ids of env must be avoided in bound variables *) -let gentermpr_core at_top env t = - gentermpr (Termast.ast_of_constr at_top env t) - let prterm_env_at_top env = gentermpr_core true env let prterm_env env = gentermpr_core false env let prterm = prterm_env (Global.env()) @@ -123,8 +63,9 @@ let fterm0 = fprterm_env let pr_constant env cst = gentermpr (ast_of_constant env cst) let pr_existential env ev = gentermpr (ast_of_existential env ev) let pr_inductive env ind = gentermpr (ast_of_inductive env ind) -let pr_constructor env cstr = - gentermpr (ast_of_constructor env cstr) +let pr_constructor env cstr = gentermpr (ast_of_constructor env cstr) + +let pr_global = pr_global open Pattern let pr_ref_label = function (* On triche sur le contexte *) @@ -135,13 +76,14 @@ let pr_ref_label = function (* On triche sur le contexte *) let pr_cases_pattern t = gentermpr (Termast.ast_of_cases_pattern t) let pr_rawterm t = gentermpr (Termast.ast_of_rawconstr t) -let pr_pattern_env tenv env t = - gentermpr (Termast.ast_of_pattern tenv env t) +let pr_pattern_env tenv env t = gentermpr (Termast.ast_of_pattern tenv env t) let pr_pattern t = pr_pattern_env (Global.env()) empty_names_context t -let rec gentacpr gt = - Esyntax.genprint default_tacpr tactic_syntax_universe tactic_initial_prec gt +(* +let gentacpr x = Pptactic.prtac x +*) +(* and default_tacpr = function | Nvar(_,s) -> (pr_id s) @@ -159,8 +101,10 @@ and default_tacpr = function (* This should be tactics *) | Node(_,s,[]) -> (str s) +(* TODO | Node(_,s,ta) -> (str s ++ brk(1,2) ++ hov 0 (prlist_with_sep pr_spc gentacpr ta)) +*) | Dynamic(_,d) as gt -> let tg = Dyn.tag d in if tg = "tactic" then (str"") @@ -168,6 +112,7 @@ and default_tacpr = function else if tg = "constr" then (str"") else dfltpr gt | gt -> dfltpr gt +*) let pr_var_decl env (id,c,typ) = let pbody = match c with diff --git a/parsing/printer.mli b/parsing/printer.mli index ff5929279..17b6341fd 100644 --- a/parsing/printer.mli +++ b/parsing/printer.mli @@ -21,7 +21,9 @@ open Termops (*i*) (* These are the entry points for printing terms, context, tac, ... *) -val gentacpr : Coqast.t -> std_ppcmds +(* +val gentacpr : Tacexpr.raw_tactic_expr -> std_ppcmds +*) val prterm_env : env -> constr -> std_ppcmds val prterm_env_at_top : env -> constr -> std_ppcmds diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4 index 5670b2ce5..14bd4498b 100644 --- a/parsing/q_coqast.ml4 +++ b/parsing/q_coqast.ml4 @@ -8,19 +8,12 @@ (* $Id$ *) +open Q_util + let dummy_loc = (0,0) let is_meta s = String.length s > 0 && s.[0] == '$' -let not_impl name x = - let desc = - if Obj.is_block (Obj.repr x) then - "tag = " ^ string_of_int (Obj.tag (Obj.repr x)) - else - "int_val = " ^ string_of_int (Obj.magic x) - in - failwith (" '$' then s else String.sub s 1 (String.length s - 1) @@ -31,11 +24,11 @@ let anti loc x = in <:expr< $anti:e$ >> -(* [expr_of_ast] contributes to translate g_*.ml4 files into g_*.ppo *) +(* [mlexpr_of_ast] contributes to translate g_*.ml4 files into g_*.ppo *) (* This is where $id's (and macros) in ast are translated in ML variables *) (* which will bind their actual ast value *) -let rec expr_of_ast = function +let rec mlexpr_of_ast = function | Coqast.Nmeta loc id -> anti loc id | Coqast.Id loc id when is_meta id -> <:expr< Coqast.Id loc $anti loc id$ >> | Coqast.Node _ "$VAR" [Coqast.Nmeta loc x] -> @@ -53,35 +46,36 @@ let rec expr_of_ast = function (* Obsolte | Coqast.Node _ "$SLAM" [Coqast.Nmeta loc idl; y] -> <:expr< - List.fold_right (Pcoq.slam_ast loc) $anti loc idl$ $expr_of_ast y$ >> + List.fold_right (Pcoq.slam_ast loc) $anti loc idl$ $mlexpr_of_ast y$ >> *) | Coqast.Node loc "$ABSTRACT" [Coqast.Str _ s; x; y] -> - <:expr< - Pcoq.abstract_binders_ast loc $str:s$ $expr_of_ast x$ $expr_of_ast y$ >> + let x = mlexpr_of_ast x in + let y = mlexpr_of_ast y in + <:expr< Ast.abstract_binders_ast loc $str:s$ $x$ $y$ >> | Coqast.Node loc nn al -> let e = expr_list_of_ast_list al in <:expr< Coqast.Node loc $str:nn$ $e$ >> | Coqast.Nvar loc id -> <:expr< Coqast.Nvar loc (Names.id_of_string $str:Names.string_of_id id$) >> | Coqast.Slam loc None a -> - <:expr< Coqast.Slam loc None $expr_of_ast a$ >> + <:expr< Coqast.Slam loc None $mlexpr_of_ast a$ >> | Coqast.Smetalam loc s a -> <:expr< match $anti loc s$ with - [ Coqast.Nvar _ id -> Coqast.Slam loc (Some id) $expr_of_ast a$ - | Coqast.Nmeta _ s -> Coqast.Smetalam loc s $expr_of_ast a$ + [ Coqast.Nvar _ id -> Coqast.Slam loc (Some id) $mlexpr_of_ast a$ + | Coqast.Nmeta _ s -> Coqast.Smetalam loc s $mlexpr_of_ast a$ | _ -> failwith "Slam expects a var or a metavar" ] >> | Coqast.Slam loc (Some s) a -> let se = <:expr< Names.id_of_string $str:Names.string_of_id s$ >> in - <:expr< Coqast.Slam loc (Some $se$) $expr_of_ast a$ >> + <:expr< Coqast.Slam loc (Some $se$) $mlexpr_of_ast a$ >> | Coqast.Num loc i -> <:expr< Coqast.Num loc $int:string_of_int i$ >> | Coqast.Id loc id -> <:expr< Coqast.Id loc $str:id$ >> | Coqast.Str loc str -> <:expr< Coqast.Str loc $str:str$ >> | Coqast.Path loc qid -> let l,a = Names.repr_path qid in - let expr_of_modid id = + let mlexpr_of_modid id = <:expr< Names.id_of_string $str:Names.string_of_id id$ >> in - let e = List.map expr_of_modid (Names.repr_dirpath l) in + let e = List.map mlexpr_of_modid (Names.repr_dirpath l) in let e = expr_list_of_var_list e in <:expr< Coqast.Path loc (Names.make_path (Names.make_dirpath $e$) (Names.id_of_string $str:Names.string_of_id a$)) >> @@ -98,7 +92,7 @@ and expr_list_of_ast_list al = then <:expr< $e1$ >> else <:expr< ( $lid:"@"$ $e1$ $e2$) >> | _ -> - let e1 = expr_of_ast a in + let e1 = mlexpr_of_ast a in let loc = (fst(MLast.loc_of_expr e1), snd(MLast.loc_of_expr e2)) in <:expr< [$e1$ :: $e2$] >> ) al (let loc = dummy_loc in <:expr< [] >>) @@ -110,31 +104,437 @@ and expr_list_of_var_list sl = let loc = (fst (MLast.loc_of_expr e1), snd (MLast.loc_of_expr e2)) in <:expr< [$e1$ :: $e2$] >>) sl <:expr< [] >> - -let rec patt_of_expr e = - let loc = MLast.loc_of_expr e in - match e with - | <:expr< $e1$.$e2$ >> -> <:patt< $patt_of_expr e1$.$patt_of_expr e2$ >> - | <:expr< $e1$ $e2$ >> -> <:patt< $patt_of_expr e1$ $patt_of_expr e2$ >> - | <:expr< loc >> -> <:patt< _ >> - | <:expr< $lid:s$ >> -> <:patt< $lid:s$ >> - | <:expr< $uid:s$ >> -> <:patt< $uid:s$ >> - | <:expr< $str:s$ >> -> <:patt< $str:s$ >> - | <:expr< $anti:e$ >> -> <:patt< $anti:patt_of_expr e$ >> - | _ -> not_impl "patt_of_expr" e + +(* We don't give location for tactic quotation! *) +let loc = dummy_loc + +let mlexpr_of_ident id = + <:expr< Names.id_of_string $str:Names.string_of_id id$ >> + +let mlexpr_of_name = function + | Names.Anonymous -> <:expr< Names.Anonymous >> + | Names.Name id -> + <:expr< Names.Names (Names.id_of_string $str:Names.string_of_id id$) >> + +let mlexpr_of_dirpath dir = + let l = Names.repr_dirpath dir in + <:expr< Names.make_dirpath $mlexpr_of_list mlexpr_of_ident l$ >> + +let mlexpr_of_qualid qid = + let (dir, id) = Nametab.repr_qualid qid in + <:expr< Nametab.make_qualid $mlexpr_of_dirpath dir$ $mlexpr_of_ident id$ >> + +let mlexpr_of_reference = function + | Tacexpr.RQualid (loc,qid) -> <:expr< Tacexpr.RQualid loc $mlexpr_of_qualid qid$ >> + | Tacexpr.RIdent (loc,id) -> <:expr< Tacexpr.RIdent loc $mlexpr_of_ident id$ >> + +let mlexpr_of_bool = function + | true -> <:expr< True >> + | false -> <:expr< False >> + +let mlexpr_of_intro_pattern = function + | Tacexpr.IntroOrAndPattern _ -> failwith "mlexpr_of_intro_pattern: TODO" + | Tacexpr.IntroWildcard -> <:expr< Tacexpr.IntroWildcard >> + | Tacexpr.IntroIdentifier id -> + <:expr< Tacexpr.IntroIdentifier (mlexpr_of_ident loc id) >> + +let mlexpr_of_ident_option = mlexpr_of_option (mlexpr_of_ident) + +let mlexpr_of_or_metanum f = function + | Rawterm.AN (_,a) -> <:expr< Rawterm.AN loc $f a$ >> + | Rawterm.MetaNum (_,n) -> + <:expr< Rawterm.MetaNum loc $mlexpr_of_int n$ >> + +let mlexpr_of_or_metaid f = function + | Tacexpr.AI a -> <:expr< Tacexpr.AI $f a$ >> + | Tacexpr.MetaId (_,id) -> <:expr< Tacexpr.AI $anti loc id$ >> + +let mlexpr_of_quantified_hypothesis = function + | Rawterm.AnonHyp n -> <:expr< Rawterm.AnonHyp $mlexpr_of_int n$ >> + | Rawterm.NamedHyp id -> <:expr< Rawterm.NamedHyp $mlexpr_of_ident id$ >> + +let mlexpr_of_located f (loc,x) = <:expr< (loc, $f x$) >> + +let mlexpr_of_hyp_location = function + | Tacexpr.InHyp id -> + <:expr< Tacexpr.InHyp $mlexpr_of_or_metaid (mlexpr_of_located mlexpr_of_ident) id$ >> + | Tacexpr.InHypType id -> + <:expr< Tacexpr.InHypType $mlexpr_of_or_metaid (mlexpr_of_located mlexpr_of_ident) id$ >> + +(* +let mlexpr_of_red_mode = function + | Closure.UNIFORM -> <:expr< Closure.UNIFORM >> + | Closure.SIMPL -> <:expr< Closure.SIMPL >> + | Closure.WITHBACK -> <:expr< Closure.WITHBACK >> +*) + +let mlexpr_of_red_flags { + Rawterm.rBeta = bb; + Rawterm.rIota = bi; + Rawterm.rZeta = bz; + Rawterm.rDelta = bd; + Rawterm.rConst = l +} = <:expr< { + Rawterm.rBeta = $mlexpr_of_bool bb$; + Rawterm.rIota = $mlexpr_of_bool bi$; + Rawterm.rZeta = $mlexpr_of_bool bz$; + Rawterm.rDelta = $mlexpr_of_bool bd$; + Rawterm.rConst = $mlexpr_of_list (mlexpr_of_or_metanum mlexpr_of_qualid) l$ +} >> + +let mlexpr_of_constr = mlexpr_of_ast + +let mlexpr_of_red_expr = function + | Rawterm.Red b -> <:expr< Rawterm.Red $mlexpr_of_bool b$ >> + | Rawterm.Hnf -> <:expr< Rawterm.Hnf >> + | Rawterm.Simpl -> <:expr< Rawterm.Simpl >> + | Rawterm.Cbv f -> + <:expr< Rawterm.Cbv $mlexpr_of_red_flags f$ >> + | Rawterm.Lazy f -> + <:expr< Rawterm.Lazy $mlexpr_of_red_flags f$ >> + | Rawterm.Unfold l -> + let f1 = mlexpr_of_list mlexpr_of_int in + let f2 = mlexpr_of_or_metanum mlexpr_of_qualid in + let f = mlexpr_of_list (mlexpr_of_pair f1 f2) in + <:expr< Rawterm.Unfold $f l$ >> + | Rawterm.Fold l -> + <:expr< Rawterm.Fold $mlexpr_of_list mlexpr_of_constr l$ >> + | Rawterm.Pattern l -> + let f1 = mlexpr_of_list mlexpr_of_int in + let f = mlexpr_of_list (mlexpr_of_pair f1 mlexpr_of_constr) in + <:expr< Rawterm.Pattern $f l$ >> + | Rawterm.ExtraRedExpr (s,l) -> + let l = mlexpr_of_list mlexpr_of_constr l in + <:expr< Rawterm.ExtraRedExpr $mlexpr_of_string s$ $l$ >> + +let rec mlexpr_of_argtype loc = function + | Genarg.BoolArgType -> <:expr< Genarg.BoolArgType >> + | Genarg.IntArgType -> <:expr< Genarg.IntArgType >> + | Genarg.IntOrVarArgType -> <:expr< Genarg.IntOrVarArgType >> + | Genarg.QualidArgType -> <:expr< Genarg.QualidArgType >> + | Genarg.PreIdentArgType -> <:expr< Genarg.PreIdentArgType >> + | Genarg.IdentArgType -> <:expr< Genarg.IdentArgType >> + | Genarg.StringArgType -> <:expr< Genarg.StringArgType >> + | Genarg.QuantHypArgType -> <:expr< Genarg.QuantHypArgType >> + | Genarg.CastedOpenConstrArgType -> <:expr< Genarg.CastedOpenConstrArgType >> + | Genarg.ConstrWithBindingsArgType -> <:expr< Genarg.ConstrWithBindingsArgType >> + | Genarg.RedExprArgType -> <:expr< Genarg.RedExprArgType >> + | Genarg.TacticArgType -> <:expr< Genarg.TacticArgType >> + | Genarg.ConstrArgType -> <:expr< Genarg.ConstrArgType >> + | Genarg.ConstrMayEvalArgType -> <:expr< Genarg.ConstrMayEvalArgType >> + | Genarg.List0ArgType t -> <:expr< Genarg.List0ArgType $mlexpr_of_argtype loc t$ >> + | Genarg.List1ArgType t -> <:expr< Genarg.List1ArgType $mlexpr_of_argtype loc t$ >> + | Genarg.OptArgType t -> <:expr< Genarg.OptArgType $mlexpr_of_argtype loc t$ >> + | Genarg.PairArgType (t1,t2) -> + let t1 = mlexpr_of_argtype loc t1 in + let t2 = mlexpr_of_argtype loc t2 in + <:expr< Genarg.PairArgType $t1$ $t2$ >> + | Genarg.ExtraArgType s -> <:expr< Genarg.ExtraArgType $str:s$ >> + +let rec mlexpr_of_may_eval f = function + | Rawterm.ConstrEval (r,c) -> + <:expr< Rawterm.ConstrEval $mlexpr_of_red_expr r$ $f c$ >> + | Rawterm.ConstrContext ((loc,id),c) -> + let id = mlexpr_of_ident id in + <:expr< Rawterm.ConstrContext (loc,$id$) $f c$ >> + | Rawterm.ConstrTypeOf c -> + <:expr< Rawterm.ConstrTypeOf $mlexpr_of_constr c$ >> + | Rawterm.ConstrTerm c -> + <:expr< Rawterm.ConstrTerm $mlexpr_of_constr c$ >> + +let mlexpr_of_binding_kind = function + | Rawterm.ExplicitBindings l -> + let l = mlexpr_of_list (mlexpr_of_pair mlexpr_of_quantified_hypothesis mlexpr_of_constr) l in + <:expr< Rawterm.ExplicitBindings $l$ >> + | Rawterm.ImplicitBindings l -> + let l = mlexpr_of_list mlexpr_of_constr l in + <:expr< Rawterm.ImplicitBindings $l$ >> + | Rawterm.NoBindings -> + <:expr< Rawterm.NoBindings >> + +let mlexpr_of_induction_arg = function + | Tacexpr.ElimOnConstr c -> + <:expr< Tacexpr.ElimOnConstr $mlexpr_of_constr c$ >> + | Tacexpr.ElimOnIdent (_,id) -> + <:expr< Tacexpr.ElimOnIdent loc $mlexpr_of_ident id$ >> + | Tacexpr.ElimOnAnonHyp n -> + <:expr< Tacexpr.ElimOnAnonHyp $mlexpr_of_int n$ >> + +let mlexpr_of_binding = mlexpr_of_pair mlexpr_of_binding_kind mlexpr_of_constr + +let mlexpr_of_constr_with_binding = + mlexpr_of_pair mlexpr_of_constr mlexpr_of_binding_kind + +let mlexpr_of_clause_pattern _ = failwith "mlexpr_of_clause_pattern: TODO" + +let mlexpr_of_pattern_ast = mlexpr_of_ast + +let mlexpr_of_entry_type = function + _ -> failwith "mlexpr_of_entry_type: TODO" + +let mlexpr_of_match_pattern = function + | Tacexpr.Term t -> <:expr< Tacexpr.Term $mlexpr_of_pattern_ast t$ >> + | Tacexpr.Subterm (ido,t) -> + <:expr< Tacexpr.Subterm $mlexpr_of_option mlexpr_of_ident ido$ $mlexpr_of_pattern_ast t$ >> + +let mlexpr_of_match_context_hyps = function + | Tacexpr.NoHypId l -> <:expr< Tacexpr.NoHypId $mlexpr_of_match_pattern l$ >> + | Tacexpr.Hyp (id,l) -> + let f = mlexpr_of_located mlexpr_of_ident in + <:expr< Tacexpr.Hyp $f id$ $mlexpr_of_match_pattern l$ >> + +let mlexpr_of_match_rule f = function + | Tacexpr.Pat (l,mp,t) -> <:expr< Tacexpr.Pat $mlexpr_of_list mlexpr_of_match_context_hyps l$ $mlexpr_of_match_pattern mp$ $f t$ >> + | Tacexpr.All t -> <:expr< Tacexpr.All $f t$ >> + +let rec mlexpr_of_atomic_tactic = function + (* Basic tactics *) + | Tacexpr.TacIntroPattern pl -> + let pl = mlexpr_of_list mlexpr_of_intro_pattern pl in + <:expr< Tacexpr.TacIntroPattern $pl$ >> + | Tacexpr.TacIntrosUntil h -> + <:expr< Tacexpr.TacIntrosUntil $mlexpr_of_quantified_hypothesis h$ >> + | Tacexpr.TacIntroMove (idopt,idopt') -> + let idopt = mlexpr_of_ident_option idopt in + let idopt'=mlexpr_of_option (mlexpr_of_located mlexpr_of_ident) idopt' in + <:expr< Tacexpr.TacIntroMove $idopt$ $idopt'$ >> + | Tacexpr.TacAssumption -> + <:expr< Tacexpr.TacAssumption >> + | Tacexpr.TacExact c -> + <:expr< Tacexpr.TacExact $mlexpr_of_constr c$ >> + | Tacexpr.TacApply cb -> + <:expr< Tacexpr.TacApply $mlexpr_of_constr_with_binding cb$ >> + | Tacexpr.TacElim (cb,cbo) -> + let cb = mlexpr_of_constr_with_binding cb in + let cbo = mlexpr_of_option mlexpr_of_constr_with_binding cbo in + <:expr< Tacexpr.TacElim $cb$ $cbo$ >> + | Tacexpr.TacElimType c -> + <:expr< Tacexpr.TacElimType $mlexpr_of_constr c$ >> + | Tacexpr.TacCase cb -> + let cb = mlexpr_of_constr_with_binding cb in + <:expr< Tacexpr.TacCase $cb$ >> + | Tacexpr.TacCaseType c -> + <:expr< Tacexpr.TacCaseType $mlexpr_of_constr c$ >> + | Tacexpr.TacFix (ido,n) -> + let ido = mlexpr_of_ident_option ido in + let n = mlexpr_of_int n in + <:expr< Tacexpr.TacFix $ido$ $n$ >> + | Tacexpr.TacMutualFix (id,n,l) -> + let id = mlexpr_of_ident id in + let n = mlexpr_of_int n in + let f =mlexpr_of_triple mlexpr_of_ident mlexpr_of_int mlexpr_of_constr in + let l = mlexpr_of_list f l in + <:expr< Tacexpr.TacMutualFix $id$ $n$ $l$ >> + | Tacexpr.TacCofix ido -> + let ido = mlexpr_of_ident_option ido in + <:expr< Tacexpr.TacCofix $ido$ >> + | Tacexpr.TacMutualCofix (id,l) -> + let id = mlexpr_of_ident id in + let f = mlexpr_of_pair mlexpr_of_ident mlexpr_of_constr in + let l = mlexpr_of_list f l in + <:expr< Tacexpr.TacMutualCofix $id$ $l$ >> + + | Tacexpr.TacCut c -> + <:expr< Tacexpr.TacCut $mlexpr_of_constr c$ >> + | Tacexpr.TacTrueCut (ido,c) -> + let ido = mlexpr_of_ident_option ido in + <:expr< Tacexpr.TacTrueCut $ido$ $mlexpr_of_constr c$ >> + | Tacexpr.TacForward (b,na,c) -> + <:expr< Tacexpr.TacForward $mlexpr_of_bool b$ $mlexpr_of_name na$ $mlexpr_of_constr c$ >> + | Tacexpr.TacGeneralize cl -> + <:expr< Tacexpr.TacGeneralize $mlexpr_of_list mlexpr_of_constr cl$ >> + | Tacexpr.TacGeneralizeDep c -> + <:expr< Tacexpr.TacGeneralizeDep $mlexpr_of_constr c$ >> + | Tacexpr.TacLetTac (id,c,cl) -> + let id = mlexpr_of_ident id in + let cl = mlexpr_of_clause_pattern cl in + <:expr< Tacexpr.TacLetTac $id$ $mlexpr_of_constr c$ $cl$ >> + + (* Derived basic tactics *) + | Tacexpr.TacOldInduction h -> + <:expr< Tacexpr.TacOldInduction $mlexpr_of_quantified_hypothesis h$ >> + | Tacexpr.TacNewInduction c -> + <:expr< Tacexpr.TacNewInduction $mlexpr_of_induction_arg c$ >> + | Tacexpr.TacOldDestruct h -> + <:expr< Tacexpr.TacOldDestruct $mlexpr_of_quantified_hypothesis h$ >> + | Tacexpr.TacNewDestruct c -> + <:expr< Tacexpr.TacNewDestruct $mlexpr_of_induction_arg c$ >> + + (* Context management *) + | Tacexpr.TacClear l -> + let l = mlexpr_of_list (mlexpr_of_or_metanum mlexpr_of_ident) l in + <:expr< Tacexpr.TacClear $l$ >> + | Tacexpr.TacClearBody l -> + let l = mlexpr_of_list (mlexpr_of_or_metanum mlexpr_of_ident) l in + <:expr< Tacexpr.TacClearBody $l$ >> + | Tacexpr.TacMove (dep,id1,id2) -> + <:expr< Tacexpr.TacMove $mlexpr_of_bool dep$ + $mlexpr_of_located mlexpr_of_ident id1$ + $mlexpr_of_located mlexpr_of_ident id2$ >> + + (* Constructors *) + | Tacexpr.TacLeft l -> + <:expr< Tacexpr.TacLeft $mlexpr_of_binding_kind l$>> + | Tacexpr.TacRight l -> + <:expr< Tacexpr.TacRight $mlexpr_of_binding_kind l$>> + | Tacexpr.TacSplit l -> + <:expr< Tacexpr.TacSplit $mlexpr_of_binding_kind l$>> + | Tacexpr.TacAnyConstructor t -> + <:expr< Tacexpr.TacAnyConstructor $mlexpr_of_option mlexpr_of_tactic t$>> + | Tacexpr.TacConstructor (n,l) -> + let n = mlexpr_of_or_metaid mlexpr_of_int n in + <:expr< Tacexpr.TacConstructor $n$ $mlexpr_of_binding_kind l$>> + + (* Conversion *) + | Tacexpr.TacReduce (r,cl) -> + let l = mlexpr_of_list mlexpr_of_hyp_location cl in + <:expr< Tacexpr.TacReduce $mlexpr_of_red_expr r$ $l$ >> + | Tacexpr.TacChange (c,cl) -> + let l = mlexpr_of_list mlexpr_of_hyp_location cl in + <:expr< Tacexpr.TacChange $mlexpr_of_constr c$ $l$ >> + + (* Equivalence relations *) + | Tacexpr.TacReflexivity -> <:expr< Tacexpr.TacReflexivity >> + | Tacexpr.TacSymmetry -> <:expr< Tacexpr.TacSymmetry >> + | Tacexpr.TacTransitivity c -> <:expr< Tacexpr.TacTransitivity $mlexpr_of_constr c$ >> + + (* Automation tactics *) + | Tacexpr.TacAuto (n,l) -> + let n = mlexpr_of_option mlexpr_of_int n in + let l = mlexpr_of_option (mlexpr_of_list mlexpr_of_string) l in + <:expr< Tacexpr.TacAuto $n$ $l$ >> +(* + | Tacexpr.TacTrivial l -> + let l = mlexpr_of_option (mlexpr_of_list mlexpr_of_string) l in + <:expr< Tacexpr.TacTrivial $l$ >> +*) + +(* + | Tacexpr.TacExtend (s,l) -> + let l = mlexpr_of_list mlexpr_of_tactic_arg l in + let loc = MLast.loc_of_expr l in + <:expr< Tacexpr.TacExtend $mlexpr_of_string s$ $l$ >> +*) + | _ -> failwith "Quotation of atomic tactic expressions: TODO" + +and mlexpr_of_tactic = function + | Tacexpr.TacAtom (loc,t) -> + <:expr< Tacexpr.TacAtom loc $mlexpr_of_atomic_tactic t$ >> + | Tacexpr.TacThen (t1,t2) -> + <:expr< Tacexpr.TacThen $mlexpr_of_tactic t1$ $mlexpr_of_tactic t2$ >> + | Tacexpr.TacThens (t,tl) -> + <:expr< Tacexpr.TacThens $mlexpr_of_tactic t$ $mlexpr_of_list mlexpr_of_tactic tl$>> + | Tacexpr.TacFirst tl -> + <:expr< Tacexpr.TacFirst $mlexpr_of_list mlexpr_of_tactic tl$ >> + | Tacexpr.TacSolve tl -> + <:expr< Tacexpr.TacSolve $mlexpr_of_list mlexpr_of_tactic tl$ >> + | Tacexpr.TacTry t -> + <:expr< Tacexpr.TacTry $mlexpr_of_tactic t$ >> + | Tacexpr.TacOrelse (t1,t2) -> + <:expr< Tacexpr.TacOrelse $mlexpr_of_tactic t1$ $mlexpr_of_tactic t2$ >> + | Tacexpr.TacDo (n,t) -> + <:expr< Tacexpr.TacDo $int:string_of_int n$ $mlexpr_of_tactic t$ >> + | Tacexpr.TacRepeat t -> + <:expr< Tacexpr.TacRepeat $mlexpr_of_tactic t$ >> + | Tacexpr.TacProgress t -> + <:expr< Tacexpr.TacProgress $mlexpr_of_tactic t$ >> + | Tacexpr.TacId -> let loc = dummy_loc in <:expr< Tacexpr.TacId >> + | Tacexpr.TacFail n -> + let loc = dummy_loc in <:expr< Tacexpr.TacFail $int:string_of_int n$ >> +(* + | Tacexpr.TacInfo t -> TacInfo (loc,f t) + + | Tacexpr.TacRec (id,(idl,t)) -> TacRec (loc,(id,(idl,f t))) + | Tacexpr.TacRecIn (l,t) -> TacRecIn(loc,List.map (fun (id,t) -> (id,f t)) l,f t) +*) + | Tacexpr.TacLetIn (l,t) -> + let f = + mlexpr_of_triple + (mlexpr_of_pair (fun _ -> <:expr< loc >>) mlexpr_of_ident) + (mlexpr_of_option (mlexpr_of_may_eval mlexpr_of_constr)) + mlexpr_of_tactic_arg in + <:expr< Tacexpr.TacLetIn $mlexpr_of_list f l$ $mlexpr_of_tactic t$ >> +(* + | Tacexpr.TacLetCut of (identifier * t * tactic_expr) list + | Tacexpr.TacMatch of t * (t * tactic_expr) list +*) + | Tacexpr.TacMatchContext (lr,l) -> + <:expr< Tacexpr.TacMatchContext + $mlexpr_of_bool lr$ + $mlexpr_of_list (mlexpr_of_match_rule mlexpr_of_tactic) l$>> +(* + | Tacexpr.TacFun of loc * tactic_fun_ast + | Tacexpr.TacFunRec of loc * identifier * tactic_fun_ast +*) +(* + | Tacexpr.TacArg (Tacexpr.AstTacArg (Coqast.Nmeta loc id)) -> anti loc id +*) + | Tacexpr.TacArg (Tacexpr.MetaIdArg (_,id)) -> anti loc id + | Tacexpr.TacArg t -> + <:expr< Tacexpr.TacArg $mlexpr_of_tactic_arg t$ >> + | _ -> failwith "Quotation of tactic expressions: TODO" + +and mlexpr_of_tactic_arg = function + | Tacexpr.MetaIdArg (loc,id) -> anti loc id + | Tacexpr.MetaNumArg (loc,n) -> + <:expr< Tacexpr.MetaNumArg loc $mlexpr_of_int n$ >> +(* + | Tacexpr.Identifier id -> + <:expr< Tacexpr.Identifier $mlexpr_of_ident id$ >> +*) +(* + | Tacexpr.AstTacArg t -> + <:expr< Tacexpr.AstTacArg $mlexpr_of_ast t$ >> +*) + | Tacexpr.TacCall (loc,t,tl) -> + <:expr< Tacexpr.TacCall loc $mlexpr_of_tactic_arg t$ $mlexpr_of_list mlexpr_of_tactic_arg tl$>> + | Tacexpr.Tacexp t -> + <:expr< Tacexpr.Tacexp $mlexpr_of_tactic t$ >> + | Tacexpr.ConstrMayEval c -> + <:expr< Tacexpr.ConstrMayEval $mlexpr_of_may_eval mlexpr_of_constr c$ >> +(* + | Tacexpr.Constr c -> + <:expr< Tacexpr.Constr $mlexpr_of_constr c$ >> +*) +(* + | Tacexpr.Qualid q -> + <:expr< Tacexpr.Qualid $mlexpr_of_qualid q$ >> +*) + | Tacexpr.Reference r -> + <:expr< Tacexpr.Reference $mlexpr_of_reference r$ >> +(* + | Tacexpr.TacArgGen q -> + <:expr< Tacexpr.TacArgGen $mlexpr_of_argtype q$ >> +*) + | _ -> failwith "mlexpr_of_tactic_arg: TODO" let f e = let ee s = - expr_of_ast (Pcoq.Gram.Entry.parse e + mlexpr_of_ast (Pcoq.Gram.Entry.parse e + (Pcoq.Gram.parsable (Stream.of_string s))) + in + let ep s = patt_of_expr (ee s) in + Quotation.ExAst (ee, ep) + +let fconstr e = + let ee s = + mlexpr_of_constr (Pcoq.Gram.Entry.parse e + (Pcoq.Gram.parsable (Stream.of_string s))) + in + let ep s = patt_of_expr (ee s) in + Quotation.ExAst (ee, ep) + +let ftac e = + let ee s = + mlexpr_of_tactic (Pcoq.Gram.Entry.parse e (Pcoq.Gram.parsable (Stream.of_string s))) in let ep s = patt_of_expr (ee s) in Quotation.ExAst (ee, ep) let _ = - Quotation.add "constr" (f Pcoq.Constr.constr_eoi); - Quotation.add "tactic" (f Pcoq.Tactic.tactic_eoi); - Quotation.add "vernac" (f Pcoq.Vernac_.vernac_eoi); + Quotation.add "constr" (fconstr Pcoq.Constr.constr_eoi); + Quotation.add "tactic" (ftac Pcoq.Tactic.tactic_eoi); +(* Quotation.add "vernac" (f Pcoq.Vernac_.vernac_eoi);*) Quotation.add "ast" (f Pcoq.Prim.ast_eoi); Quotation.default := "constr" - diff --git a/parsing/q_util.ml4 b/parsing/q_util.ml4 new file mode 100644 index 000000000..85259d6e4 --- /dev/null +++ b/parsing/q_util.ml4 @@ -0,0 +1,63 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* > -> <:patt< $patt_of_expr e1$.$patt_of_expr e2$ >> + | <:expr< $e1$ $e2$ >> -> <:patt< $patt_of_expr e1$ $patt_of_expr e2$ >> + | <:expr< loc >> -> <:patt< _ >> + | <:expr< $lid:s$ >> -> <:patt< $lid:s$ >> + | <:expr< $uid:s$ >> -> <:patt< $uid:s$ >> + | <:expr< $str:s$ >> -> <:patt< $str:s$ >> + | <:expr< $anti:e$ >> -> <:patt< $anti:patt_of_expr e$ >> + | _ -> not_impl "patt_of_expr" e + +let mlexpr_of_list f l = + List.fold_right + (fun e1 e2 -> + let e1 = f e1 in + let loc = (fst (MLast.loc_of_expr e1), snd (MLast.loc_of_expr e2)) in + <:expr< [$e1$ :: $e2$] >>) + l (let loc = dummy_loc in <:expr< [] >>) + +let mlexpr_of_pair m1 m2 (a1,a2) = + let e1 = m1 a1 and e2 = m2 a2 in + let loc = (fst (MLast.loc_of_expr e1), snd (MLast.loc_of_expr e2)) in + <:expr< ($e1$, $e2$) >> + +let mlexpr_of_triple m1 m2 m3 (a1,a2,a3)= + let e1 = m1 a1 and e2 = m2 a2 and e3 = m3 a3 in + let loc = (fst (MLast.loc_of_expr e1), snd (MLast.loc_of_expr e3)) in + <:expr< ($e1$, $e2$, $e3$) >> + +(* We don't give location for tactic quotation! *) +let loc = dummy_loc + +let mlexpr_of_int n = <:expr< $int:string_of_int n$ >> + +let mlexpr_of_string s = <:expr< $str:s$ >> + +let mlexpr_of_option f = function + | None -> <:expr< None >> + | Some e -> <:expr< Some $f e$ >> diff --git a/parsing/q_util.mli b/parsing/q_util.mli new file mode 100644 index 000000000..6be871a0a --- /dev/null +++ b/parsing/q_util.mli @@ -0,0 +1,28 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* MLast.patt + +val mlexpr_of_list : ('a -> MLast.expr) -> 'a list -> MLast.expr + +val mlexpr_of_pair : + ('a -> MLast.expr) -> ('b -> MLast.expr) + -> 'a * 'b -> MLast.expr + +val mlexpr_of_triple : + ('a -> MLast.expr) -> ('b -> MLast.expr) -> ('c -> MLast.expr) + -> 'a * 'b * 'c -> MLast.expr + +val mlexpr_of_int : int -> MLast.expr + +val mlexpr_of_string : string -> MLast.expr + +val mlexpr_of_option : ('a -> MLast.expr) -> 'a option -> MLast.expr + diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 9590c3b76..fb8a6c8a4 100755 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -20,6 +20,7 @@ open Declare open Term open Termops open Rawterm +open Nametab (* usage qque peu general: utilise aussi dans record *) @@ -44,8 +45,7 @@ type coe_info_typ = { coe_value : unsafe_judgment; coe_strength : strength; coe_is_identity : bool; - coe_param : int; - mutable coe_hide : bool } + coe_param : int } type cl_index = int type coe_index = int @@ -137,18 +137,6 @@ let coercion_exists coe = let coe_of_reference x = x -let hide_coercion coe = - let _,coe_info = coercion_info coe in - if coe_info.coe_hide then Some coe_info.coe_param else None - -let set_coercion_visibility b coe = - let _,coe_info = coercion_info coe in - coe_info.coe_hide <- not b - -let is_coercion_visible coe = - let _,coe_info = coercion_info coe in - not coe_info.coe_hide - let coercion_params coe_info = coe_info.coe_param (* coercion_info_from_index : int -> coe_typ * coe_info_typ *) @@ -244,9 +232,12 @@ let strength_of_cl = function let string_of_class = function | CL_FUN -> "FUNCLASS" | CL_SORT -> "SORTCLASS" - | CL_CONST sp -> string_of_id (id_of_global (Global.env()) (ConstRef sp)) - | CL_IND sp -> string_of_id (id_of_global (Global.env()) (IndRef sp)) - | CL_SECVAR sp -> string_of_id (id_of_global (Global.env()) (VarRef sp)) + | CL_CONST sp -> + string_of_qualid (shortest_qualid_of_global (Global.env()) (ConstRef sp)) + | CL_IND sp -> + string_of_qualid (shortest_qualid_of_global (Global.env()) (IndRef sp)) + | CL_SECVAR sp -> + string_of_qualid (shortest_qualid_of_global (Global.env()) (VarRef sp)) (* coercion_value : coe_index -> unsafe_judgment * bool *) @@ -344,8 +335,7 @@ let declare_coercion coef v stre ~isid ~src:cls ~target:clt ~params:ps = { coe_value = v; coe_strength = stre; coe_is_identity = isid; - coe_param = ps; - coe_hide = true }), + coe_param = ps }), cls, clt)) let coercion_strength v = v.coe_strength @@ -357,3 +347,32 @@ let get_coercion_value v = v.coe_value.uj_val let classes () = !class_tab let coercions () = !coercion_tab let inheritance_graph () = !inheritance_graph + +let coercion_of_qualid qid = + let ref = Nametab.global qid in + let coe = coe_of_reference ref in + if not (coercion_exists coe) then + errorlabstrm "try_add_coercion" + (Nametab.pr_global_env (Global.env()) ref ++ str" is not a coercion"); + coe + +module CoercionPrinting = + struct + type t = coe_typ + let encode = coercion_of_qualid + let printer x = pr_global_env (Global.env()) x + let key = Goptions.SecondaryTable ("Printing","Coercion") + let title = "Explicitly printed coercions: " + let member_message x b = + str "Explicit printing of coercion " ++ printer x ++ + str (if b then " is set" else " is unset") + let synchronous = true + end + +module PrintingCoercion = Goptions.MakeRefTable(CoercionPrinting) + +let hide_coercion coe = + if not (PrintingCoercion.active coe) then + let _,coe_info = coercion_info coe in + Some coe_info.coe_param + else None diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 3861d8d35..cd5f31db8 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -15,6 +15,7 @@ open Term open Evd open Environ open Declare +open Nametab (*i*) (*s This is the type of class kinds *) @@ -80,15 +81,6 @@ val coercion_info_from_index : coe_index -> coe_typ * coe_info_typ val coercion_value : coe_index -> (unsafe_judgment * bool) -(*s This is for printing purpose *) - -(* [hide_coercion] returns the number of params to skip if the coercion must - be hidden, [None] otherwise; it raises [Not_found] if not a coercion *) -val hide_coercion : coe_typ -> int option - -val set_coercion_visibility : bool -> coe_typ -> unit -val is_coercion_visible : coe_typ -> bool - (*s Lookup functions for coercion paths *) val lookup_path_between : cl_index * cl_index -> inheritance_path val lookup_path_to_fun_from : cl_index -> inheritance_path @@ -115,9 +107,15 @@ val install_path_printer : ((cl_index * cl_index) * inheritance_path -> std_ppcmds) -> unit (*i*) -(* This is for printing purpose *) +(*s This is for printing purpose *) val string_of_class : cl_typ -> string val get_coercion_value : coe_info_typ -> constr val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list val classes : unit -> (int * (cl_typ * cl_info_typ)) list val coercions : unit -> (int * (coe_typ * coe_info_typ)) list + +(* [hide_coercion] returns the number of params to skip if the coercion must + be hidden, [None] otherwise; it raises [Not_found] if not a coercion *) +val hide_coercion : coe_typ -> int option + + diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 5e718289c..faa5e9e46 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -28,13 +28,8 @@ open Nametab (****************************************************************************) (* Tools for printing of Cases *) -let encode_inductive ref = - let indsp = match ref with - | IndRef indsp -> indsp - | _ -> - errorlabstrm "indsp_of_id" - (pr_global_env (Global.env()) ref ++ - str" is not an inductive type") in +let encode_inductive qid = + let indsp = global_inductive qid in let constr_lengths = mis_constr_nargs indsp in (indsp,constr_lengths) @@ -47,59 +42,64 @@ let isomorphic_to_bool lc = let isomorphic_to_tuple lc = (Array.length lc = 1) +let encode_bool (loc,_ as locqid) = + let (_,lc as x) = encode_inductive locqid in + if not (isomorphic_to_bool lc) then + user_err_loc (loc,"encode_if", + str "This type cannot be seen as a boolean type"); + x + +let encode_tuple (loc,_ as locqid) = + let (_,lc as x) = encode_inductive locqid in + if not (isomorphic_to_tuple lc) then + user_err_loc (loc,"encode_tuple", + str "This type cannot be seen as a tuple type"); + x + module PrintingCasesMake = functor (Test : sig - val test : int array -> bool - val error_message : string - val member_message : global_reference -> bool -> string + val encode : qualid located -> inductive * int array + val member_message : std_ppcmds -> bool -> std_ppcmds val field : string val title : string end) -> struct type t = inductive * int array - let encode = encode_inductive - let check (_,lc) = - if not (Test.test lc) then - errorlabstrm "check_encode" (str Test.error_message) - let printer (ind,_) = - pr_id (basename (path_of_inductive (Global.env()) ind)) + let encode = Test.encode + let printer (ind,_) = pr_global_env (Global.env()) (IndRef ind) let key = Goptions.SecondaryTable ("Printing",Test.field) let title = Test.title - let member_message = Test.member_message + let member_message x = Test.member_message (printer x) let synchronous = true end module PrintingCasesIf = PrintingCasesMake (struct - let test = isomorphic_to_bool - let error_message = "This type cannot be seen as a boolean type" + let encode = encode_bool let field = "If" let title = "Types leading to pretty-printing of Cases using a `if' form: " - let member_message ref b = - let s = string_of_qualid(shortest_qualid_of_global (Global.env()) ref) in - if b then - "Cases on elements of " ^ s ^ " are printed using a `if' form" - else - "Cases on elements of " ^ s ^ " are not printed using `if' form" + let member_message s b = + str "Cases on elements of " ++ s ++ + str + (if b then " are printed using a `if' form" + else " are not printed using a `if' form") end) module PrintingCasesLet = PrintingCasesMake (struct - let test = isomorphic_to_tuple - let error_message = "This type cannot be seen as a tuple type" + let encode = encode_tuple let field = "Let" let title = "Types leading to a pretty-printing of Cases using a `let' form:" - let member_message ref b = - let s = string_of_qualid(shortest_qualid_of_global (Global.env()) ref) in - if b then - "Cases on elements of " ^ s ^ " are printed using a `let' form" - else - "Cases on elements of " ^ s ^ " are not printed using a `let' form" + let member_message s b = + str "Cases on elements of " ++ s ++ + str + (if b then " are printed using a `let' form" + else " are not printed using a `let' form") end) -module PrintingIf = Goptions.MakeIdentTable(PrintingCasesIf) -module PrintingLet = Goptions.MakeIdentTable(PrintingCasesLet) +module PrintingIf = Goptions.MakeRefTable(PrintingCasesIf) +module PrintingLet = Goptions.MakeRefTable(PrintingCasesLet) let force_let ci = let indsp = ci.ci_ind in diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli index 14b9cd5e1..65fefb83f 100644 --- a/pretyping/indrec.mli +++ b/pretyping/indrec.mli @@ -52,3 +52,4 @@ val make_rec_branch_arg : val declare_eliminations : mutual_inductive -> unit val lookup_eliminator : inductive -> sorts_family -> constr val elimination_suffix : sorts_family -> string +val make_elimination_ident : identifier -> sorts_family -> identifier diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml index f69d22dfc..c7e2b0eb3 100644 --- a/pretyping/rawterm.ml +++ b/pretyping/rawterm.ml @@ -32,6 +32,17 @@ type fix_kind = RFix of (int array * int) | RCoFix of int type binder_kind = BProd | BLambda | BLetIn +type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier + +type 'a explicit_substitution = (quantified_hypothesis * 'a) list + +type 'a substitution = + | ImplicitBindings of 'a list + | ExplicitBindings of 'a explicit_substitution + | NoBindings + +type 'a with_bindings = 'a * 'a substitution + type hole_kind = | ImplicitArg of global_reference * int | AbstractionType of name @@ -46,7 +57,6 @@ type 'ctxt reference = | RVar of identifier | REVar of int * 'ctxt -(*i Pas beau ce constr dans rawconstr, mais mal compris ce ctxt des ref i*) type rawconstr = | RRef of loc * global_reference | RVar of loc * identifier @@ -115,3 +125,30 @@ let set_loc_of_rawconstr loc = function | RDynamic (_,d) -> RDynamic (loc,d) let join_loc (deb1,_) (_,fin2) = (deb1,fin2) + +type 'a raw_red_flag = { + rBeta : bool; + rIota : bool; + rZeta : bool; + rDelta : bool; (* true = delta all but rConst; false = delta only on rConst*) + rConst : 'a list +} + +type ('a,'b) red_expr_gen = + | Red of bool + | Hnf + | Simpl + | Cbv of 'b raw_red_flag + | Lazy of 'b raw_red_flag + | Unfold of (int list * 'b) list + | Fold of 'a list + | Pattern of (int list * 'a) list + | ExtraRedExpr of string * 'a list + +type 'a or_metanum = AN of loc * 'a | MetaNum of loc * int + +type 'a may_eval = + | ConstrTerm of 'a + | ConstrEval of ('a, qualid or_metanum) red_expr_gen * 'a + | ConstrContext of (loc * identifier) * 'a + | ConstrTypeOf of 'a diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli index 973bac719..3a6d68115 100644 --- a/pretyping/rawterm.mli +++ b/pretyping/rawterm.mli @@ -31,6 +31,17 @@ type fix_kind = RFix of (int array * int) | RCoFix of int type binder_kind = BProd | BLambda | BLetIn +type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier + +type 'a explicit_substitution = (quantified_hypothesis * 'a) list + +type 'a substitution = + | ImplicitBindings of 'a list + | ExplicitBindings of 'a explicit_substitution + | NoBindings + +type 'a with_bindings = 'a * 'a substitution + type hole_kind = | ImplicitArg of global_reference * int | AbstractionType of name @@ -81,3 +92,30 @@ val dummy_loc : loc val loc_of_rawconstr : rawconstr -> loc val set_loc_of_rawconstr : loc -> rawconstr -> rawconstr val join_loc : loc -> loc -> loc + +type 'a raw_red_flag = { + rBeta : bool; + rIota : bool; + rZeta : bool; + rDelta : bool; (* true = delta all but rConst; false = delta only on rConst*) + rConst : 'a list +} + +type ('a,'b) red_expr_gen = + | Red of bool + | Hnf + | Simpl + | Cbv of 'b raw_red_flag + | Lazy of 'b raw_red_flag + | Unfold of (int list * 'b) list + | Fold of 'a list + | Pattern of (int list * 'a) list + | ExtraRedExpr of string * 'a list + +type 'a or_metanum = AN of loc * 'a | MetaNum of loc * int + +type 'a may_eval = + | ConstrTerm of 'a + | ConstrEval of ('a, qualid or_metanum) red_expr_gen * 'a + | ConstrContext of (loc * identifier) * 'a + | ConstrTypeOf of 'a diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ad7e02b77..9bdf5822f 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -21,6 +21,7 @@ open Reductionops open Closure open Instantiate open Cbv +open Rawterm exception Elimconst exception Redelimination @@ -743,7 +744,8 @@ let compute = cbv_betadeltaiota (* gives [na:ta]c' such that c converts to ([na:ta]c' a), abstracting only * the specified occurrences. *) -let abstract_scheme env (locc,a,ta) t = +let abstract_scheme env sigma (locc,a) t = + let ta = Retyping.get_type_of env sigma a in let na = named_hd env ta Anonymous in if occur_meta ta then error "cannot find a type for the generalisation"; if occur_meta a then @@ -752,31 +754,60 @@ let abstract_scheme env (locc,a,ta) t = mkLambda (na, ta,subst_term_occ env locc a t) -let pattern_occs loccs_trm_typ env sigma c = - let abstr_trm = List.fold_right (abstract_scheme env) loccs_trm_typ c in - applist(abstr_trm, List.map (fun (_,t,_) -> t) loccs_trm_typ) +let pattern_occs loccs_trm env sigma c = + let abstr_trm = List.fold_right (abstract_scheme env sigma) loccs_trm c in + applist(abstr_trm, List.map snd loccs_trm) (* Generic reduction: reduction functions used in reduction tactics *) -type red_expr = - | Red of bool - | Hnf - | Simpl - | Cbv of Closure.RedFlags.reds - | Lazy of Closure.RedFlags.reds - | Unfold of (int list * evaluable_global_reference) list - | Fold of constr list - | Pattern of (int list * constr * constr) list +type red_expr = (constr, evaluable_global_reference) red_expr_gen + +open RedFlags + +let make_flag_constant = function + | EvalVarRef id -> fVAR id + | EvalConstRef sp -> fCONST sp + +let make_flag f = + let red = no_red in + let red = if f.rBeta then red_add red fBETA else red in + let red = if f.rIota then red_add red fIOTA else red in + let red = if f.rZeta then red_add red fZETA else red in + let red = + if f.rDelta then (* All but rConst *) + let red = red_add red fDELTA in + let red = red_add_transparent red (Conv_oracle.freeze ()) in + List.fold_right + (fun v red -> red_sub red (make_flag_constant v)) + f.rConst red + else (* Only rConst *) + let red = red_add_transparent (red_add red fDELTA) all_opaque in + List.fold_right + (fun v red -> red_add red (make_flag_constant v)) + f.rConst red + in red + +let red_expr_tab = ref Stringmap.empty + +type generic_reduction_function = constr list -> reduction_function + +let declare_red_expr s f = + try + let _ = Stringmap.find s in + error ("There is already a reduction expression of name "^s) + with Not_found -> + red_expr_tab := Stringmap.add s f !red_expr_tab let reduction_of_redexp = function | Red internal -> if internal then internal_red_product else red_product | Hnf -> hnf_constr | Simpl -> nf - | Cbv f -> cbv_norm_flags f - | Lazy f -> clos_norm_flags f + | Cbv f -> cbv_norm_flags (make_flag f) + | Lazy f -> clos_norm_flags (make_flag f) | Unfold ubinds -> unfoldn ubinds | Fold cl -> fold_commands cl | Pattern lp -> pattern_occs lp + | ExtraRedExpr (s,cl) -> Stringmap.find s !red_expr_tab cl (* Used in several tactics. *) diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 4100a31ae..c03c67c09 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -40,7 +40,7 @@ val unfoldn : val fold_commands : constr list -> reduction_function (* Pattern *) -val pattern_occs : (int list * constr * constr) list -> reduction_function +val pattern_occs : (int list * constr) list -> reduction_function (* Rem: Lazy strategies are defined in Reduction *) (* Call by value strategy (uses Closures) *) @@ -60,6 +60,8 @@ val reduce_to_atomic_ind : env -> evar_map -> types -> inductive * types returns [I] and [t'] or fails with a user error *) val reduce_to_quantified_ind : env -> evar_map -> types -> inductive * types +open Rawterm +(* type red_expr = | Red of bool (* raise Redelimination if true otherwise UserError *) | Hnf @@ -69,8 +71,14 @@ type red_expr = | Unfold of (int list * evaluable_global_reference) list | Fold of constr list | Pattern of (int list * constr * constr) list +*) +type red_expr = (constr, evaluable_global_reference) red_expr_gen -val reduction_of_redexp : red_expr -> reduction_function +val reduction_of_redexp : red_expr -> reduction_function + +type generic_reduction_function = constr list -> reduction_function + +val declare_red_expr : string -> generic_reduction_function -> unit (* Opaque and Transparent commands. *) val set_opaque_const : section_path -> unit diff --git a/proofs/clenv.ml b/proofs/clenv.ml index ce56a980d..44c5aafb0 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -19,11 +19,14 @@ open Instantiate open Environ open Evd open Proof_type +open Refiner open Proof_trees open Logic open Reductionops open Tacmach open Evar_refiner +open Rawterm +open Tacexpr (* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms, gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *) @@ -48,9 +51,9 @@ let abstract_list_all env sigma typ c l = raise (RefinerError (CannotGeneralize typ)) (* Generator of metavariables *) -let meta_ctr=ref 0;; +let meta_ctr = ref 0;; -let new_meta ()=incr meta_ctr;!meta_ctr;; +let new_meta () = incr meta_ctr;!meta_ctr;; (* replaces a mapping of existentials into a mapping of metas. Problem if an evar appears in the type of another one (pops anomaly) *) @@ -964,23 +967,22 @@ let clenv_match_args s clause = | [] -> clause | (b,c)::t -> let k = - (match b with - | Dep s -> - if List.mem_assoc b t then - errorlabstrm "clenv_match_args" - (str "The variable " ++ pr_id s ++ - str " occurs more than once in binding") - else - clenv_lookup_name clause s - | NoDep n -> - if List.mem_assoc b t then errorlabstrm "clenv_match_args" - (str "The position " ++ int n ++ - str " occurs more than once in binding"); - (try - List.nth mvs (n-1) - with (Failure _|Invalid_argument _) -> - errorlabstrm "clenv_match_args" (str "No such binder")) - | Com -> anomaly "No free term in clenv_match_args") + match b with + | NamedHyp s -> + if List.mem_assoc b t then + errorlabstrm "clenv_match_args" + (str "The variable " ++ pr_id s ++ + str " occurs more than once in binding") + else + clenv_lookup_name clause s + | AnonHyp n -> + if List.mem_assoc b t then errorlabstrm "clenv_match_args" + (str "The position " ++ int n ++ + str " occurs more than once in binding"); + try + List.nth mvs (n-1) + with (Failure _|Invalid_argument _) -> + errorlabstrm "clenv_match_args" (str "No such binder") in let k_typ = w_hnf_constr clause.hook (clenv_instance_type clause k) and c_typ = w_hnf_constr clause.hook (w_type_of clause.hook c) in @@ -1036,21 +1038,15 @@ let e_res_pf kONT clenv gls = (* Clausal environment for an application *) -let collect_com lbind = - map_succeed (function (Com,c)->c | _ -> failwith "Com") lbind - -let make_clenv_binding_gen n wc (c,t) lbind = - let largs = collect_com lbind in - let lcomargs = List.length largs in - if lcomargs = List.length lbind then - let clause = mk_clenv_from_n wc n (c,t) in - clenv_constrain_dep_args (n <> None) clause largs - else if lcomargs = 0 then - let clause = mk_clenv_rename_from_n wc n (c,t) in - clenv_match_args lbind clause - else - errorlabstrm "make_clenv_bindings" - (str "Cannot mix bindings and free associations") +let make_clenv_binding_gen n wc (c,t) = function + | ImplicitBindings largs -> + let clause = mk_clenv_from_n wc n (c,t) in + clenv_constrain_dep_args (n <> None) clause largs + | ExplicitBindings lbind -> + let clause = mk_clenv_rename_from_n wc n (c,t) in + clenv_match_args lbind clause + | NoBindings -> + mk_clenv_from_n wc n (c,t) let make_clenv_binding_apply wc n = make_clenv_binding_gen (Some n) wc let make_clenv_binding = make_clenv_binding_gen None diff --git a/proofs/clenv.mli b/proofs/clenv.mli index 203353876..54334c45a 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -93,15 +93,17 @@ val clenv_constrain_missing_args : (* Used in user contrib Lannion *) val clenv_constrain_dep_args : constr list -> wc clausenv -> wc clausenv *) val clenv_lookup_name : 'a clausenv -> identifier -> int -val clenv_match_args : (bindOcc * constr) list -> wc clausenv -> wc clausenv +val clenv_match_args : constr Rawterm.explicit_substitution -> wc clausenv -> wc clausenv val e_res_pf : (wc -> tactic) -> wc clausenv -> tactic val clenv_type_of : wc clausenv -> constr -> constr val clenv_unique_resolver : bool -> wc clausenv -> goal sigma -> wc clausenv val make_clenv_binding_apply : - wc -> int -> constr * types -> constr substitution -> wc clausenv + named_context sigma -> int -> constr * constr -> + types Rawterm.substitution -> named_context sigma clausenv val make_clenv_binding : - wc -> constr * types -> constr substitution -> wc clausenv + named_context sigma -> constr * constr -> + types Rawterm.substitution -> named_context sigma clausenv (* Exported for program.ml only *) val clenv_add_sign : @@ -110,10 +112,6 @@ val unify_to_subterm : wc clausenv -> constr * constr -> wc clausenv * constr val unify_to_subterm_list : bool -> wc clausenv -> constr list -> constr -> wc clausenv * constr list -(* -val clenv_constrain_dep_args_of : - int -> constr list -> wc clausenv -> wc clausenv -*) val clenv_typed_unify : Reductionops.conv_pb -> constr -> constr -> wc clausenv -> wc clausenv diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 521cd9b4d..91e13aeed 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -23,6 +23,7 @@ open Proof_trees open Proof_type open Logic open Refiner +open Tacexpr type wc = named_context sigma (* for a better reading of the following *) @@ -148,12 +149,14 @@ let pfic gls c = let evc = gls.sigma in Astterm.interp_constr evc (Global.env_of_context gls.it.evar_hyps) c +(* let instantiate_tac = function | [Integer n; Command com] -> (fun gl -> instantiate n (pfic gl com) gl) | [Integer n; Constr c] -> (fun gl -> instantiate n c gl) | _ -> invalid_arg "Instantiate called with bad arguments" +*) (* vernac command existential *) diff --git a/proofs/evar_refiner.mli b/proofs/evar_refiner.mli index f791f3aba..46b0db62e 100644 --- a/proofs/evar_refiner.mli +++ b/proofs/evar_refiner.mli @@ -53,5 +53,7 @@ val w_const_value : wc -> constant -> constr val w_defined_evar : wc -> existential_key -> bool val instantiate : evar -> constr -> tactic +(* val instantiate_tac : tactic_arg list -> tactic +*) val instantiate_pf_com : int -> Coqast.t -> pftreestate -> pftreestate diff --git a/proofs/logic.ml b/proofs/logic.ml index eb0b94671..2391ece78 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -43,7 +43,6 @@ type refiner_error = (* Errors raised by the tactics *) | CannotUnify of constr * constr | CannotGeneralize of constr - | BadTacticArgs of string * tactic_arg list | IntroNeedsProduct | DoesNotOccurIn of constr * identifier diff --git a/proofs/logic.mli b/proofs/logic.mli index ba1ca7031..dfe40228f 100644 --- a/proofs/logic.mli +++ b/proofs/logic.mli @@ -55,7 +55,6 @@ type refiner_error = (*i Errors raised by the tactics i*) | CannotUnify of constr * constr | CannotGeneralize of constr - | BadTacticArgs of string * tactic_arg list | IntroNeedsProduct | DoesNotOccurIn of constr * identifier diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 035e7af58..410e80093 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -34,7 +34,8 @@ open Safe_typing type proof_topstate = { top_hyps : named_context * named_context; top_goal : goal; - top_strength : strength } + top_strength : bool * Nametab.strength; + top_hook : declaration_hook } let proof_edits = (Edit.empty() : (identifier,pftreestate,proof_topstate) Edit.t) @@ -105,7 +106,7 @@ let get_current_proof_name () = | Some(na,_,_) -> na let add_proof (na,pfs,ts) = - Edit.create proof_edits (na,pfs,ts,Some !undo_limit) + Edit.create proof_edits (na,pfs,ts,Some (!undo_limit+1)) let delete_proof na = try @@ -149,14 +150,16 @@ let subtree_solved () = (*********************************************************************) (* Undo functions *) (*********************************************************************) - -let set_undo n = - if n>=0 then - undo_limit := n+1 - else - error "Cannot set a negative undo limit" - -let reset_undo () = set_undo undo_default + +let set_undo = function + | None -> undo_limit := undo_default + | Some n -> + if n>=0 then + undo_limit := n + else + error "Cannot set a negative undo limit" + +let get_undo () = Some !undo_limit let undo n = try @@ -182,7 +185,7 @@ let cook_proof () = ({ const_entry_body = pfterm; const_entry_type = Some concl; const_entry_opaque = true }, - strength)) + strength, ts.top_hook)) (*********************************************************************) (* Abort functions *) @@ -204,12 +207,13 @@ let delete_all_proofs = init_proofs (* Modifying the current prooftree *) (*********************************************************************) -let start_proof na str sign concl = +let start_proof na str sign concl hook = let top_goal = mk_goal sign concl in let ts = { top_hyps = (sign,empty_named_context); top_goal = top_goal; - top_strength = str } + top_strength = str; + top_hook = hook} in start(na,ts); set_current_proof na diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 462134fce..bb828e287 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -14,7 +14,7 @@ open Names open Term open Sign open Environ -open Declare +open Nametab open Proof_type open Tacmach (*i*) @@ -56,18 +56,17 @@ val delete_all_proofs : unit -> unit val undo : int -> unit -(* [set_undo n] sets the size of the ``undo'' stack *) +(* [set_undo (Some n)] sets the size of the ``undo'' stack; [set_undo None] + sets the size to the default value (12) *) -val set_undo : int -> unit - -(* [reset_undo n] sets the size of the ``undo'' stack to the default - value (12) *) - -val reset_undo : unit -> unit +val set_undo : int option -> unit +val get_undo : unit -> int option (*s [start_proof s str env t] starts a proof of name [s] and conclusion [t] *) -val start_proof : identifier -> strength -> named_context -> constr -> unit +val start_proof : + identifier -> bool * strength -> named_context -> constr + -> declaration_hook -> unit (* [restart_proof ()] restarts the current focused proof from the beginning or fails if no proof is focused *) @@ -93,7 +92,9 @@ val suspend_proof : unit -> unit into a constant with its name and strength; it fails if there is no current proof of if it is not completed *) -val cook_proof : unit -> identifier * (Safe_typing.constant_entry * strength) +val cook_proof : unit -> + identifier * + (Safe_typing.constant_entry * (bool * strength) * declaration_hook) (*s [get_pftreestate ()] returns the current focused pending proof or raises [UserError "no focused proof"] *) diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 115367cd7..7777f3463 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -24,6 +24,7 @@ open Proof_trees open Proof_type open Logic open Refiner +open Tacexpr let re_sig it gc = { it = it; sigma = gc } @@ -164,40 +165,6 @@ let top_of_tree = top_of_tree let frontier = frontier let change_constraints_pftreestate = change_constraints_pftreestate -(*************************************************) -(* Tacticals re-exported from the Refiner module.*) -(*************************************************) - -(* A special exception for levels for the Fail tactic *) -exception FailError = Refiner.FailError - -let tclIDTAC = tclIDTAC -let tclORELSE = tclORELSE -let tclTHEN = tclTHEN -let tclTHENLIST = tclTHENLIST -let tclTHEN_i = tclTHEN_i -let tclTHENL = tclTHENL -let tclTHENS = tclTHENS -let tclTHENSi = tclTHENSi -let tclTHENST = tclTHENST -let tclTHENSI = tclTHENSI -let tclREPEAT = tclREPEAT -let tclREPEAT_MAIN = tclREPEAT_MAIN -let tclFIRST = tclFIRST -let tclSOLVE = tclSOLVE -let tclTRY = tclTRY -let tclTHENTRY = tclTHENTRY -let tclCOMPLETE = tclCOMPLETE -let tclAT_LEAST_ONCE = tclAT_LEAST_ONCE -let tclFAIL = tclFAIL -let tclDO = tclDO -let tclPROGRESS = tclPROGRESS -let tclWEAK_PROGRESS = tclWEAK_PROGRESS -let tclNOTSAMEGOAL = tclNOTSAMEGOAL -let tclINFO = tclINFO - -let unTAC = unTAC - (********************************************) (* Definition of the most primitive tactics *) @@ -258,11 +225,13 @@ let rename_bound_var_goal gls = (* The interpreter of defined tactics *) (***************************************) +(* let vernac_tactic = vernac_tactic let add_tactic = Refiner.add_tactic let overwriting_tactic = Refiner.overwriting_add_tactic +*) (* Some combinators for parsing tactic arguments. @@ -302,26 +271,38 @@ let tactic_com_list = let translate = pf_interp_constr x in tac (List.map translate tl) x +open Rawterm + let tactic_bind_list = fun tac tl x -> let translate = pf_interp_constr x in - tac (List.map (fun (b,c)->(b,translate c)) tl) x + let tl = + match tl with + | NoBindings -> NoBindings + | ImplicitBindings l -> ImplicitBindings (List.map translate l) + | ExplicitBindings l -> + ExplicitBindings (List.map (fun (b,c)->(b,translate c)) l) + in tac tl x + +let translate_bindings gl = function + | NoBindings -> NoBindings + | ImplicitBindings l -> ImplicitBindings (List.map (pf_interp_constr gl) l) + | ExplicitBindings l -> + ExplicitBindings (List.map (fun (b,c)->(b,pf_interp_constr gl c)) l) let tactic_com_bind_list = - fun tac (c,tl) x -> - let translate = pf_interp_constr x in - tac (translate c,List.map (fun (b,c')->(b,translate c')) tl) x + fun tac (c,tl) gl -> + let translate = pf_interp_constr gl in + tac (translate c,translate_bindings gl tl) gl let tactic_com_bind_list_list = fun tac args gl -> - let translate (c,tl) = - (pf_interp_constr gl c, - List.map (fun (b,c')->(b,pf_interp_constr gl c')) tl) + let translate (c,tl) = (pf_interp_constr gl c, translate_bindings gl tl) in tac (List.map translate args) gl (* Some useful combinators for hiding tactic implementations *) - +(* type 'a hide_combinator = string -> ('a -> tactic) -> ('a -> tactic) let hide_atomic_tactic s tac = @@ -331,45 +312,48 @@ let hide_atomic_tactic s tac = let overwrite_hidden_atomic_tactic s tac = overwriting_tactic s (function [] -> tac | _ -> assert false); vernac_tactic(s,[]) - - +*) +(* let hide_constr_comarg_tactic s tac = let tacfun = function | [Constr c] -> tac c - | [Command com] -> tactic_com tac com - | _ -> anomaly "hide_constr_comarg_tactic : neither CONSTR nor COMMAND" +(* | [Command com] -> tactic_com tac com*) + | _ -> anomaly "hide_constr_comarg_tactic : neither CONSTR nor CONSTR" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[Constr c]), - fun com -> vernac_tactic(s,[Command com])) - + (* fun com -> vernac_tactic(s,[Command com]) *) fun _ -> failwith "Command unsupported") +*) +(* let overwrite_hidden_constr_comarg_tactic s tac = let tacfun = function | [Constr c] -> tac c - | [Command com] -> - (fun gls -> tac (pf_interp_constr gls com) gls) +(* | [Command com] -> + (fun gls -> tac (pf_interp_constr gls com) gls)*) | _ -> anomaly - "overwrite_hidden_constr_comarg_tactic : neither CONSTR nor COMMAND" + "overwrite_hidden_constr_comarg_tactic : neither CONSTR nor CONSTR" in overwriting_tactic s tacfun; (fun c -> vernac_tactic(s,[(Constr c)]), - fun c -> vernac_tactic(s,[(Command c)])) - + (*fun c -> vernac_tactic(s,[(Command c)])*) fun _ -> failwith "Command unsupported") +*) +(* let hide_constr_tactic s tac = let tacfun = function | [Constr c] -> tac c - | [Command com] -> tactic_com tac com - | _ -> anomaly "hide_constr_tactic : neither CONSTR nor COMMAND" +(* | [Command com] -> tactic_com tac com*) + | _ -> anomaly "hide_constr_tactic : neither CONSTR nor CONSTR" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[(Constr c)])) - +*) +(* let hide_openconstr_tactic s tac = let tacfun = function | [OpenConstr c] -> tac c - | [Command com] -> tactic_opencom tac com - | _ -> anomaly "hide_openconstr_tactic : neither OPENCONSTR nor COMMAND" +(* | [Command com] -> tactic_opencom tac com*) + | _ -> anomaly "hide_openconstr_tactic : neither OPENCONSTR nor CONSTR" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[(OpenConstr c)])) @@ -393,40 +377,44 @@ let hide_identl_tactic s tac = let tacfun = (function [Clause idl] -> tac idl | _ -> assert false) in add_tactic s tacfun; fun idl -> vernac_tactic(s,[Clause idl]) - +*) +(* let hide_constrl_tactic s tac = let tacfun = function - | ((Command com)::_) as al -> +(* | ((Command com)::_) as al -> tactic_com_list tac - (List.map (function (Command com) -> com | _ -> assert false) al) + (List.map (function (Command com) -> com | _ -> assert false) al)*) | ((Constr com)::_) as al -> tac (List.map (function (Constr c) -> c | _ -> assert false) al) - | _ -> anomaly "hide_constrl_tactic : neither CONSTR nor COMMAND" + | _ -> anomaly "hide_constrl_tactic : neither CONSTR nor CONSTR" in add_tactic s tacfun; fun ids -> vernac_tactic(s,(List.map (fun id -> Constr id) ids)) - +*) +(* let hide_bindl_tactic s tac = let tacfun = function - | [Bindings al] -> tactic_bind_list tac al +(* | [Bindings al] -> tactic_bind_list tac al*) | [Cbindings al] -> tac al | _ -> anomaly "hide_bindl_tactic : neither BINDINGS nor CBINDINGS" in add_tactic s tacfun; fun bindl -> vernac_tactic(s,[Cbindings bindl]) - +*) +(* let hide_cbindl_tactic s tac = let tacfun = function - | [Command com; Bindings al] -> tactic_com_bind_list tac (com,al) +(* | [Command com; Bindings al] -> tactic_com_bind_list tac (com,al)*) | [Constr c; Cbindings al] -> tac (c,al) - | _ -> anomaly "hide_cbindl_tactic : neither CONSTR nor COMMAND" + | _ -> anomaly "hide_cbindl_tactic : neither CONSTR nor CONSTR" in add_tactic s tacfun; fun (c,bindl) -> vernac_tactic(s,[Constr c; Cbindings bindl]) - +*) +(* let hide_cbindll_tactic s tac = let rec getcombinds = function - | ((Command com)::(Bindings al)::l) -> (com,al)::(getcombinds l) +(* | ((Command com)::(Bindings al)::l) -> (com,al)::(getcombinds l)*) | [] -> [] | _ -> anomaly "hide_cbindll_tactic : not the expected form" in @@ -440,19 +428,21 @@ let hide_cbindll_tactic s tac = | [] -> [] in let tacfun = function - | ((Command com)::_) as args -> - tactic_com_bind_list_list tac (getcombinds args) +(* | ((Command com)::_) as args -> + tactic_com_bind_list_list tac (getcombinds args)*) | ((Constr com)::_) as args -> tac (getconstrbinds args) - | _ -> anomaly "hide_cbindll_tactic : neither CONSTR nor COMMAND" + | _ -> anomaly "hide_cbindll_tactic : neither CONSTR nor CONSTR" in add_tactic s tacfun; fun l -> vernac_tactic(s,putconstrbinds l) - +*) (* Pretty-printers *) open Pp open Printer +open Tacexpr +open Rawterm let pr_com sigma goal com = prterm (rename_bound_var (Global.env()) @@ -460,15 +450,17 @@ let pr_com sigma goal com = (Astterm.interp_constr sigma (Evarutil.evar_env goal) com)) let pr_one_binding sigma goal = function - | (Dep id,com) -> pr_id id ++ str ":=" ++ pr_com sigma goal com - | (NoDep n,com) -> int n ++ str ":=" ++ pr_com sigma goal com - | (Com,com) -> pr_com sigma goal com + | (NamedHyp id,com) -> pr_id id ++ str ":=" ++ pr_com sigma goal com + | (AnonHyp n,com) -> int n ++ str ":=" ++ pr_com sigma goal com let pr_bindings sigma goal lb = let prf = pr_one_binding sigma goal in match lb with - | [] -> prlist_with_sep pr_spc prf lb - | _ -> str "with" ++ spc () ++ prlist_with_sep pr_spc prf lb + | ImplicitBindings l -> + str "with" ++ spc () ++ prlist_with_sep pr_spc (pr_com sigma goal) l + | ExplicitBindings l -> + str "with" ++ spc () ++ prlist_with_sep pr_spc prf l + | NoBindings -> mt () let rec pr_list f = function | [] -> mt () @@ -480,5 +472,4 @@ let pr_gls gls = let pr_glls glls = hov 0 (pr_decls (sig_sig glls) ++ fnl () ++ prlist_with_sep pr_fnl pr_seq (sig_it glls)) - -let pr_tactic = Refiner.pr_tactic + diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index ee09b7ada..49a2db419 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -19,6 +19,8 @@ open Proof_trees open Proof_type open Refiner open Tacred +open Tacexpr +open Rawterm (*i*) (* Operations for handling terms under a local typing context. *) @@ -115,39 +117,9 @@ val top_of_tree : pftreestate -> pftreestate val change_constraints_pftreestate : evar_map -> pftreestate -> pftreestate -(*s Tacticals re-exported from the Refiner module. *) - -(* A special exception for levels for the Fail tactic *) -exception FailError of int - -val tclIDTAC : tactic -val tclORELSE : tactic -> tactic -> tactic -val tclTHEN : tactic -> tactic -> tactic -val tclTHENLIST : tactic list -> tactic -val tclTHEN_i : tactic -> (int -> tactic) -> tactic -val tclTHENL : tactic -> tactic -> tactic -val tclTHENS : tactic -> tactic list -> tactic -val tclTHENSi : tactic -> tactic list -> (int -> tactic) -> tactic -val tclTHENST : tactic -> tactic list -> tactic -> tactic -val tclTHENSI : tactic -> tactic list -> tactic -val tclREPEAT : tactic -> tactic -val tclREPEAT_MAIN : tactic -> tactic -val tclFIRST : tactic list -> tactic -val tclSOLVE : tactic list -> tactic -val tclTRY : tactic -> tactic -val tclTHENTRY : tactic -> tactic -> tactic -val tclCOMPLETE : tactic -> tactic -val tclAT_LEAST_ONCE : tactic -> tactic -val tclFAIL : int -> tactic -val tclDO : int -> tactic -> tactic -val tclPROGRESS : tactic -> tactic -val tclWEAK_PROGRESS : tactic -> tactic -val tclNOTSAMEGOAL : tactic -> tactic -val tclINFO : tactic -> tactic - -val unTAC : tactic -> goal sigma -> proof_tree sigma -val vernac_tactic : tactic_expression -> tactic - +(* +val vernac_tactic : string * tactic_arg list -> tactic +*) (*s The most primitive tactics. *) val refiner : rule -> tactic @@ -185,9 +157,10 @@ val tclIDTAC_list : tactic_list (*s Tactic Registration. *) +(* val add_tactic : string -> (tactic_arg list -> tactic) -> unit val overwriting_tactic : string -> (tactic_arg list -> tactic) -> unit - +*) (*s Transformation of tactic arguments. *) @@ -197,42 +170,16 @@ val tactic_com : (constr,Coqast.t) parse_combinator val tactic_com_sort : (constr,Coqast.t) parse_combinator val tactic_com_list : (constr list, Coqast.t list) parse_combinator -val tactic_bind_list : ((bindOcc * constr) list, - (bindOcc * Coqast.t) list) parse_combinator +val tactic_bind_list : + (constr substitution, Coqast.t substitution) parse_combinator val tactic_com_bind_list : - (constr * (bindOcc * constr) list, - Coqast.t * (bindOcc * Coqast.t) list) parse_combinator + (constr * constr substitution, + Coqast.t * Coqast.t substitution) parse_combinator val tactic_com_bind_list_list : - ((constr * (bindOcc * constr) list) list, - (Coqast.t * (bindOcc * Coqast.t) list) list) parse_combinator - - -(*s Hiding the implementation of tactics. *) - -val hide_tactic : - string -> (tactic_arg list -> tactic) -> (tactic_arg list -> tactic) - -val overwrite_hidden_tactic : - string -> (tactic_arg list -> tactic) -> (tactic_arg list -> tactic) - - -type 'a hide_combinator = string -> ('a -> tactic) -> ('a -> tactic) - -val hide_atomic_tactic : string -> tactic -> tactic -val hide_constr_tactic : constr hide_combinator -val hide_openconstr_tactic : Pretyping.open_constr hide_combinator -val hide_constrl_tactic : (constr list) hide_combinator -val hide_numarg_tactic : int hide_combinator -val hide_ident_tactic : identifier hide_combinator -val hide_identl_tactic : hyp_location list hide_combinator -val hide_string_tactic : string hide_combinator -val hide_bindl_tactic : ((bindOcc * constr) list) hide_combinator -val hide_cbindl_tactic : (constr * (bindOcc * constr) list) hide_combinator -val hide_cbindll_tactic : - ((constr * (bindOcc * constr) list) list) hide_combinator - + ((constr * constr substitution) list, + (Coqast.t * Coqast.t substitution) list) parse_combinator (*s Pretty-printing functions. *) @@ -243,4 +190,3 @@ open Pp val pr_com : evar_map -> goal -> Coqast.t -> std_ppcmds val pr_gls : goal sigma -> std_ppcmds val pr_glls : goal list sigma -> std_ppcmds -val pr_tactic : tactic_expression -> std_ppcmds diff --git a/proofs/tactic_debug.ml b/proofs/tactic_debug.ml index 61ccbcac7..43162f05d 100644 --- a/proofs/tactic_debug.ml +++ b/proofs/tactic_debug.ml @@ -7,6 +7,7 @@ (***********************************************************************) open Ast open Pp +open Pptactic open Printer (* This module intends to be a beginning of debugger for tactic expressions. @@ -33,7 +34,7 @@ let help () = (* Prints the state and waits for an instruction *) let debug_prompt goalopt tac_ast = db_pr_goal goalopt; - msg (str "Going to execute:" ++ fnl () ++ (gentacpr tac_ast) ++ fnl ()); + msg (str "Going to execute:" ++ fnl () ++ (pr_raw_tactic tac_ast) ++ fnl ()); (* str "Commands: =Continue, s=Skip, x=Exit" >];*) (* mSG (str "Going to execute:" ++ fnl () ++ (gentacpr tac_ast) ++ fnl () ++ fnl () ++ str "----=Continue----s=Skip----x=Exit----");*) @@ -62,7 +63,7 @@ let db_constr debug env c = (* Prints a matched hypothesis *) let db_matched_hyp debug env (id,c) = if debug = DebugOn then - msgnl (str "Matched hypothesis --> " ++ str (id^": ") ++ + msgnl (str "Matched hypothesis --> " ++ str (Names.string_of_id id^": ") ++ prterm_env env c) (* Prints the matched conclusion *) diff --git a/proofs/tactic_debug.mli b/proofs/tactic_debug.mli index 4a6aecfde..b05a33fdf 100644 --- a/proofs/tactic_debug.mli +++ b/proofs/tactic_debug.mli @@ -9,6 +9,7 @@ (*i $Id$ i*) open Proof_type +open Names open Term (* This module intends to be a beginning of debugger for tactic expressions. @@ -22,13 +23,13 @@ type debug_info = | Exit (* Prints the state and waits *) -val debug_prompt : goal sigma option -> Coqast.t -> debug_info +val debug_prompt : goal sigma option -> Tacexpr.raw_tactic_expr -> debug_info (* Prints a constr *) val db_constr : debug_info -> Environ.env -> constr -> unit (* Prints a matched hypothesis *) -val db_matched_hyp : debug_info -> Environ.env -> string * constr -> unit +val db_matched_hyp : debug_info -> Environ.env -> identifier * constr -> unit (* Prints the matched conclusion *) val db_matched_concl : debug_info -> Environ.env -> constr -> unit diff --git a/tactics/auto.ml b/tactics/auto.ml index b9d09f717..256914d4c 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -32,10 +32,10 @@ open Clenv open Hiddentac open Libobject open Library -open Vernacinterp open Printer open Nametab open Declarations +open Tacexpr (****************************************************************************) (* The Type of Constructions Autotactic Hints *) @@ -47,7 +47,7 @@ type auto_tactic = | Give_exact of constr | Res_pf_THEN_trivial_fail of constr * unit clausenv (* Hint Immediate *) | Unfold_nth of global_reference (* Hint Unfold *) - | Extern of Coqast.t (* Hint Extern *) + | Extern of raw_tactic_expr (* Hint Extern *) type pri_auto_tactic = { hname : identifier; (* name of the hint *) @@ -136,6 +136,8 @@ type frozen_hint_db_table = Hint_db.t Stringmap.t type hint_db_table = Hint_db.t Stringmap.t ref +type hint_db_name = string + let searchtable = (ref Stringmap.empty : hint_db_table) let searchtable_map name = @@ -300,7 +302,10 @@ let make_extern name pri pat tacast = let add_extern name pri (patmetas,pat) tacast dbname = (* We check that all metas that appear in tacast have at least one occurence in the left pattern pat *) - let tacmetas = Coqast.collect_metas tacast in +(* TODO + let tacmetas = Coqast.collect_metas tacast in +*) + let tacmetas = [] in match (list_subtract tacmetas patmetas) with | i::_ -> errorlabstrm "add_extern" @@ -328,150 +333,8 @@ let add_trivials l dbnames = Lib.add_anonymous_leaf (inAutoHint(dbname, List.map make_trivial l))) dbnames -let _ = - vinterp_add - "HintUnfold" - (function - | [ VARG_IDENTIFIER hintname; VARG_VARGLIST l; VARG_QUALID qid] -> - let dbnames = if l = [] then ["core"] else - List.map - (function VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintUnfold") l in - fun () -> - let ref = Nametab.global dummy_loc qid in - add_unfolds [(hintname, ref)] dbnames - | _-> bad_vernac_args "HintUnfold") - -let _ = - vinterp_add - "HintResolve" - (function - | [VARG_IDENTIFIER hintname; VARG_VARGLIST l; VARG_CONSTR c] -> - let env = Global.env() and sigma = Evd.empty in - let c1 = Astterm.interp_constr sigma env c in - let dbnames = if l = [] then ["core"] else - List.map (function VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintResolve") l in - fun () -> add_resolves env sigma [hintname, c1] dbnames - | _-> bad_vernac_args "HintResolve" ) - -let _ = - vinterp_add - "HintImmediate" - (function - | [VARG_IDENTIFIER hintname; VARG_VARGLIST l; VARG_CONSTR c] -> - let c1 = Astterm.interp_constr Evd.empty (Global.env()) c in - let dbnames = if l = [] then ["core"] else - List.map (function VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintImmediate") l in - fun () -> add_trivials [hintname, c1] dbnames - | _ -> bad_vernac_args "HintImmediate") - - -let _ = - vinterp_add - "HintConstructors" - (function - | [VARG_IDENTIFIER idr; VARG_VARGLIST l; VARG_QUALID qid] -> - begin - try - let env = Global.env() and sigma = Evd.empty in - let isp = destInd (Declare.global_qualified_reference qid) in - let conspaths = - let (mib,mip) = Global.lookup_inductive isp in - mip.mind_consnames in - let lcons = - array_map_to_list - (fun id -> - let sp = make_path (dirpath (fst isp)) id in - let c = Declare.global_absolute_reference sp in - (id, c)) - conspaths in - let dbnames = if l = [] then ["core"] else - List.map (function VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintConstructors") l in - fun () -> add_resolves env sigma lcons dbnames - with Invalid_argument("mind_specif_of_mind") -> - error ((Nametab.string_of_qualid qid) ^ " is not an inductive type") - end - | _ -> bad_vernac_args "HintConstructors") - -let _ = - vinterp_add - "HintExtern" - (function - | [VARG_IDENTIFIER hintname; VARG_VARGLIST l; - VARG_NUMBER pri; VARG_CONSTR patcom; VARG_TACTIC tacexp] -> - let pat = - Astterm.interp_constrpattern Evd.empty (Global.env()) patcom in - let dbnames = if l = [] then ["core"] else - List.map (function VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintConstructors") l in - fun () -> add_externs hintname pri pat tacexp dbnames - | _ -> bad_vernac_args "HintExtern") - -let _ = - vinterp_add - "HintsResolve" - (function - | (VARG_VARGLIST l)::lh -> - let env = Global.env() and sigma = Evd.empty in - let lhints = - List.map - (function - | VARG_QUALID qid -> - let ref = Nametab.global dummy_loc qid in - let env = Global.env() in - let c = Declare.constr_of_reference ref in - let _,i = Nametab.repr_qualid qid in - (i, c) - | _-> bad_vernac_args "HintsResolve") lh in - let dbnames = if l = [] then ["core"] else - List.map (function VARG_IDENTIFIER i -> string_of_id i - | _-> bad_vernac_args "HintsResolve") l in - fun () -> add_resolves env sigma lhints dbnames - | _-> bad_vernac_args "HintsResolve") - -let _ = - vinterp_add - "HintsUnfold" - (function - | (VARG_VARGLIST l)::lh -> - let lhints = - List.map (function - | VARG_QUALID qid -> - let _,n = Nametab.repr_qualid qid in - (n, Nametab.global dummy_loc qid) - | _ -> bad_vernac_args "HintsUnfold") lh in - let dbnames = if l = [] then ["core"] else - List.map (function - | VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintsUnfold") l in - fun () -> add_unfolds lhints dbnames - | _ -> bad_vernac_args "HintsUnfold") - -let _ = - vinterp_add - "HintsImmediate" - (function - | (VARG_VARGLIST l)::lh -> - let lhints = - List.map - (function - | VARG_QUALID qid -> - let _,n = Nametab.repr_qualid qid in - let ref = Nametab.locate qid in - let env = Global.env () in - let c = Declare.constr_of_reference ref in - (n, c) - | _ -> bad_vernac_args "HintsImmediate") lh in - let dbnames = if l = [] then ["core"] else - List.map (function - | VARG_IDENTIFIER i -> string_of_id i - | _ -> bad_vernac_args "HintsImmediate") l in - fun () -> add_trivials lhints dbnames - | _-> bad_vernac_args "HintsImmediate") - +open Vernacexpr + (**************************************************************************) (* Functions for printing the hints *) (**************************************************************************) @@ -483,7 +346,7 @@ let fmt_autotactic = function | Res_pf_THEN_trivial_fail (c,clenv) -> (str"Apply " ++ prterm c ++ str" ; Trivial") | Unfold_nth c -> (str"Unfold " ++ pr_global c) - | Extern coqast -> (str "Extern " ++ gentacpr coqast) + | Extern coqast -> (str "Extern " ++ Pptactic.pr_raw_tactic coqast) let fmt_hint v = (fmt_autotactic v.code ++ str"(" ++ int v.pri ++ str")" ++ spc ()) @@ -514,7 +377,7 @@ let fmt_hint_list_for_head c = let fmt_hint_ref ref = fmt_hint_list_for_head (label_of_ref ref) (* Print all hints associated to head id in any database *) -let print_hint_qid qid = ppnl(fmt_hint_ref (Nametab.global dummy_loc qid)) +let print_hint_ref ref = ppnl(fmt_hint_ref ref) let fmt_hint_term cl = try @@ -574,30 +437,44 @@ let print_searchtable () = print_hint_db db) !searchtable -let _ = - vinterp_add "PrintHint" - (function - | [] -> fun () -> print_searchtable() - | _ -> bad_vernac_args "PrintHint") - -let _ = - vinterp_add "PrintHintDb" - (function - | [(VARG_IDENTIFIER id)] -> - fun () -> print_hint_db_by_name (string_of_id id) - | _ -> bad_vernac_args "PrintHintDb") - -let _ = - vinterp_add "PrintHintGoal" - (function - | [] -> fun () -> print_applicable_hint() - | _ -> bad_vernac_args "PrintHintGoal") - -let _ = - vinterp_add "PrintHintId" - (function - | [(VARG_QUALID qid)] -> fun () -> print_hint_qid qid - | _ -> bad_vernac_args "PrintHintId") +let add_hints dbnames h = + let dbnames = if dbnames = [] then ["core"] else dbnames in match h with + | HintsResolve lhints -> + let env = Global.env() and sigma = Evd.empty in + let f (n,c) = + let c = Astterm.interp_constr sigma env c in + let n = match n with + | None -> basename (sp_of_global env (Declare.reference_of_constr c)) + | Some n -> n in + (n,c) in + add_resolves env sigma (List.map f lhints) dbnames + | HintsImmediate lhints -> + let env = Global.env() and sigma = Evd.empty in + let f (n,c) = + let c = Astterm.interp_constr sigma env c in + let n = match n with + | None -> basename (sp_of_global env (Declare.reference_of_constr c)) + | Some n -> n in + (n,c) in + add_trivials (List.map f lhints) dbnames + | HintsUnfold lhints -> + let f (n,locqid) = + let r = Nametab.global locqid in + let n = match n with + | None -> basename (sp_of_global (Global.env()) r) + | Some n -> n in + (n,r) in + add_unfolds (List.map f lhints) dbnames + | HintsConstructors (hintname, qid) -> + let env = Global.env() and sigma = Evd.empty in + let isp = global_inductive qid in + let consnames = (snd (Global.lookup_inductive isp)).mind_consnames in + let lcons = list_tabulate (fun i -> mkConstruct (isp,i+1)) (Array.length consnames) in + let lcons = List.map2 (fun id c -> (id,c)) (Array.to_list consnames) lcons in + add_resolves env sigma lcons dbnames + | HintsExtern (hintname, pri, patcom, tacexp) -> + let pat = Astterm.interp_constrpattern Evd.empty (Global.env()) patcom in + add_externs hintname pri pat tacexp dbnames (**************************************************************************) (* Automatic tactics *) @@ -627,6 +504,29 @@ let make_local_hint_db g = in Hint_db.add_list hintlist Hint_db.empty +(* Serait-ce possible de compiler d'abord la tactique puis de faire la + substitution sans passer par bdize dont l'objectif est de prparer un + terme pour l'affichage ? (HH) *) + +(* Si on enlve le dernier argument (gl) conclPattern est calcul une +fois pour toutes : en particulier si Pattern.somatch produit une UserError +Ce qui fait que si la conclusion ne matche pas le pattern, Auto choue, mme +si aprs Intros la conclusion matche le pattern. +*) + +(* conclPattern doit chouer avec error car il est rattraper par tclFIRST *) + +let forward_tac_interp = + ref (fun _ -> failwith "tac_interp is not installed for Auto") + +let set_extern_interp f = forward_tac_interp := f + +let conclPattern concl pat tac gl = + let constr_bindings = + try Pattern.matches pat concl + with PatternMatchingFailure -> error "conclPattern" in + !forward_tac_interp constr_bindings tac gl + (**************************************************************************) (* The Trivial tactic *) (**************************************************************************) @@ -697,16 +597,11 @@ let full_trivial gl = let db_list = List.map (fun x -> searchtable_map x) dbnames in tclTRY (trivial_fail_db db_list (make_local_hint_db gl)) gl -let dyn_trivial = function - | [] -> trivial [] - | [Quoted_string "*"] -> full_trivial - | l -> trivial (List.map - (function - | Identifier id -> (string_of_id id) - | other -> bad_tactic_args "dyn_trivial" [other]) - l) - -let h_trivial = hide_tactic "Trivial" dyn_trivial +let gen_trivial = function + | None -> full_trivial + | Some l -> trivial l + +let h_trivial l = Refiner.abstract_tactic (TacTrivial l) (gen_trivial l) (**************************************************************************) (* The classical Auto tactic *) @@ -807,24 +702,13 @@ let full_auto n gl = let default_full_auto gl = full_auto !default_search_depth gl -let dyn_auto l = match l with - | [] -> auto !default_search_depth [] - | [Integer n] -> auto n [] - | [Quoted_string "*"] -> default_full_auto - | [Integer n; Quoted_string "*"] -> full_auto n - | (Integer n)::l1 -> - auto n (List.map - (function - | Identifier id -> (string_of_id id) - | other -> bad_tactic_args "dyn_auto" [other]) l1) - | _ -> - auto !default_search_depth - (List.map - (function - | Identifier id -> (string_of_id id) - | other -> bad_tactic_args "dyn_auto" [other]) l) +let gen_auto n dbnames = + let n = match n with None -> !default_search_depth | Some n -> n in + match dbnames with + | None -> full_auto n + | Some l -> auto n l -let h_auto = hide_tactic "Auto" dyn_auto +let h_auto n l = Refiner.abstract_tactic (TacAuto (n,l)) (gen_auto n l) (**************************************************************************) (* The "destructing Auto" from Eduardo *) @@ -845,19 +729,13 @@ let dautomatic des_opt n = tclTRY (destruct_auto des_opt n) let default_dauto = dautomatic !default_search_decomp !default_search_depth -let dyn_dauto = function - | [] -> default_dauto - | [Integer n] -> dautomatic !default_search_decomp n - | [Integer n; Integer p] -> dautomatic p n - | _ -> invalid_arg "DAuto: non numeric arguments" - -let dauto = - let gentac = hide_tactic "DAuto" dyn_dauto in - function - | (None, None) -> gentac [] - | (Some n, None) -> gentac [Integer n] - | (Some n, Some p) -> gentac [Integer n; Integer p] - | _ -> assert false +let dauto = function + | None, None -> default_dauto + | Some n, None -> dautomatic !default_search_decomp n + | Some n, Some p -> dautomatic p n + | None, Some p -> dautomatic p !default_search_depth + +let h_dauto (n,p) = Refiner.abstract_tactic (TacDAuto (n,p)) (dauto (n,p)) (***************************************) (*** A new formulation of Auto *********) @@ -888,8 +766,8 @@ let compileAutoArg contac = function (tclTHEN (Tacticals.tryAllClauses (function - | Some id -> Dhyp.dHyp id - | None -> Dhyp.dConcl)) + | Some id -> Dhyp.h_destructHyp false id + | None -> Dhyp.h_destructConcl)) contac) let compileAutoArgList contac = List.map (compileAutoArg contac) @@ -931,33 +809,17 @@ let superauto n to_add argl = let default_superauto g = superauto !default_search_depth [] [] g -let cvt_autoArg = function - | "Destructing" -> [Destructing] - | "UsingTDB" -> [UsingTDB] - | "NoAutoArg" -> [] - | x -> errorlabstrm "cvt_autoArg" - (str "Unexpected argument for Auto!" ++ str x) - -let cvt_autoArgs = - list_join_map - (function - | Quoted_string s -> (cvt_autoArg s) - | _ -> errorlabstrm "cvt_autoArgs" (str "String expected")) - -let interp_to_add gl = function - | Qualid qid -> - let _,id = Nametab.repr_qualid qid in - (next_ident_away id (pf_ids_of_hyps gl), - Declare.constr_of_reference (Nametab.global dummy_loc qid)) - | _ -> errorlabstrm "cvt_autoArgs" (str "Qualid expected") - -let dyn_superauto l g = - match l with - | (Integer n)::a::b::c::to_add -> - superauto n (List.map (interp_to_add g) to_add) (cvt_autoArgs [a;b;c])g - | _::a::b::c::to_add -> - superauto !default_search_depth (List.map (interp_to_add g) to_add) - (cvt_autoArgs [a;b;c]) g - | l -> bad_tactic_args "SuperAuto" l g - -let h_superauto = hide_tactic "SuperAuto" dyn_superauto +let interp_to_add gl locqid = + let r = Nametab.global locqid in + let id = basename (sp_of_global (Global.env()) r) in + (next_ident_away id (pf_ids_of_hyps gl), Declare.constr_of_reference r) + +let gen_superauto nopt l a b gl = + let n = match nopt with Some n -> n | None -> !default_search_depth in + let al = (if a then [Destructing] else [])@(if b then [UsingTDB] else []) in + superauto n (List.map (interp_to_add gl) l) al gl + +let h_superauto no l a b = + Refiner.abstract_tactic (TacSuperAuto (no,l,a,b)) (gen_superauto no l a b) + + diff --git a/tactics/auto.mli b/tactics/auto.mli index 20f770b83..c5266bc58 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -28,7 +28,7 @@ type auto_tactic = | Give_exact of constr | Res_pf_THEN_trivial_fail of constr * unit clausenv (* Hint Immediate *) | Unfold_nth of global_reference (* Hint Unfold *) - | Extern of Coqast.t (* Hint Extern *) + | Extern of Tacexpr.raw_tactic_expr (* Hint Extern *) open Rawterm @@ -59,6 +59,18 @@ type frozen_hint_db_table = Hint_db.t Stringmap.t type hint_db_table = Hint_db.t Stringmap.t ref +type hint_db_name = string + +val add_hints : hint_db_name list -> Vernacexpr.hints -> unit + +val print_searchtable : unit -> unit + +val print_applicable_hint : unit -> unit + +val print_hint_ref : global_reference -> unit + +val print_hint_db_by_name : hint_db_name -> unit + val searchtable : hint_db_table (* [make_exact_entry hint_name (c, ctyp)]. @@ -99,12 +111,15 @@ val make_resolve_hyp : env -> evar_map -> named_declaration -> (constr_label * pri_auto_tactic) list -(* [make_extern name pri pattern tactic_ast] *) +(* [make_extern name pri pattern tactic_expr] *) val make_extern : - identifier -> int -> constr_pattern -> Coqast.t + identifier -> int -> constr_pattern -> Tacexpr.raw_tactic_expr -> constr_label * pri_auto_tactic +val set_extern_interp : + ((int * constr) list -> Tacexpr.raw_tactic_expr -> tactic) -> unit + (* Create a Hint database from the pairs (name, constr). Useful to take the current goal hypotheses as hints *) @@ -117,9 +132,16 @@ val default_search_depth : int ref (* Try unification with the precompiled clause, then use registered Apply *) val unify_resolve : (constr * unit clausenv) -> tactic +(* [ConclPattern concl pat tacast]: + if the term concl matches the pattern pat, (in sense of + [Pattern.somatches], then replace [?1] [?2] metavars in tacast by the + right values to build a tactic *) + +val conclPattern : constr -> constr_pattern -> Tacexpr.raw_tactic_expr -> tactic + (* The Auto tactic *) -val auto : int -> string list -> tactic +val auto : int -> hint_db_name list -> tactic (* auto with default search depth and with the hint database "core" *) val default_auto : tactic @@ -131,14 +153,19 @@ val full_auto : int -> tactic except the "v62" compatibility database *) val default_full_auto : tactic +(* The generic form of auto (second arg [None] means all bases) *) +val gen_auto : int option -> hint_db_name list option -> tactic + (* The hidden version of auto *) -val h_auto : tactic_arg list -> tactic +val h_auto : int option -> hint_db_name list option -> tactic (* Trivial *) -val trivial : string list -> tactic +val trivial : hint_db_name list -> tactic +val gen_trivial : hint_db_name list option -> tactic val full_trivial : tactic -val h_trivial : tactic_arg list -> tactic +val h_trivial : hint_db_name list option -> tactic +val fmt_autotactic : auto_tactic -> Pp.std_ppcmds (*s The following is not yet up to date -- Papageno. *) @@ -147,13 +174,15 @@ val dauto : int option * int option -> tactic val default_search_decomp : int ref val default_dauto : tactic +val h_dauto : int option * int option -> tactic (* SuperAuto *) type autoArguments = | UsingTDB | Destructing -val fmt_autotactic : auto_tactic -> Pp.std_ppcmds - +(* val superauto : int -> (identifier * constr) list -> autoArguments list -> tactic +*) +val h_superauto : int option -> qualid located list -> bool -> bool -> tactic diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 3216a6065..75c7509ac 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -5,6 +5,7 @@ (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) + open Ast open Coqast open Equality @@ -12,12 +13,13 @@ open Hipattern open Names open Pp open Proof_type -open Tacmach +open Tacticals open Tacinterp open Tactics open Term open Util open Vernacinterp +open Tacexpr (* Rewriting rules *) type rew_rule = constr * bool * tactic @@ -37,7 +39,7 @@ let _ = Summary.survive_section = false } (* Rewriting rules before tactic interpretation *) -type raw_rew_rule = constr * bool * t +type raw_rew_rule = constr * bool * raw_tactic_expr (* Applies all the rules of one base *) let one_base tac_main bas = @@ -48,8 +50,8 @@ let one_base tac_main bas = else tclREPEAT_MAIN (tclPROGRESS (List.fold_left (fun tac (csr,dir,tc) -> tclTHEN tac - (if dir then tclREPEAT_MAIN (tclTHENST (rewriteLR csr) [tac_main] tc) - else tclREPEAT_MAIN (tclTHENST (rewriteRL csr) [tac_main] tc))) + (tclREPEAT_MAIN + (tclTHENSFIRSTn (general_rewrite dir csr) [|tac_main|] tc))) tclIDTAC lrul)) (* The AutoRewrite tactic *) @@ -77,39 +79,3 @@ let (in_hintrewrite,out_hintrewrite)= (* To add rewriting rules to a base *) let add_rew_rules base lrul = Lib.add_anonymous_leaf (in_hintrewrite (base,lrul)) - -(* The vernac declaration of HintRewrite *) -let _ = vinterp_add "HintRewrite" - (function - | [VARG_STRING ort;VARG_CONSTRLIST lcom;VARG_IDENTIFIER id;VARG_TACTIC t] - when ort = "LR" || ort = "RL" -> - (fun () -> - let (evc,env) = Command.get_current_context () in - let lcsr = - List.map (function - | Node(loc,"CONSTR",l) -> - let ist = { evc=evc; env=env; lfun=[]; lmatch=[]; - goalopt=None; debug=Tactic_debug.DebugOff } in - constr_of_Constr (interp_tacarg ist (Node(loc,"COMMAND",l))) - | _ -> bad_vernac_args "HintRewrite") lcom in - add_rew_rules (string_of_id id) - (List.map (fun csr -> (csr,ort = "LR",t)) lcsr)) - | _ -> bad_vernac_args "HintRewrite") - -(* To get back the tactic arguments and call AutoRewrite *) -let v_autorewrite = function - | (Tac (t,_))::l -> - let lbas = - List.map (function - | Identifier id -> string_of_id id - | _ -> Tacinterp.bad_tactic_args "AutoRewrite") l in - autorewrite t lbas - | l -> - let lbas = - List.map (function - | Identifier id -> string_of_id id - | _ -> Tacinterp.bad_tactic_args "AutoRewrite") l in - autorewrite tclIDTAC lbas - -(* Declaration of AutoRewrite *) -let _ = hide_tactic "AutoRewrite" v_autorewrite diff --git a/tactics/autorewrite.mli b/tactics/autorewrite.mli index ce9265866..b24ecb18b 100644 --- a/tactics/autorewrite.mli +++ b/tactics/autorewrite.mli @@ -8,10 +8,12 @@ (*i $Id$ i*) +(*i*) open Tacmach +(*i*) (* Rewriting rules before tactic interpretation *) -type raw_rew_rule = Term.constr * bool * Coqast.t +type raw_rew_rule = Term.constr * bool * Tacexpr.raw_tactic_expr (* To add rewriting rules to a base *) val add_rew_rules : string -> raw_rew_rule list -> unit diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli new file mode 100644 index 000000000..27b926d7a --- /dev/null +++ b/tactics/contradiction.mli @@ -0,0 +1,19 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* tactic +val contradiction_on_hyp : identifier -> tactic +val contradiction : tactic diff --git a/tactics/dhyp.ml b/tactics/dhyp.ml index 10e4230d6..720cb6f5f 100644 --- a/tactics/dhyp.ml +++ b/tactics/dhyp.ml @@ -120,25 +120,24 @@ open Reduction open Proof_type open Rawterm open Tacmach +open Refiner open Tactics open Clenv open Tactics open Tacticals open Libobject open Library -open Vernacinterp open Pattern open Coqast open Ast open Pcoq +open Tacexpr (* two patterns - one for the type, and one for the type of the type *) type destructor_pattern = { d_typ: constr_pattern; d_sort: constr_pattern } -type ('a,'b) location = Hyp of 'a | Concl of 'b - (* hypothesis patterns might need to do matching on the conclusion, too. * conclusion-patterns only need to do matching on the hypothesis *) type located_destructor_pattern = @@ -150,7 +149,7 @@ type located_destructor_pattern = type destructor_data = { d_pat : located_destructor_pattern; d_pri : int; - d_code : Ast.act } (* should be of phylum tactic *) + d_code : raw_tactic_expr } (* should be of phylum tactic *) type t = (identifier,destructor_data) Nbtermdn.t type frozen_t = (identifier,destructor_data) Nbtermdn.frozen_t @@ -170,8 +169,8 @@ let rollback f x = let add (na,dd) = let pat = match dd.d_pat with - | Hyp(_,p,_) -> p.d_typ - | Concl p -> p.d_typ + | HypLocation(_,p,_) -> p.d_typ + | ConclLocation p -> p.d_typ in if Nbtermdn.in_dn tactab na then begin msgnl (str "Warning [Overriding Destructor Entry " ++ @@ -207,67 +206,59 @@ let ((inDD : destructor_data_object->obj), open_function = cache_dd; export_function = export_dd }) -let add_destructor_hint na pat pri code = +let add_destructor_hint na loc pat pri code = + begin match loc, code with + | HypLocation _, TacFun ([id],body) -> () + | ConclLocation _, _ -> () + | _ -> + errorlabstrm "add_destructor_hint" + (str "The tactic should be a function of the hypothesis name") end; + let (_,pat) = Astterm.interp_constrpattern Evd.empty (Global.env()) pat in + let pat = match loc with + | HypLocation b -> + HypLocation + (b,{d_typ=pat;d_sort=PMeta(Some (Clenv.new_meta()))}, + {d_typ=PMeta(Some (Clenv.new_meta())); + d_sort=PMeta(Some (Clenv.new_meta())) }) + | ConclLocation () -> + ConclLocation({d_typ=pat;d_sort=PMeta(Some (Clenv.new_meta()))}) in Lib.add_anonymous_leaf (inDD (na,{ d_pat = pat; d_pri=pri; d_code=code })) -let _ = - vinterp_add "HintDestruct" - (function - | [VARG_IDENTIFIER na; VARG_AST location; VARG_CONSTR patcom; - VARG_NUMBER pri; VARG_AST tacexp] -> - let loc = match location with - | Node(_,"CONCL",[]) -> Concl() - | Node(_,"DiscardableHYP",[]) -> Hyp true - | Node(_,"PreciousHYP",[]) -> Hyp false - | _ -> assert false - in - fun () -> - let (_,pat) = Astterm.interp_constrpattern - Evd.empty (Global.env()) patcom in - let code = Ast.to_act_check_vars ["$0",ETast] ETast tacexp in - add_destructor_hint na - (match loc with - | Hyp b -> - Hyp(b,{d_typ=pat;d_sort=PMeta(Some (Clenv.new_meta()))}, - { d_typ=PMeta(Some (Clenv.new_meta())); - d_sort=PMeta(Some (Clenv.new_meta())) }) - | Concl () -> - Concl({d_typ=pat;d_sort=PMeta(Some (Clenv.new_meta()))})) - pri code - | _ -> bad_vernac_args "HintDestruct") - let match_dpat dp cls gls = let cltyp = clause_type cls gls in match (cls,dp) with - | (Some id,Hyp(_,hypd,concld)) -> + | (Some id,HypLocation(_,hypd,concld)) -> (matches hypd.d_typ cltyp)@ (matches hypd.d_sort (pf_type_of gls cltyp))@ (matches concld.d_typ (pf_concl gls))@ (matches concld.d_sort (pf_type_of gls (pf_concl gls))) - | (None,Concl concld) -> + | (None,ConclLocation concld) -> (matches concld.d_typ (pf_concl gls))@ (matches concld.d_sort (pf_type_of gls (pf_concl gls))) | _ -> error "ApplyDestructor" +let forward_tac_interp = + ref (fun _ -> failwith "tac_interp is not installed for DHyp") + +let set_extern_interp f = forward_tac_interp := f + let applyDestructor cls discard dd gls = let mvb = match_dpat dd.d_pat cls gls in - let astb = match cls with - | Some id -> ["$0", Vast (nvar id)] - | None -> ["$0", Vast (nvar (id_of_string "$0"))] in - (* TODO: find the real location *) - let tcom = match Ast.eval_act dummy_loc astb dd.d_code with - | Vast tcom -> tcom - | _ -> assert false - in + let tac = match cls with + | Some id -> + let arg = Reference (RIdent (dummy_loc,id)) in + TacCall (dummy_loc, Tacexp dd.d_code, [arg]) + | None -> Tacexp dd.d_code in let discard_0 = match (cls,dd.d_pat) with - | (Some id,Hyp(discardable,_,_)) -> + | (Some id,HypLocation(discardable,_,_)) -> if discard & discardable then thin [id] else tclIDTAC - | (None,Concl _) -> tclIDTAC + | (None,ConclLocation _) -> tclIDTAC | _ -> error "ApplyDestructor" in - (tclTHEN (Tacinterp.interp tcom) discard_0) gls + tclTHEN (!forward_tac_interp (TacArg tac)) discard_0 gls + (* [DHyp id gls] @@ -284,19 +275,8 @@ let destructHyp discard id gls = let cDHyp id gls = destructHyp true id gls let dHyp id gls = destructHyp false id gls -open Tacinterp - -let _= - add_tactic "DHyp" - (function - | [Identifier id] -> dHyp id - | l -> bad_tactic_args "DHyp" l) - -let _= - add_tactic "CDHyp" - (function - | [Identifier id] -> cDHyp id - | l -> bad_tactic_args "CDHyp" l) +let h_destructHyp b id = + abstract_tactic (TacDestructHyp (b,(dummy_loc,id))) (destructHyp b id) (* [DConcl gls] @@ -309,11 +289,7 @@ let dConcl gls = let sorted_ddl = Sort.list (fun dd1 dd2 -> dd1.d_pri > dd2.d_pri) ddl in tclFIRST (List.map (applyDestructor None false) sorted_ddl) gls -let _= - add_tactic "DConcl" - (function - | [] -> dConcl - | l -> bad_tactic_args "DConcl" l) +let h_destructConcl = abstract_tactic TacDestructConcl dConcl let to2Lists (table : t) = Nbtermdn.to2lists table @@ -331,11 +307,10 @@ let rec search n = let auto_tdb n = tclTRY (tclCOMPLETE (search n)) -let sarch_depth_tdb = ref(5) - -let dyn_auto_tdb = function - | [Integer n] -> auto_tdb n - | [] -> auto_tdb !sarch_depth_tdb - | l -> bad_tactic_args "AutoTDB" l - -let h_auto_tdb = hide_tactic "AutoTDB" dyn_auto_tdb +let search_depth_tdb = ref(5) + +let depth_tdb = function + | None -> !search_depth_tdb + | Some n -> n + +let h_auto_tdb n = abstract_tactic (TacAutoTDB n) (auto_tdb (depth_tdb n)) diff --git a/tactics/dhyp.mli b/tactics/dhyp.mli index 4879bafc7..bedbb26c9 100644 --- a/tactics/dhyp.mli +++ b/tactics/dhyp.mli @@ -15,5 +15,16 @@ open Tacmach (* Programmable destruction of hypotheses and conclusions. *) +val set_extern_interp : (Tacexpr.raw_tactic_expr -> tactic) -> unit + +(* val dHyp : identifier -> tactic val dConcl : tactic +*) +val h_destructHyp : bool -> identifier -> tactic +val h_destructConcl : tactic +val h_auto_tdb : int option -> tactic + +val add_destructor_hint : + identifier -> (bool,unit) Tacexpr.location -> + Genarg.constr_ast -> int -> Tacexpr.raw_tactic_expr -> unit diff --git a/tactics/elim.ml b/tactics/elim.ml index 4008f10f7..cbe18ba36 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -23,6 +23,7 @@ open Tacmach open Tacticals open Tactics open Hiddentac +open Tacexpr let introElimAssumsThen tac ba = let nassums = @@ -85,12 +86,12 @@ let up_to_delta = ref false (* true *) let general_decompose recognizer c gl = let typc = pf_type_of gl c in - (tclTHENS (cut typc) - [tclTHEN (intro_using tmphyp_name) - (onLastHyp - (ifOnHyp recognizer (general_decompose recognizer) - (fun id -> clear [id]))); - exact_no_check c]) gl + tclTHENSV (cut typc) + [| tclTHEN (intro_using tmphyp_name) + (onLastHyp + (ifOnHyp recognizer (general_decompose recognizer) + (fun id -> clear [id]))); + exact_no_check c |] gl let head_in gls indl t = try @@ -101,19 +102,14 @@ let head_in gls indl t = in List.mem ity indl with Not_found -> false -let inductive_of_qualid gls qid = - let c = - try Declare.construct_qualified_reference qid - with Not_found -> Nametab.error_global_not_found qid - in - match kind_of_term c with - | Ind ity -> ity - | _ -> - errorlabstrm "Decompose" - (Nametab.pr_qualid qid ++ str " is not an inductive type") +let inductive_of = function + | Nametab.IndRef ity -> ity + | r -> + errorlabstrm "Decompose" + (Printer.pr_global r ++ str " is not an inductive type") let decompose_these c l gls = - let indl = List.map (inductive_of_qualid gls) l in + let indl = (*List.map inductive_of*) l in general_decompose (fun (_,t) -> head_in gls indl t) c gls let decompose_nonrec c gls = @@ -131,27 +127,23 @@ let decompose_or c gls = (fun (_,t) -> is_disjunction t) c gls -let dyn_decompose args gl = - let out_qualid = function - | Qualid qid -> qid - | l -> bad_tactic_args "DecomposeThese" [l] gl in - match args with - | Command c :: ids -> - decompose_these (pf_interp_constr gl c) (List.map out_qualid ids) gl - | Constr c :: ids -> - decompose_these c (List.map out_qualid ids) gl - | l -> bad_tactic_args "DecomposeThese" l gl - -let h_decompose = - let v_decompose = hide_tactic "DecomposeThese" dyn_decompose in - fun ids c -> - v_decompose - (Constr c :: List.map (fun x -> Qualid (Nametab.qualid_of_sp x)) ids) +let inj x = Rawterm.AN (Rawterm.dummy_loc,x) +let h_decompose l c = + Refiner.abstract_tactic + (TacDecompose (List.map inj l,c)) (decompose_these c l) +let h_decompose_or c = + Refiner.abstract_tactic (TacDecomposeOr c) (decompose_or c) + +let h_decompose_and c = + Refiner.abstract_tactic (TacDecomposeAnd c) (decompose_and c) + +(* let vernac_decompose_and = hide_constr_tactic "DecomposeAnd" decompose_and let vernac_decompose_or = hide_constr_tactic "DecomposeOr" decompose_or +*) (* The tactic Double performs a double induction *) @@ -181,7 +173,13 @@ let induction_trailer abs_i abs_j bargs = [bring_hyps hyps; clear ids; simple_elimination (mkVar id)]) gls)) -let double_ind abs_i abs_j gls = +let double_ind h1 h2 gls = + let abs_i = depth_of_quantified_hypothesis true h1 gls in + let abs_j = depth_of_quantified_hypothesis true h2 gls in + let (abs_i,abs_j) = + if abs_i < abs_j then (abs_i,abs_j) else + if abs_i > abs_j then (abs_j,abs_i) else + error "Both hypotheses are the same" in let cl = pf_concl gls in (tclTHEN (tclDO abs_i intro) (onLastHyp @@ -190,38 +188,32 @@ let double_ind abs_i abs_j gls = (introElimAssumsThen (induction_trailer abs_i abs_j)) ([],[]) (mkVar id)))) gls -let dyn_double_ind = function - | [Integer i; Integer j] -> double_ind i j - | _ -> assert false - -let _ = add_tactic "DoubleInd" dyn_double_ind - +let h_double_induction h1 h2 = + Refiner.abstract_tactic (TacDoubleInduction (h1,h2)) (double_ind h1 h2) (*****************************) (* Decomposing introductions *) (*****************************) -let rec intro_pattern p = - let clear_last = tclLAST_HYP (fun c -> (clear [destVar c])) - and case_last = tclLAST_HYP h_simplest_case in - match p with - | WildPat -> tclTHEN intro clear_last - | IdPat id -> intro_mustbe_force id - | DisjPat l -> tclTHEN introf - (tclTHENS - (tclTHEN case_last clear_last) - (List.map intro_pattern l)) - | ConjPat l -> - tclTHENSEQ [introf; case_last; clear_last; intros_pattern l] - | ListPat l -> intros_pattern l +let clear_last = tclLAST_HYP (fun c -> (clear [destVar c])) +let case_last = tclLAST_HYP h_simplest_case + +let rec intro_pattern = function + | IntroWildcard -> + tclTHEN intro clear_last + | IntroIdentifier id -> + intro_mustbe_force id + | IntroOrAndPattern l -> + tclTHEN introf + (tclTHENS + (tclTHEN case_last clear_last) + (List.map intros_pattern l)) and intros_pattern l = tclMAP intro_pattern l -let dyn_intro_pattern = function - | [] -> intros - | [Intropattern p] -> intro_pattern p - | l -> bad_tactic_args "Elim.dyn_intro_pattern" l +let intro_patterns = function + | [] -> tclREPEAT intro + | l -> tclMAP intro_pattern l -let v_intro_pattern = hide_tactic "Intros" dyn_intro_pattern +let h_intro_patterns l = Refiner.abstract_tactic (TacIntroPattern l) (intro_patterns l) -let h_intro_pattern p = v_intro_pattern [Intropattern p] diff --git a/tactics/elim.mli b/tactics/elim.mli index b67055d27..c42b27edc 100644 --- a/tactics/elim.mli +++ b/tactics/elim.mli @@ -28,12 +28,17 @@ val general_decompose : (identifier * constr -> bool) -> constr -> tactic val decompose_nonrec : constr -> tactic val decompose_and : constr -> tactic val decompose_or : constr -> tactic -val h_decompose : section_path list -> constr -> tactic +val h_decompose : inductive list -> constr -> tactic +val h_decompose_or : constr -> tactic +val h_decompose_and : constr -> tactic -val double_ind : int -> int -> tactic +val double_ind : Rawterm.quantified_hypothesis -> Rawterm.quantified_hypothesis -> tactic +val h_double_induction : Rawterm.quantified_hypothesis -> Rawterm.quantified_hypothesis->tactic -val intro_pattern : intro_pattern -> tactic -val intros_pattern : intro_pattern list -> tactic +val intro_pattern : Tacexpr.intro_pattern_expr -> tactic +val intros_pattern : Tacexpr.intro_pattern_expr list -> tactic +(* val dyn_intro_pattern : tactic_arg list -> tactic val v_intro_pattern : tactic_arg list -> tactic -val h_intro_pattern : intro_pattern -> tactic +*) +val h_intro_patterns : Tacexpr.intro_pattern_expr list -> tactic diff --git a/tactics/equality.ml b/tactics/equality.ml index c3eb15846..346b9dccb 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -34,8 +34,10 @@ open Tacticals open Tactics open Tacinterp open Tacred +open Rawterm open Vernacinterp open Coqlib +open Vernacexpr open Setoid_replace open Declarations @@ -56,7 +58,7 @@ let general_rewrite_bindings lft2rgt (c,l) gl = let _,t = splay_prod env sigma ctype in match match_with_equation t with | None -> - if l = [] + if l = NoBindings then general_s_rewrite lft2rgt c gl else error "The term provided does not end with an equation" | Some (hdcncl,_) -> @@ -68,18 +70,19 @@ let general_rewrite_bindings lft2rgt (c,l) gl = else pf_global gl (id_of_string (hdcncls^suffix)) in - tclNOTSAMEGOAL (general_elim (c,l) (elim,[])) gl + tclNOTSAMEGOAL (general_elim (c,l) (elim,NoBindings)) gl (* was tclWEAK_PROGRESS which only fails for tactics generating one subgoal and did not fail for useless conditional rewritings generating an extra condition *) (* Conditional rewriting, the success of a rewriting is related to the resolution of the conditions by a given tactic *) + let conditional_rewrite lft2rgt tac (c,bl) = - tclTHEN_i (general_rewrite_bindings lft2rgt (c,bl)) - (fun i -> if i=1 then tclIDTAC else tclCOMPLETE tac) + tclTHENSFIRSTn (general_rewrite_bindings lft2rgt (c,bl)) + [|tclIDTAC|] (tclCOMPLETE tac) -let general_rewrite lft2rgt c = general_rewrite_bindings lft2rgt (c,[]) +let general_rewrite lft2rgt c = general_rewrite_bindings lft2rgt (c,NoBindings) let rewriteLR_bindings = general_rewrite_bindings true let rewriteRL_bindings = general_rewrite_bindings false @@ -87,43 +90,6 @@ let rewriteRL_bindings = general_rewrite_bindings false let rewriteLR = general_rewrite true let rewriteRL = general_rewrite false -let dyn_rewriteLR = function - | [Command com; Bindings binds] -> - tactic_com_bind_list rewriteLR_bindings (com,binds) - | [Constr c; Cbindings binds] -> - rewriteLR_bindings (c,binds) - | _ -> assert false - -let dyn_rewriteRL = function - | [Command com; Bindings binds] -> - tactic_com_bind_list rewriteRL_bindings (com,binds) - | [Constr c; Cbindings binds] -> - rewriteRL_bindings (c,binds) - | _ -> assert false - -let dyn_conditional_rewrite lft2rgt = function - | [(Tacexp tac); (Command com);(Bindings binds)] -> - tactic_com_bind_list - (conditional_rewrite lft2rgt (Tacinterp.interp tac)) - (com,binds) - | [(Tac (tac,_)); (Constr c);(Cbindings binds)] -> - conditional_rewrite lft2rgt tac (c,binds) - | _ -> assert false - -let v_rewriteLR = hide_tactic "RewriteLR" dyn_rewriteLR -let h_rewriteLR_bindings (c,bl) = v_rewriteLR [(Constr c);(Cbindings bl)] -let h_rewriteLR c = h_rewriteLR_bindings (c,[]) - -let v_rewriteRL = hide_tactic "RewriteRL" dyn_rewriteRL -let h_rewriteRL_bindings (c,bl) = v_rewriteRL [(Constr c);(Cbindings bl)] -let h_rewriteRL c = h_rewriteRL_bindings (c,[]) - -let v_conditional_rewriteLR = - hide_tactic "CondRewriteLR" (dyn_conditional_rewrite true) -let v_conditional_rewriteRL = - hide_tactic "CondRewriteRL" (dyn_conditional_rewrite false) - - (* The Rewrite in tactic *) let general_rewrite_in lft2rgt id (c,l) gl = let ctype = pf_type_of gl c in @@ -146,37 +112,15 @@ let general_rewrite_in lft2rgt id (c,l) gl = try pf_global gl (id_of_string rwr_thm) with Not_found -> error ("Cannot find rewrite principle "^rwr_thm) in - general_elim_in id (c,l) (elim,[]) gl - -let conditional_rewrite_in lft2rgt id tac (c,bl) = - tclTHEN_i (general_rewrite_in lft2rgt id (c,bl)) - (fun i -> if i=1 then tclIDTAC else tclCOMPLETE tac) + general_elim_in id (c,l) (elim,NoBindings) gl +let rewriteLRin = general_rewrite_in true +let rewriteRLin = general_rewrite_in false -let dyn_rewrite_in lft2rgt = function - | [Identifier id;(Command com);(Bindings binds)] -> - tactic_com_bind_list (general_rewrite_in lft2rgt id) (com,binds) - | [Identifier id;(Constr c);(Cbindings binds)] -> - general_rewrite_in lft2rgt id (c,binds) - | _ -> assert false +let conditional_rewrite_in lft2rgt id tac (c,bl) = + tclTHENSFIRSTn (general_rewrite_in lft2rgt id (c,bl)) + [|tclIDTAC|] (tclCOMPLETE tac) -let dyn_conditional_rewrite_in lft2rgt = function - | [(Tacexp tac); Identifier id; (Command com);(Bindings binds)] -> - tactic_com_bind_list - (conditional_rewrite_in lft2rgt id (Tacinterp.interp tac)) - (com,binds) - | [(Tac (tac,_)); Identifier id; (Constr c);(Cbindings binds)] -> - conditional_rewrite_in lft2rgt id tac (c,binds) - | _ -> assert false - -let rewriteLR_in_tac = hide_tactic "RewriteLRin" (dyn_rewrite_in true) -let rewriteRL_in_tac = hide_tactic "RewriteRLin" (dyn_rewrite_in false) -let v_conditional_rewriteLR_in = - hide_tactic "CondRewriteLRin" (dyn_conditional_rewrite_in true) -let v_conditional_rewriteRL_in = - hide_tactic "CondRewriteRLin" (dyn_conditional_rewrite_in false) - - (* Replacing tactics *) (* eq,symeq : equality on Set and its symmetry theorem @@ -195,7 +139,7 @@ let abstract_replace (eq,sym_eq) (eqt,sym_eqt) c2 c1 unsafe gl = | Sort (Type(_)) -> (eqt,sym_eqt) | _ -> error "replace" in - (tclTHENL (elim_type (applist (e, [t1;c1;c2]))) + (tclTHENLAST (elim_type (applist (e, [t1;c1;c2]))) (tclORELSE assumption (tclTRY (tclTHEN (apply sym) assumption)))) gl else @@ -207,17 +151,6 @@ let replace c2 c1 gl = let eqT = build_coq_eqT_data.eq () in let eqT_sym = build_coq_eqT_data.sym () in abstract_replace (eq,eq_sym) (eqT,eqT_sym) c2 c1 false gl - -let dyn_replace args gl = - match args with - | [(Command c1);(Command c2)] -> - replace (pf_interp_constr gl c1) (pf_interp_constr gl c2) gl - | [(Constr c1);(Constr c2)] -> - replace c1 c2 gl - | _ -> assert false - -let v_replace = hide_tactic "Replace" dyn_replace -let h_replace c1 c2 = v_replace [(Constr c1);(Constr c2)] (* End of Eduardo's code. The rest of this file could be improved using the functions match_with_equation, etc that I defined @@ -616,14 +549,18 @@ let discrEverywhere = (fun gls -> errorlabstrm "DiscrEverywhere" (str" No discriminable equalities")) +let discr_tac = function + | None -> discrEverywhere + | Some id -> discr id + let discrConcl gls = discrClause None gls let discrHyp id gls = discrClause (Some id) gls -(**) +(* let h_discr = hide_atomic_tactic "Discr" discrEverywhere let h_discrConcl = hide_atomic_tactic "DiscrConcl" discrConcl let h_discrHyp = hide_ident_or_numarg_tactic "DiscrHyp" discrHyp -(**) +*) (* returns the sigma type (sigS, sigT) with the respective constructor depending on the sort *) @@ -870,10 +807,10 @@ let injClause = function let injConcl gls = injClause None gls let injHyp id gls = injClause (Some id) gls -(**) +(* let h_injConcl = hide_atomic_tactic "Inj" injConcl let h_injHyp = hide_ident_or_numarg_tactic "InjHyp" injHyp -(**) +*) let decompEqThen ntac id gls = let eqn = pf_whd_betadeltaiota gls (clause_type (Some id) gls) in @@ -935,10 +872,10 @@ let dEq = dEqThen (fun x -> tclIDTAC) let dEqConcl gls = dEq None gls let dEqHyp id gls = dEq (Some id) gls -(**) +(* let dEqConcl_tac = hide_atomic_tactic "DEqConcl" dEqConcl let dEqHyp_tac = hide_ident_or_numarg_tactic "DEqHyp" dEqHyp -(**) +*) let rewrite_msg = function | None -> @@ -1094,7 +1031,9 @@ let subst_tuple_term env sigma dep_pair b = |- (P e1) |- (eq T e1 e2) *) -let revSubstInConcl eqn gls = +(* Redondant avec Replace ! *) + +let substInConcl_RL eqn gls = let (lbeq,(t,e1,e2)) = find_eq_data_decompose eqn in let body = subst_tuple_term (pf_env gls) (project gls) e2 (pf_concl gls) in assert (dependent (mkRel 1) body); @@ -1105,12 +1044,14 @@ let revSubstInConcl eqn gls = |- (P e2) |- (eq T e1 e2) *) -let substInConcl eqn gls = - (tclTHENS (revSubstInConcl (swap_equands gls eqn)) +let substInConcl_LR eqn gls = + (tclTHENS (substInConcl_RL (swap_equands gls eqn)) ([tclIDTAC; swapEquandsInConcl])) gls -let substInHyp eqn id gls = +let substInConcl l2r = if l2r then substInConcl_LR else substInConcl_RL + +let substInHyp_LR eqn id gls = let (lbeq,(t,e1,e2)) = (find_eq_data_decompose eqn) in let body = subst_term e1 (clause_type (Some id) gls) in if not (dependent (mkRel 1) body) then errorlabstrm "SubstInHyp" (mt ()); @@ -1119,11 +1060,13 @@ let substInHyp eqn id gls = (tclTHENS (bareRevSubstInConcl lbeq body (t,e1,e2)) ([exact_no_check (mkVar id);tclIDTAC]))])) gls -let revSubstInHyp eqn id gls = - (tclTHENS (substInHyp (swap_equands gls eqn) id) +let substInHyp_RL eqn id gls = + (tclTHENS (substInHyp_LR (swap_equands gls eqn) id) ([tclIDTAC; swapEquandsInConcl])) gls +let substInHyp l2r = if l2r then substInHyp_LR else substInHyp_RL + let try_rewrite tac gls = try tac gls @@ -1138,20 +1081,21 @@ let try_rewrite tac gls = errorlabstrm "try_rewrite" (str "Cannot find a well type generalisation of the goal that" ++ str " makes progress the proof.") - -let subst eqn cls gls = +let subst l2r eqn cls gls = match cls with - | None -> substInConcl eqn gls - | Some id -> substInHyp eqn id gls + | None -> substInConcl l2r eqn gls + | Some id -> substInHyp l2r eqn id gls (* |- (P a) - * Subst_Concl a=b + * SubstConcl_LR a=b * |- (P b) * |- a=b *) -let substConcl_LR eqn gls = try_rewrite (subst eqn None) gls +let substConcl l2r eqn gls = try_rewrite (subst l2r eqn None) gls +let substConcl_LR = substConcl true +(* let substConcl_LR_tac = let gentac = hide_tactic "SubstConcl_LR" @@ -1162,6 +1106,7 @@ let substConcl_LR_tac = | _ -> assert false) in fun eqn -> gentac [Command eqn] +*) (* id:(P a) |- G * SubstHyp a=b id @@ -1169,20 +1114,24 @@ let substConcl_LR_tac = * id:(P a) |-a=b *) -let hypSubst id cls gls = +let hypSubst l2r id cls gls = match cls with | None -> - (tclTHENS (substInConcl (clause_type (Some id) gls)) + (tclTHENS (substInConcl l2r (clause_type (Some id) gls)) ([tclIDTAC; exact_no_check (mkVar id)])) gls | Some hypid -> - (tclTHENS (substInHyp (clause_type (Some id) gls) hypid) + (tclTHENS (substInHyp l2r (clause_type (Some id) gls) hypid) ([tclIDTAC;exact_no_check (mkVar id)])) gls +let hypSubst_LR = hypSubst true + (* id:a=b |- (P a) * HypSubst id. * id:a=b |- (P b) *) -let substHypInConcl_LR id gls = try_rewrite (hypSubst id None) gls +let substHypInConcl l2r id gls = try_rewrite (hypSubst l2r id None) gls +let substHypInConcl_LR = substHypInConcl true +(* let substHypInConcl_LR_tac = let gentac = hide_tactic "SubstHypInConcl_LR" @@ -1191,22 +1140,19 @@ let substHypInConcl_LR_tac = | _ -> assert false) in fun id -> gentac [Identifier id] +*) (* id:a=b H:(P a) |- G SubstHypInHyp id H. id:a=b H:(P b) |- G *) -let revSubst eqn cls gls = - match cls with - | None -> revSubstInConcl eqn gls - | Some id -> revSubstInHyp eqn id gls - (* |- (P b) SubstConcl_RL a=b |- (P a) |- a=b *) -let substConcl_RL eqn gls = try_rewrite (revSubst eqn None) gls +let substConcl_RL = substConcl false +(* let substConcl_RL_tac = let gentac = hide_tactic "SubstConcl_RL" @@ -1217,28 +1163,24 @@ let substConcl_RL_tac = | _ -> assert false) in fun eqn -> gentac [Command eqn] +*) (* id:(P b) |-G SubstHyp_RL a=b id id:(P a) |- G |- a=b *) -let substHyp_RL eqn id gls = try_rewrite (revSubst eqn (Some id)) gls +let substHyp l2r eqn id gls = try_rewrite (subst l2r eqn (Some id)) gls +let substHyp_RL = substHyp false -let revHypSubst id cls gls = - match cls with - | None -> - (tclTHENS (revSubstInConcl (clause_type (Some id) gls)) - ([tclIDTAC; exact_no_check (mkVar id)])) gls - | Some hypid -> - (tclTHENS (revSubstInHyp (clause_type (Some id) gls) hypid) - ([tclIDTAC;exact_no_check (mkVar id)])) gls +let hypSubst_RL = hypSubst false (* id:a=b |- (P b) * HypSubst id. * id:a=b |- (P a) *) -let substHypInConcl_RL id gls = try_rewrite (revHypSubst id None) gls +let substHypInConcl_RL = substHypInConcl false +(* let substHypInConcl_RL_tac = let gentac = hide_tactic "SubstHypInConcl_RL" @@ -1247,6 +1189,7 @@ let substHypInConcl_RL_tac = | _ -> assert false) in fun id -> gentac [Identifier id] +*) (* id:a=b H:(P b) |- G SubstHypInHyp id H. @@ -1301,26 +1244,6 @@ let (in_autorewrite_rule,out_autorewrite_rule)= Libobject.cache_function = cache_autorewrite_rule; Libobject.export_function = export_autorewrite_rule }) -(* Semantic of the HintRewrite vernacular command *) -let _ = - vinterp_add "HintRewrite" - (let rec lrules_arg lrl = function - | [] -> lrl - | (VARG_VARGLIST [VARG_CONSTR rule; VARG_STRING ort])::a - when ort="LR" or ort="RL" -> - lrules_arg (lrl@[(Astterm.interp_constr Evd.empty - (Global.env()) rule,ort="LR")]) a - | _ -> bad_vernac_args "HintRewrite" - and lbases_arg lbs = function - | [] -> lbs - | (VARG_VARGLIST ((VARG_IDENTIFIER rbase)::b))::a -> - lbases_arg (lbs@[(rbase,lrules_arg [] b)]) a - | _ -> bad_vernac_args "HintRewrite" - in - fun largs () -> - List.iter (fun c -> Lib.add_anonymous_leaf - (in_autorewrite_rule c)) (lbases_arg [] largs)) - (****The tactic****) (*To build the validation function. Length=number of unproven goals, Valid=a diff --git a/tactics/equality.mli b/tactics/equality.mli index cfda6dc34..47ec78374 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -21,29 +21,24 @@ open Pattern open Wcclausenv open Tacticals open Tactics +open Tacexpr +open Rawterm (*i*) val find_eq_pattern : sorts -> sorts -> constr -val general_rewrite_bindings : bool -> (constr * constr substitution) -> tactic +val general_rewrite_bindings : bool -> constr with_bindings -> tactic val general_rewrite : bool -> constr -> tactic -val rewriteLR_bindings : (constr * constr substitution) -> tactic -val h_rewriteLR_bindings : (constr * constr substitution) -> tactic -val rewriteRL_bindings : (constr * constr substitution) -> tactic -val h_rewriteRL_bindings : (constr * constr substitution) -> tactic +val rewriteLR_bindings : constr with_bindings -> tactic +val rewriteRL_bindings : constr with_bindings -> tactic val rewriteLR : constr -> tactic -val h_rewriteLR : constr -> tactic val rewriteRL : constr -> tactic -val h_rewriteRL : constr -> tactic -val conditional_rewrite : - bool -> tactic -> (constr * constr substitution) -> tactic -val general_rewrite_in : - bool -> identifier -> (constr * constr substitution) -> tactic +val conditional_rewrite : bool -> tactic -> constr with_bindings -> tactic +val general_rewrite_in : bool -> identifier -> constr with_bindings -> tactic val conditional_rewrite_in : - bool -> identifier -> tactic -> (constr * constr substitution) -> tactic - + bool -> identifier -> tactic -> constr with_bindings -> tactic (* usage : abstract_replace (eq,sym_eq) (eqt,sym_eqt) c2 c1 unsafe gl @@ -58,7 +53,6 @@ val abstract_replace : constr * constr -> constr * constr -> constr -> constr -> bool -> tactic val replace : constr -> constr -> tactic -val h_replace : constr -> constr -> tactic type elimination_types = | Set_Type @@ -75,13 +69,9 @@ val discrClause : clause -> tactic val discrConcl : tactic val discrHyp : identifier -> tactic val discrEverywhere : tactic -val h_discrConcl : tactic -val h_discrHyp : identifier -> tactic -val h_discrConcl : tactic -val h_discr : tactic +val discr_tac : identifier option -> tactic val inj : identifier -> tactic -val h_injHyp : identifier -> tactic -val h_injConcl : tactic +val injClause : clause -> tactic val dEq : clause -> tactic val dEqThen : (int -> tactic) -> clause -> tactic @@ -90,10 +80,12 @@ val make_iterated_tuple : env -> evar_map -> (constr * constr) -> (constr * constr) -> constr * constr * constr -val subst : constr -> clause -> tactic -val hypSubst : identifier -> clause -> tactic -val revSubst : constr -> clause -> tactic -val revHypSubst : identifier -> clause -> tactic +val substHypInConcl : bool -> identifier -> tactic +val substConcl : bool -> constr -> tactic +val substHyp : bool -> constr -> identifier -> tactic + +val hypSubst_LR : identifier -> clause -> tactic +val hypSubst_RL : identifier -> clause -> tactic val discriminable : env -> evar_map -> constr -> constr -> bool @@ -132,4 +124,3 @@ val explicit_hint_base : goal sigma -> hint_base -> rewriting_rule list val autorewrite : hint_base list -> tactic list option -> option_step -> tactic list option -> bool -> int -> tactic - diff --git a/tactics/extraargs.mli b/tactics/extraargs.mli new file mode 100644 index 000000000..d5a2b9886 --- /dev/null +++ b/tactics/extraargs.mli @@ -0,0 +1,21 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* tactic +val h_injHyp : identifier -> tactic +val h_rewriteLR : constr -> tactic + +val refine_tac : Genarg.open_constr -> tactic diff --git a/tactics/hiddentac.ml b/tactics/hiddentac.ml index bc07c9a4d..f745570e9 100644 --- a/tactics/hiddentac.ml +++ b/tactics/hiddentac.ml @@ -11,34 +11,110 @@ open Term open Proof_type open Tacmach -open Tacentries +open Rawterm +open Refiner +open Tacexpr +open Tactics +open Util + +let inj_id id = (dummy_loc,id) + +(* Basic tactics *) +let h_intro_move x y = + abstract_tactic (TacIntroMove (x, option_app inj_id y)) (intro_move x y) +let h_intro x = h_intro_move (Some x) None +let h_intros_until x = abstract_tactic (TacIntrosUntil x) (intros_until x) +let h_assumption = abstract_tactic TacAssumption assumption +let h_exact c = abstract_tactic (TacExact c) (exact_check c) +let h_apply cb = abstract_tactic (TacApply cb) (apply_with_bindings cb) +let h_elim cb cbo = abstract_tactic (TacElim (cb,cbo)) (elim cb cbo) +let h_elim_type c = abstract_tactic (TacElimType c) (elim_type c) +let h_case cb = abstract_tactic (TacCase cb) (general_case_analysis cb) +let h_case_type c = abstract_tactic (TacCaseType c) (case_type c) +let h_fix ido n = abstract_tactic (TacFix (ido,n)) (fix ido n) +let h_mutual_fix id n l = + abstract_tactic (TacMutualFix (id,n,l)) (mutual_fix id n l) +let h_cofix ido = abstract_tactic (TacCofix ido) (cofix ido) +let h_mutual_cofix id l = + abstract_tactic (TacMutualCofix (id,l)) (mutual_cofix id l) + +let h_cut c = abstract_tactic (TacCut c) (cut c) +let h_true_cut ido c = abstract_tactic (TacTrueCut (ido,c)) (true_cut ido c) +let h_forward b na c = abstract_tactic (TacForward (b,na,c)) (forward b na c) +let h_generalize cl = abstract_tactic (TacGeneralize cl) (generalize cl) +let h_generalize_dep c = abstract_tactic (TacGeneralizeDep c)(generalize_dep c) +let h_let_tac id c cl = + abstract_tactic (TacLetTac (id,c,cl)) (letin_tac true (Names.Name id) c cl) +let h_instantiate n c = + abstract_tactic (TacInstantiate (n,c)) (Evar_refiner.instantiate n c) + +(* Derived basic tactics *) +let h_old_induction h = abstract_tactic (TacOldInduction h) (old_induct h) +let h_old_destruct h = abstract_tactic (TacOldDestruct h) (old_destruct h) +let h_new_induction c = abstract_tactic (TacNewInduction c) (new_induct c) +let h_new_destruct c = abstract_tactic (TacNewDestruct c) (new_destruct c) +let h_specialize n (c,bl as d) = + abstract_tactic (TacSpecialize (n,d)) (new_hyp n c bl) +let h_lapply c = abstract_tactic (TacLApply c) (cut_and_apply c) + +(* Context management *) +let inj x = AN (Rawterm.dummy_loc,x) +let h_clear l = abstract_tactic (TacClear (List.map inj l)) (clear l) +let h_clear_body l = + abstract_tactic (TacClearBody (List.map inj l)) (clear_body l) +let h_move dep id1 id2 = + abstract_tactic (TacMove (dep,inj_id id1,inj_id id2)) (move_hyp dep id1 id2) +let h_rename id1 id2 = + abstract_tactic (TacRename (inj_id id1,inj_id id2)) (rename_hyp id1 id2) + +(* Constructors *) +let h_left l = abstract_tactic (TacLeft l) (left l) +let h_right l = abstract_tactic (TacLeft l) (right l) +let h_split l = abstract_tactic (TacSplit l) (split l) +(* Moved to tacinterp because of dependence in Tacinterp.interp +let h_any_constructor t = + abstract_tactic (TacAnyConstructor t) (any_constructor t) +*) +let h_constructor n l = + abstract_tactic (TacConstructor(AI n,l))(constructor_tac None n l) +let h_one_constructor n = h_constructor n NoBindings +let h_simplest_left = h_left NoBindings +let h_simplest_right = h_right NoBindings + +(* Conversion *) +let h_reduce r cl = abstract_tactic (TacReduce (r,cl)) (reduce r cl) +let h_change c cl = abstract_tactic (TacChange (c,cl)) (change c cl) + +(* Equivalence relations *) +let h_reflexivity = abstract_tactic TacReflexivity intros_reflexivity +let h_symmetry = abstract_tactic TacSymmetry intros_symmetry +let h_transitivity c = + abstract_tactic (TacTransitivity c) (intros_transitivity c) + +(* let h_clear ids = v_clear [(Clause (List.map (fun x -> InHyp x) ids))] let h_move dep id1 id2 = (if dep then v_move else v_move_dep) [Identifier id1;Identifier id2] -let h_contradiction = v_contradiction [] let h_reflexivity = v_reflexivity [] let h_symmetry = v_symmetry [] -let h_assumption = v_assumption [] -let h_absurd c = v_absurd [(Constr c)] -let h_exact c = v_exact [(Constr c)] let h_one_constructor i = v_constructor [(Integer i)] let h_any_constructor = v_constructor [] let h_transitivity c = v_transitivity [(Constr c)] let h_simplest_left = v_left [(Cbindings [])] let h_simplest_right = v_right [(Cbindings [])] let h_split c = v_split [(Constr c);(Cbindings [])] -let h_apply c s = v_apply [(Constr c);(Cbindings s)] -let h_simplest_apply c = v_apply [(Constr c);(Cbindings [])] -let h_cut c = v_cut [(Constr c)] -let h_simplest_elim c = v_elim [(Constr c);(Cbindings [])] -let h_elimType c = v_elimType [(Constr c)] +*) + +let h_simplest_apply c = h_apply (c,NoBindings) +let h_simplest_elim c = h_elim (c,NoBindings) None +(* let h_inductionInt i = v_induction[(Integer i)] let h_inductionId id = v_induction[(Identifier id)] -let h_simplest_case c = v_case [(Constr c);(Cbindings [])] -let h_caseType c = v_caseType [(Constr c)] +*) +let h_simplest_case c = h_case (c,NoBindings) +(* let h_destructInt i = v_destruct [(Integer i)] let h_destructId id = v_destruct [(Identifier id)] - - +*) diff --git a/tactics/hiddentac.mli b/tactics/hiddentac.mli index 517889c16..7a83eff5e 100644 --- a/tactics/hiddentac.mli +++ b/tactics/hiddentac.mli @@ -13,35 +13,98 @@ open Names open Term open Proof_type open Tacmach -open Tacentries +open Tacexpr +open Rawterm (*i*) (* Tactics for the interpreter. They left a trace in the proof tree when they are called. *) -val h_clear : identifier list -> tactic -val h_move : bool -> identifier -> identifier -> tactic -val h_contradiction : tactic -val h_reflexivity : tactic -val h_symmetry : tactic +(* Basic tactics *) + +val h_intro_move : identifier option -> identifier option -> tactic +val h_intro : identifier -> tactic +val h_intros_until : quantified_hypothesis -> tactic + val h_assumption : tactic -val h_absurd : constr -> tactic val h_exact : constr -> tactic + +val h_apply : constr with_bindings -> tactic + +val h_elim : constr with_bindings -> + constr with_bindings option -> tactic +val h_elim_type : constr -> tactic +val h_case : constr with_bindings -> tactic +val h_case_type : constr -> tactic + +val h_mutual_fix : identifier -> int -> + (identifier * int * constr) list -> tactic +val h_fix : identifier option -> int -> tactic +val h_mutual_cofix : identifier -> (identifier * constr) list -> tactic +val h_cofix : identifier option -> tactic + +val h_cut : constr -> tactic +val h_true_cut : identifier option -> constr -> tactic +val h_generalize : constr list -> tactic +val h_generalize_dep : constr -> tactic +val h_forward : bool -> name -> constr -> tactic +val h_let_tac : identifier -> constr -> identifier clause_pattern -> tactic +val h_instantiate : int -> constr -> tactic + +(* Derived basic tactics *) + +val h_old_induction : quantified_hypothesis -> tactic +val h_old_destruct : quantified_hypothesis -> tactic +val h_new_induction : constr induction_arg -> tactic +val h_new_destruct : constr induction_arg -> tactic +val h_specialize : int option -> constr with_bindings -> tactic +val h_lapply : constr -> tactic + +(* Automation tactic : see Auto *) + + +(* Context management *) +val h_clear : identifier list -> tactic +val h_clear_body : identifier list -> tactic +val h_move : bool -> identifier -> identifier -> tactic +val h_rename : identifier -> identifier -> tactic + + +(* Constructors *) +(* +val h_any_constructor : tactic -> tactic +*) +val h_constructor : int -> constr substitution -> tactic +val h_left : constr substitution -> tactic +val h_right : constr substitution -> tactic +val h_split : constr substitution -> tactic + val h_one_constructor : int -> tactic -val h_any_constructor : tactic -val h_transitivity : constr -> tactic val h_simplest_left : tactic val h_simplest_right : tactic -val h_split : constr -> tactic -val h_apply : constr -> constr substitution -> tactic + + +(* Conversion *) +val h_reduce : Tacred.red_expr -> hyp_location list -> tactic +val h_change : constr -> hyp_location list -> tactic + +(* Equivalence relations *) +val h_reflexivity : tactic +val h_symmetry : tactic +val h_transitivity : constr -> tactic + +(* +val h_reflexivity : tactic +val h_symmetry : tactic +val h_transitivity : constr -> tactic +*) val h_simplest_apply : constr -> tactic -val h_cut : constr -> tactic val h_simplest_elim : constr -> tactic -val h_elimType : constr -> tactic val h_simplest_case : constr -> tactic -val h_caseType : constr -> tactic +(* val h_inductionInt : int -> tactic val h_inductionId : identifier -> tactic val h_destructInt : int -> tactic val h_destructId : identifier -> tactic +*) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 39c2bd8f7..8e7648704 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -125,6 +125,17 @@ let is_unit_type t = op2bool (match_with_unit_type t) inductive binary relation R, so that R has only one constructor establishing its reflexivity. *) +(* ["(A : ?)(x:A)(? A x x)"] and ["(x : ?)(? x x)"] *) +let x = Name (id_of_string "x") +let y = Name (id_of_string "y") +let name_A = Name (id_of_string "A") +let coq_refl_rel1_pattern = + PProd + (name_A, PMeta None, + PProd (x, PRel 1, PApp (PMeta None, [|PRel 2; PRel 1; PRel 1|]))) +let coq_refl_rel2_pattern = + PProd (x, PMeta None, PApp (PMeta None, [|PRel 1; PRel 1|])) + let match_with_equation t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with @@ -133,8 +144,8 @@ let match_with_equation t = let constr_types = mip.mind_nf_lc in let nconstr = Array.length mip.mind_consnames in if nconstr = 1 && - (is_matching (build_coq_refl_rel1_pattern ()) constr_types.(0) || - is_matching (build_coq_refl_rel1_pattern ()) constr_types.(0)) + (is_matching coq_refl_rel1_pattern constr_types.(0) || + is_matching coq_refl_rel1_pattern constr_types.(0)) then Some (hdapp,args) else @@ -143,9 +154,13 @@ let match_with_equation t = let is_equation t = op2bool (match_with_equation t) +(* ["(?1 -> ?2)"] *) +let imp a b = PProd (Anonymous, a, b) +let coq_arrow_pattern = imp (PMeta (Some 1)) (PMeta (Some 2)) + let match_with_nottype t = try - match matches (build_coq_arrow_pattern ()) t with + match matches coq_arrow_pattern t with | [(1,arg);(2,mind)] -> if is_empty_type mind then Some (mind,arg) else None | _ -> anomaly "Incorrect pattern matching" diff --git a/tactics/inv.ml b/tactics/inv.ml index 07466c497..5cd54c80e 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -241,8 +241,8 @@ let generalizeRewriteIntros tac depids id gls = let projectAndApply thin id depids gls = let env = pf_env gls in - let subst_hyp_LR id = tclTRY(hypSubst id None) in - let subst_hyp_RL id = tclTRY(revHypSubst id None) in + let subst_hyp_LR id = tclTRY(hypSubst_LR id None) in + let subst_hyp_RL id = tclTRY(hypSubst_RL id None) in let subst_hyp gls = let orient_rule id = let (t,t1,t2) = dest_eq gls (pf_get_hyp_typ gls id) in @@ -372,7 +372,7 @@ let raw_inversion inv_kind indbinding id status gl = case_nodep_then_using in (tclTHENS - (true_cut_anon cut_concl) + (true_cut None cut_concl) [case_tac (introCaseAssumsThen (rewrite_equations_tac inv_kind id neqns)) (Some elim_predicate) ([],[]) newc; onLastHyp @@ -421,6 +421,11 @@ let inversion inv_kind status id gls = (* Specializing it... *) +let inv_gen gene thin status = try_intros_until (inversion (gene,thin) status) + +open Tacexpr + +(* let hinv_kind = Quoted_string "HalfInversion" let inv_kind = Quoted_string "Inversion" let invclr_kind = Quoted_string "InversionClear" @@ -429,69 +434,19 @@ let com_of_id id = if id = hinv_kind then None else if id = inv_kind then Some false else Some true +*) -(* Inv generates nodependent inversion *) -let (half_inv_tac, inv_tac, inv_clear_tac) = - let gentac = - hide_tactic "Inv" - (function - | ic :: [id_or_num] -> - tactic_try_intros_until - (inversion (false, com_of_id ic) NoDep) - id_or_num - | l -> bad_tactic_args "Inv" l) - in - ((fun id -> gentac [hinv_kind; Identifier id]), - (fun id -> gentac [inv_kind; Identifier id]), - (fun id -> gentac [invclr_kind; Identifier id])) - - -(* Inversion without intros. No vernac entry yet! *) -let named_inv = - let gentac = - hide_tactic "NamedInv" - (function - | [ic; Identifier id] -> inversion (true, com_of_id ic) NoDep id - | l -> bad_tactic_args "NamedInv" l) - in - (fun ic id -> gentac [ic; Identifier id]) - -(* Generates a dependent inversion with a certain generalisation of the goal *) -let (half_dinv_tac, dinv_tac, dinv_clear_tac) = - let gentac = - hide_tactic "DInv" - (function - | ic :: [id_or_num] -> - tactic_try_intros_until - (inversion (false, com_of_id ic) (Dep None)) - id_or_num - | l -> bad_tactic_args "DInv" l) - in - ((fun id -> gentac [hinv_kind; Identifier id]), - (fun id -> gentac [inv_kind; Identifier id]), - (fun id -> gentac [invclr_kind; Identifier id])) - -(* generates a dependent inversion using a given generalisation of the goal *) -let (half_dinv_with, dinv_with, dinv_clear_with) = - let gentac = - hide_tactic "DInvWith" - (function - | [ic; id_or_num; Command com] -> - tactic_try_intros_until - (fun id gls -> - inversion (false, com_of_id ic) - (Dep (Some (pf_interp_constr gls com))) id gls) - id_or_num - | [ic; id_or_num; Constr c] -> - tactic_try_intros_until - (fun id gls -> - inversion (false, com_of_id ic) (Dep (Some c)) id gls) - id_or_num - | l -> bad_tactic_args "DInvWith" l) - in - ((fun id c -> gentac [hinv_kind; Identifier id; Constr c]), - (fun id c -> gentac [inv_kind; Identifier id; Constr c]), - (fun id c -> gentac [invclr_kind; Identifier id; Constr c])) +let inv k id = inv_gen false k NoDep id + +let half_inv_tac id = inv None (Rawterm.NamedHyp id) +let inv_tac id = inv (Some false) (Rawterm.NamedHyp id) +let inv_clear_tac id = inv (Some true) (Rawterm.NamedHyp id) + +let dinv k c id = inv_gen false k (Dep c) id + +let half_dinv_tac id = dinv None None (Rawterm.NamedHyp id) +let dinv_tac id = dinv (Some false) None (Rawterm.NamedHyp id) +let dinv_clear_tac id = dinv (Some true) None (Rawterm.NamedHyp id) (* InvIn will bring the specified clauses into the conclusion, and then * perform inversion on the named hypothesis. After, it will intro them @@ -515,22 +470,4 @@ let invIn com id ids gls = gls with e -> wrap_inv_error id e -let invIn_tac = - let gentac = - hide_tactic "InvIn" - (function - | (com::(Identifier id)::hl as ll) -> - let hl' = - List.map - (function - | Identifier id -> id - | _ -> bad_tactic_args "InvIn" ll) hl - in - invIn (com_of_id com) id hl' - | ll -> bad_tactic_args "InvIn" ll) - in - fun com id hl -> - gentac - ((Identifier com) - ::(Identifier id) - ::(List.map (fun id -> (Identifier id)) hl)) +let invIn_gen com id idl = try_intros_until (fun id -> invIn com id idl) id diff --git a/tactics/inv.mli b/tactics/inv.mli index 792f13261..9375efdea 100644 --- a/tactics/inv.mli +++ b/tactics/inv.mli @@ -14,14 +14,25 @@ open Term open Tacmach (*i*) +type inversion_status = Dep of constr option | NoDep + +val inv_gen : + bool -> bool option -> inversion_status -> Rawterm.quantified_hypothesis -> tactic +val invIn_gen : bool option -> Rawterm.quantified_hypothesis -> identifier list -> tactic + +val inv : bool option -> Rawterm.quantified_hypothesis -> tactic +val dinv : bool option -> constr option -> Rawterm.quantified_hypothesis -> tactic val half_inv_tac : identifier -> tactic val inv_tac : identifier -> tactic val inv_clear_tac : identifier -> tactic val half_dinv_tac : identifier -> tactic val dinv_tac : identifier -> tactic val dinv_clear_tac : identifier -> tactic +(* val half_dinv_with : identifier -> constr -> tactic val dinv_with : identifier -> constr -> tactic val dinv_clear_with : identifier -> constr -> tactic - +*) +(* val invIn_tac : identifier -> identifier -> identifier list -> tactic +*) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 343361815..1a83dbb5d 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -32,6 +32,7 @@ open Wcclausenv open Tacticals open Tactics open Inv +open Vernacexpr open Safe_typing let not_work_message = "tactic fails to build the inversion lemma, may be because the predicate has arguments that depend on other arguments" @@ -246,7 +247,7 @@ let add_inversion_lemma name env sigma t sort dep inv_op = (ConstantEntry { const_entry_body = invProof; const_entry_type = None; const_entry_opaque = false }, - NeverDischarge) + Nametab.NeverDischarge) in () (* open Pfedit *) @@ -269,17 +270,6 @@ let inversion_lemma_from_goal n na id sort dep_option inv_op = str"which are not free in its instance"); *) add_inversion_lemma na env sigma t sort dep_option inv_op -open Vernacinterp - -let _ = - vinterp_add - "MakeInversionLemmaFromHyp" - (function - | [VARG_NUMBER n; VARG_IDENTIFIER na; VARG_IDENTIFIER id] -> - fun () -> - inversion_lemma_from_goal n na id mk_Prop false inv_clear_tac - | _ -> bad_vernac_args "MakeInversionLemmaFromHyp") - let add_inversion_lemma_exn na com comsort bool tac = let env = Global.env () and sigma = Evd.empty in let c = Astterm.interp_type sigma env com in @@ -290,51 +280,6 @@ let add_inversion_lemma_exn na com comsort bool tac = | UserError ("Case analysis",s) -> (* rfrence Indrec *) errorlabstrm "Inv needs Nodep Prop Set" s -let _ = - vinterp_add - "MakeInversionLemma" - (function - | [VARG_IDENTIFIER na; VARG_CONSTR com; VARG_CONSTR sort] -> - fun () -> - add_inversion_lemma_exn na com sort false inv_clear_tac - | _ -> bad_vernac_args "MakeInversionLemma") - -let _ = - vinterp_add - "MakeSemiInversionLemmaFromHyp" - (function - | [VARG_NUMBER n; VARG_IDENTIFIER na; VARG_IDENTIFIER id] -> - fun () -> - inversion_lemma_from_goal n na id mk_Prop false half_inv_tac - | _ -> bad_vernac_args "MakeSemiInversionLemmaFromHyp") - -let _ = - vinterp_add - "MakeSemiInversionLemma" - (function - | [VARG_IDENTIFIER na; VARG_CONSTR com; VARG_CONSTR sort] -> - fun () -> - add_inversion_lemma_exn na com sort false half_inv_tac - | _ -> bad_vernac_args "MakeSemiInversionLemma") - -let _ = - vinterp_add - "MakeDependentInversionLemma" - (function - | [VARG_IDENTIFIER na; VARG_CONSTR com; VARG_CONSTR sort] -> - fun () -> - add_inversion_lemma_exn na com sort true dinv_clear_tac - | _ -> bad_vernac_args "MakeDependentInversionLemma") - -let _ = - vinterp_add - "MakeDependentSemiInversionLemma" - (function - | [VARG_IDENTIFIER na; VARG_CONSTR com; VARG_CONSTR sort] -> - fun () -> - add_inversion_lemma_exn na com sort true half_dinv_tac - | _ -> bad_vernac_args "MakeDependentSemiInversionLemma") - (* ================================= *) (* Applying a given inversion lemma *) (* ================================= *) @@ -355,21 +300,7 @@ let lemInv id c gls = (str "Cannot refine current goal with the lemma " ++ prterm_env (Global.env()) c) -let useInversionLemma = - let gentac = - hide_tactic "UseInversionLemma" - (function - | [id_or_num; Command com] -> - tactic_try_intros_until - (fun id gls -> lemInv id (pf_interp_constr gls com) gls) - id_or_num - | [id_or_num; Constr c] -> - tactic_try_intros_until - (fun id gls -> lemInv id c gls) - id_or_num - | l -> bad_vernac_args "useInversionLemma" l) - in - fun id c -> gentac [Identifier id;Constr c] +let lemInv_gen id c = try_intros_until (fun id -> lemInv id c) id let lemInvIn id c ids gls = let hyps = List.map (pf_get_hyp gls) ids in @@ -387,26 +318,4 @@ let lemInvIn id c ids gls = | UserError(a,b) -> errorlabstrm "LemInvIn" b *) -let useInversionLemmaIn = - let gentac = - hide_tactic "UseInversionLemmaIn" - (function - | ((Identifier id)::(Command com)::hl as ll) -> - fun gls -> - lemInvIn id (pf_interp_constr gls com) - (List.map - (function - | (Identifier id) -> id - | _ -> bad_vernac_args "UseInversionLemmaIn" ll) hl) gls - | ((Identifier id)::(Constr c)::hl as ll) -> - fun gls -> - lemInvIn id c - (List.map - (function - | (Identifier id) -> id - | _ -> bad_vernac_args "UseInversionLemmaIn" ll) hl) gls - | ll -> bad_vernac_args "UseInversionLemmaIn" ll) - in - fun id c hl -> - gentac ((Identifier id)::(Constr c) - ::(List.map (fun id -> (Identifier id)) hl)) +let lemInvIn_gen id c l = try_intros_until (fun id -> lemInvIn id c l) id diff --git a/tactics/leminv.mli b/tactics/leminv.mli new file mode 100644 index 000000000..3d5f33c66 --- /dev/null +++ b/tactics/leminv.mli @@ -0,0 +1,15 @@ + +open Names +open Term +open Rawterm +open Proof_type + +val lemInv_gen : quantified_hypothesis -> constr -> tactic +val lemInvIn_gen : quantified_hypothesis -> constr -> identifier list -> tactic + +val inversion_lemma_from_goal : + int -> identifier -> identifier -> sorts -> bool -> + (identifier -> tactic) -> unit +val add_inversion_lemma_exn : + identifier -> Coqast.t -> Coqast.t -> bool -> (identifier -> tactic) -> unit + diff --git a/tactics/refine.ml b/tactics/refine.ml index a942b37b7..d1f380e49 100644 --- a/tactics/refine.ml +++ b/tactics/refine.ml @@ -340,20 +340,3 @@ let refine oc gl = let th = compute_metamap env gmm c in tcc_aux th gl -let refine_tac = Tacmach.hide_openconstr_tactic "Refine" refine - -open Proof_type - -let dyn_tcc args gl = match args with - | [Command com] -> - let env = pf_env gl in - refine - (Astterm.interp_casted_openconstr (project gl) env com (pf_concl gl)) - gl - | [OpenConstr c] -> - refine c gl - | _ -> assert false - -let tcc_tac = hide_tactic "Tcc" dyn_tcc - - diff --git a/tactics/refine.mli b/tactics/refine.mli index ec251b8ee..4f72a2622 100644 --- a/tactics/refine.mli +++ b/tactics/refine.mli @@ -12,4 +12,6 @@ open Term open Tacmach val refine : Pretyping.open_constr -> tactic +(* val refine_tac : Pretyping.open_constr -> tactic +*) diff --git a/tactics/setoid_replace.ml b/tactics/setoid_replace.ml index 1362ba2cc..a147997ba 100644 --- a/tactics/setoid_replace.ml +++ b/tactics/setoid_replace.ml @@ -24,6 +24,8 @@ open Environ open Termast open Command open Tactics +open Tacticals +open Vernacexpr open Safe_typing open Nametab @@ -237,14 +239,14 @@ let add_setoid a aeq th = let eq_ext_name2 = gen_eq_lem_name () in let _ = Declare.declare_constant eq_ext_name ((ConstantEntry {const_entry_body = eq_morph; - const_entry_type = None; + const_entry_type = None; const_entry_opaque = true}), - Declare.NeverDischarge) in + Nametab.NeverDischarge) in let _ = Declare.declare_constant eq_ext_name2 ((ConstantEntry {const_entry_body = eq_morph2; const_entry_type = None; const_entry_opaque = true}), - Declare.NeverDischarge) in + Nametab.NeverDischarge) in let eqmorph = (current_constant eq_ext_name) in let eqmorph2 = (current_constant eq_ext_name2) in (Lib.add_anonymous_leaf @@ -257,21 +259,9 @@ let add_setoid a aeq th = else errorlabstrm "Add Setoid" (str "Not a valid setoid theory") (* The vernac command "Add Setoid" *) +let add_setoid a aeq th = + add_setoid (constr_of a) (constr_of aeq) (constr_of th) -let constr_of_comarg = function - | VARG_CONSTR c -> constr_of c - | _ -> anomaly "Add Setoid" - -let _ = - vinterp_add "AddSetoid" - (function - | [(VARG_CONSTR a); (VARG_CONSTR aeq); (VARG_CONSTR th)] -> - (fun () -> (add_setoid - (constr_of a) - (constr_of aeq) - (constr_of th))) - | _ -> anomaly "AddSetoid") - (***************** Adding a morphism to the database ****************************) (* We maintain a table of the currently edited proofs of morphism lemma @@ -347,7 +337,7 @@ let gen_compat_lemma env m body larg lisset = | _ -> assert false in aux larg lisset 0 -let new_morphism m id = +let new_morphism m id hook = if morphism_table_mem m then errorlabstrm "New Morphism" (str "The term " ++ prterm m ++ str " is already declared as a morphism") @@ -368,9 +358,9 @@ let new_morphism m id = let poss = (List.map setoid_table_mem args_t) in let lem = (gen_compat_lemma env m body args_t poss) in let lemast = (ast_of_constr true env lem) in - new_edited id m poss; - start_proof_com (Some id) Declare.NeverDischarge lemast; - (Options.if_verbose Vernacentries.show_open_subgoals ())) + new_edited id m poss; + start_proof_com (Some id) (false,Nametab.NeverDischarge) lemast hook; + (Options.if_verbose Vernacentries.show_open_subgoals ())) let rec sub_bool l1 n = function | [] -> [] @@ -464,7 +454,7 @@ let add_morphism lem_name (m,profil) = ((ConstantEntry {const_entry_body = lem_2; const_entry_type = None; const_entry_opaque = true}), - Declare.NeverDischarge) in + Nametab.NeverDischarge) in let lem2 = (current_constant lem2_name) in (Lib.add_anonymous_leaf (morphism_to_obj (m, @@ -481,54 +471,13 @@ let add_morphism lem_name (m,profil) = arg_types = args_t; lem2 = None})))); Options.if_verbose ppnl (prterm m ++str " is registered as a morphism") - -let _ = - let current_save = vinterp_map "SaveNamed" in - overwriting_vinterp_add "SaveNamed" - (function - |[] -> (fun () -> - let pf_id = Pfedit.get_current_proof_name () in - current_save [] (); - if (is_edited pf_id) - then - (add_morphism pf_id (what_edited pf_id); - no_more_edited pf_id)) - | _ -> assert false) +let morphism_hook stre ref = + let pf_id = basename (sp_of_global (Global.env()) ref) in + if (is_edited pf_id) + then + (add_morphism pf_id (what_edited pf_id); no_more_edited pf_id) -let _ = - let current_defined = vinterp_map "DefinedNamed" in - overwriting_vinterp_add "DefinedNamed" - (function - |[] -> (fun () -> let pf_id = Pfedit.get_current_proof_name () in - current_defined [] (); - if (is_edited pf_id) - then - (add_morphism pf_id (what_edited pf_id); - no_more_edited pf_id)) - | _ -> assert false) - -(* -let _ = - vinterp_add "NewMorphism" - (function - | [(VARG_IDENTIFIER s)] -> - (fun () -> - try - (let m = Declare.global_reference CCI s in - new_morphism m (gen_lem_name m)) - with - Not_found -> - errorlabstrm "New Morphism" - (str "The term " ++ str(string_of_id s) ++ str" is not a known name")) - | _ -> anomaly "NewMorphism") -*) - -let _ = - vinterp_add "NamedNewMorphism" - (function - | [(VARG_IDENTIFIER s);(VARG_CONSTR m)] -> - (fun () -> new_morphism (constr_of m) s) - | _ -> anomaly "NewMorphism") +let new_named_morphism id m = new_morphism (constr_of m) id morphism_hook (****************************** The tactic itself *******************************) @@ -670,21 +619,3 @@ let general_s_rewrite lft2rgt c gl = let setoid_rewriteLR = general_s_rewrite true let setoid_rewriteRL = general_s_rewrite false - -let dyn_setoid_replace = function - | [(Constr c1);(Constr c2)] -> (fun gl -> setoid_replace c1 c2 None gl) - | _ -> invalid_arg "Setoid_replace : Bad arguments" - -let h_setoid_replace = hide_tactic "Setoid_replace" dyn_setoid_replace - -let dyn_setoid_rewriteLR = function - | [(Constr c)] -> setoid_rewriteLR c - | _ -> invalid_arg "Setoid_rewrite : Bad arguments" - -let h_setoid_rewriteLR = hide_tactic "Setoid_rewriteLR" dyn_setoid_rewriteLR - -let dyn_setoid_rewriteRL = function - | [(Constr c)] -> setoid_rewriteRL c - | _ -> invalid_arg "Setoid_rewrite : Bad arguments" - -let h_setoid_rewriteRL = hide_tactic "Setoid_rewriteRL" dyn_setoid_rewriteRL diff --git a/tactics/setoid_replace.mli b/tactics/setoid_replace.mli index c280bb269..d8bc55656 100644 --- a/tactics/setoid_replace.mli +++ b/tactics/setoid_replace.mli @@ -10,6 +10,7 @@ open Term open Proof_type +open Genarg val equiv_list : unit -> constr list @@ -20,3 +21,7 @@ val setoid_rewriteLR : constr -> tactic val setoid_rewriteRL : constr -> tactic val general_s_rewrite : bool -> constr -> tactic + +val add_setoid : constr_ast -> constr_ast -> constr_ast -> unit + +val new_named_morphism : Names.identifier -> constr_ast -> unit diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml new file mode 100644 index 000000000..2d8f7c904 --- /dev/null +++ b/tactics/tacinterp.ml @@ -0,0 +1,1738 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* x + | MetaId (loc,_) -> error_syntactic_metavariables_not_allowed loc + +(* Values for interpretation *) +type value = + | VClosure of interp_sign * raw_tactic_expr + | VTactic of tactic (* For mixed ML/Ltac tactics (e.g. Tauto) *) + | VFTactic of value list * (value list->tactic) + | VRTactic of (goal list sigma * validation) + | VContext of interp_sign * direction_flag + * (pattern_ast,raw_tactic_expr) match_rule list + | VFun of (identifier * value) list * identifier option list *raw_tactic_expr + | VVoid + | VInteger of int + | VIdentifier of identifier (* idents which are not refs, as in "Intro H" *) + | VConstr of constr + | VConstr_context of constr + | VRec of value ref + +(* Signature for interpretation: val_interp and interpretation functions *) +and interp_sign = + { evc : Evd.evar_map; + env : Environ.env; + lfun : (identifier * value) list; + lmatch : (int * constr) list; + goalopt : goal sigma option; + debug : debug_info } + +(* For tactic_of_value *) +exception NotTactic + +(* Gives the constr corresponding to a Constr tactic_arg *) +let constr_of_VConstr = function + | VConstr c -> c + | _ -> anomalylabstrm "constr_of_VConstr" (str "Not a Constr tactic_arg") + +(* Gives the constr corresponding to a Constr_context tactic_arg *) +let constr_of_VConstr_context = function + | VConstr_context c -> c + | _ -> + anomalylabstrm "constr_of_VConstr_context" (str + "Not a Constr_context tactic_arg") + +(* +(* Gives identifiers and makes the possible injection constr -> ident *) +let make_ids ast = function + | Identifier id -> id + | Constr c -> + (try destVar c with + | Invalid_argument "destVar" -> + anomalylabstrm "make_ids" + (str "This term cannot be reduced to an identifier" ++ fnl () ++ + print_ast ast)) + | _ -> anomalylabstrm "make_ids" (str "Not an identifier") +*) + +let pr_value env = function + | VVoid -> str "()" + | VInteger n -> int n + | VIdentifier id -> pr_id id + | VConstr c -> Printer.prterm_env env c + | VConstr_context c -> Printer.prterm_env env c + | (VClosure _ | VTactic _ | VFTactic _ | VRTactic _ | + VContext _ | VFun _ | VRec _) -> str "" + +(* Transforms a named_context into a (string * constr) list *) +let make_hyps = List.map (fun (id,_,typ) -> (id,body_of_type typ)) + +(* Transforms an id into a constr if possible *) +let constr_of_id ist id = + match ist.goalopt with + | None -> construct_reference ist.env id + | Some goal -> + let hyps = pf_hyps goal in + if mem_named_context id hyps then + mkVar id + else + let csr = global_qualified_reference (make_short_qualid id) in + (match kind_of_term csr with + | Var _ -> raise Not_found + | _ -> csr) + +(* To embed several objects in Coqast.t *) +let ((tactic_in : (interp_sign -> raw_tactic_expr) -> Dyn.t), + (tactic_out : Dyn.t -> (interp_sign -> raw_tactic_expr))) = + create "tactic" + +let ((value_in : value -> Dyn.t), + (value_out : Dyn.t -> value)) = create "value" + +let tacticIn t = TacArg (TacDynamic (dummy_loc,tactic_in t)) +let tacticOut = function + | TacArg (TacDynamic (_,d)) -> + if (tag d) = "tactic" then + tactic_out d + else + anomalylabstrm "tacticOut" (str "Dynamic tag should be tactic") + | ast -> + anomalylabstrm "tacticOut" + (str "Not a Dynamic ast: " (* ++ print_ast ast*) ) + +let valueIn t = TacDynamic (dummy_loc,value_in t) +let valueOut = function + | TacDynamic (_,d) -> + if (tag d) = "value" then + value_out d + else + anomalylabstrm "valueOut" (str "Dynamic tag should be value") + | ast -> + anomalylabstrm "valueOut" + (str "Not a Dynamic ast: " (* ++ print_ast ast*) ) + +let constrIn c = constrIn c +let constrOut = constrOut + +let loc = dummy_loc + +(* Table of interpretation functions *) +let interp_tab = + (Hashtbl.create 17 : (string , interp_sign -> Coqast.t -> value) Hashtbl.t) + +(* Adds an interpretation function *) +let interp_add (ast_typ,interp_fun) = + try + Hashtbl.add interp_tab ast_typ interp_fun + with + Failure _ -> + errorlabstrm "interp_add" + (str "Cannot add the interpretation function for " ++ str ast_typ ++ str " twice") + +(* Adds a possible existing interpretation function *) +let overwriting_interp_add (ast_typ,interp_fun) = + if Hashtbl.mem interp_tab ast_typ then + begin + Hashtbl.remove interp_tab ast_typ; + warning ("Overwriting definition of tactic interpreter command " ^ ast_typ) + end; + Hashtbl.add interp_tab ast_typ interp_fun + +(* Finds the interpretation function corresponding to a given ast type *) +let look_for_interp = Hashtbl.find interp_tab + +(* Globalizes the identifier *) + +let find_reference ist qid = + (* We first look for a variable of the current proof *) + match Nametab.repr_qualid qid, ist.goalopt with + | (d,id),Some gl when repr_dirpath d = [] & List.mem id (pf_ids_of_hyps gl) + -> VarRef id + | _ -> Nametab.locate qid + +let coerce_to_reference ist = function + | VConstr c -> + (try reference_of_constr c + with Not_found -> invalid_arg_loc (loc, "Not a reference")) +(* | VIdentifier id -> VarRef id*) + | v -> errorlabstrm "coerce_to_reference" + (str "The value" ++ spc () ++ pr_value ist.env v ++ + str "cannot be coerced to a reference") + +(* turns a value into an evaluable reference *) +let error_not_evaluable s = + errorlabstrm "evalref_of_ref" + (str "Cannot coerce" ++ spc () ++ s ++ spc () ++ + str "to an evaluable reference") + +let coerce_to_evaluable_ref env c = + let ev = match c with + | VConstr c when isConst c -> EvalConstRef (destConst c) + | VConstr c when isVar c -> EvalVarRef (destVar c) +(* | VIdentifier id -> EvalVarRef id*) + | _ -> error_not_evaluable (pr_value env c) + in + if not (Tacred.is_evaluable env ev) then + error_not_evaluable (pr_value env c); + ev + +let coerce_to_inductive = function + | VConstr c when isInd c -> destInd c + | x -> + try + let r = match x with + | VConstr c -> reference_of_constr c + | _ -> failwith "" in + errorlabstrm "coerce_to_inductive" + (Printer.pr_global r ++ str " is not an inductive type") + with _ -> + errorlabstrm "coerce_to_inductive" + (str "Found an argument which should be an inductive type") + +(* Summary and Object declaration *) +let mactab = ref Gmap.empty + +let lookup qid = Gmap.find (locate_tactic qid) !mactab + +let _ = + let init () = mactab := Gmap.empty in + let freeze () = !mactab in + let unfreeze fs = mactab := fs in + Summary.declare_summary "tactic-definition" + { Summary.freeze_function = freeze; + Summary.unfreeze_function = unfreeze; + Summary.init_function = init; + Summary.survive_section = false } + +(* Interpretation of extra generic arguments *) +type genarg_interp_type = + interp_sign -> raw_generic_argument -> closed_generic_argument +let extragenargtab = ref (Gmap.empty : (string,genarg_interp_type) Gmap.t) +let add_genarg_interp id f = extragenargtab := Gmap.add id f !extragenargtab +let lookup_genarg_interp id = + try Gmap.find id !extragenargtab + with Not_found -> failwith ("No interpretation function found for entry "^id) + + +(* Unboxes VRec *) +let unrec = function + | VRec v -> !v + | a -> a + +(************* GLOBALIZE ************) + +(* We have identifier <| global_reference <| constr *) + +(* Globalize a name which can be fresh *) +let glob_ident l ist x = + (* We use identifier both for variables and new names; thus nothing to do *) + if List.mem x (fst ist) then () else l:=x::!l; + x + +let glob_name l ist = function + | Anonymous -> Anonymous + | Name id -> Name (glob_ident l ist id) + +(* Globalize a name which must be bound -- actually just check it is bound *) +let glob_hyp (lfun,_) (loc,id) = + if List.mem id lfun then id + else +(* + try let _ = lookup (make_short_qualid id) in id + with Not_found -> +*) + Pretype_errors.error_var_not_found_loc loc id + +let glob_lochyp ist (loc,_ as locid) = (loc,glob_hyp ist locid) + +let error_unbound_metanum loc n = + user_err_loc + (loc,"glob_qualid_or_metanum", str "?" ++ int n ++ str " is unbound") + +let glob_metanum ist loc n = + if List.mem n (snd ist) then n else error_unbound_metanum loc n + +let glob_hyp_or_metanum ist = function + | AN (loc,id) -> AN (loc,glob_hyp ist (loc,id)) + | MetaNum (loc,n) -> MetaNum (loc,glob_metanum ist loc n) + +let glob_qualid_or_metanum ist = function + | AN (loc,qid) -> AN (loc,qualid_of_sp (sp_of_global (Global.env())(Nametab.global (loc,qid)))) + | MetaNum (loc,n) -> MetaNum (loc,glob_metanum ist loc n) + +let glob_reference ist (_,qid as locqid) = + let dir, id = repr_qualid qid in + try + if dir = empty_dirpath && List.mem id (fst ist) then qid + else raise Not_found + with Not_found -> + qualid_of_sp (sp_of_global (Global.env()) (Nametab.global locqid)) + +let glob_ltac_qualid ist (loc,qid as locqid) = + try qualid_of_sp (locate_tactic qid) + with Not_found -> glob_reference ist locqid + +let glob_ltac_reference ist = function + | RIdent (loc,id) -> + if List.mem id (fst ist) then RIdent (loc,id) + else RQualid (loc,glob_ltac_qualid ist (loc,make_short_qualid id)) + | RQualid qid -> RQualid (loc,glob_ltac_qualid ist qid) + +let rec glob_intro_pattern lf ist = function + | IntroOrAndPattern l -> + IntroOrAndPattern (List.map (List.map (glob_intro_pattern lf ist)) l) + | IntroWildcard -> + IntroWildcard + | IntroIdentifier id -> + IntroIdentifier (glob_ident lf ist id) + +let glob_quantified_hypothesis ist x = + (* We use identifier both for variables and quantified hyps (no way to + statically check the existence of a quantified hyp); thus nothing to do *) + x + +let glob_constr ist c = + let _ = Astterm.interp_rawconstr_gen Evd.empty (Global.env()) [] false (fst ist) c in + c + +(* Globalize bindings *) +let glob_binding ist (b,c) = + (glob_quantified_hypothesis ist b,glob_constr ist c) + +let glob_bindings ist = function + | NoBindings -> NoBindings + | ImplicitBindings l -> ImplicitBindings (List.map (glob_constr ist) l) + | ExplicitBindings l -> ExplicitBindings (List.map (glob_binding ist) l) + +let glob_constr_with_bindings ist (c,bl) = + (glob_constr ist c, glob_bindings ist bl) + +let glob_clause_pattern ist (l,occl) = + let rec check = function + | (hyp,l) :: rest -> + let (loc,_ as id) = skip_metaid hyp in + (AI(loc,glob_hyp ist id),l)::(check rest) + | [] -> [] + in (l,check occl) + +let glob_induction_arg ist = function + | ElimOnConstr c -> ElimOnConstr (glob_constr ist c) + | ElimOnAnonHyp n as x -> x + | ElimOnIdent (loc,id) as x -> x + +(* Globalize a reduction expression *) +let glob_evaluable_or_metanum ist = function + | AN (loc,qid) -> AN (loc,glob_reference ist (loc,qid)) + | MetaNum (loc,n) -> MetaNum (loc,glob_metanum ist loc n) + +let glob_unfold ist (l,qid) = (l,glob_evaluable_or_metanum ist qid) + +let glob_flag ist red = + { red with rConst = List.map (glob_evaluable_or_metanum ist) red.rConst } + +let glob_pattern ist (l,c) = (l,glob_constr ist c) + +let glob_redexp ist = function + | Unfold l -> Unfold (List.map (glob_unfold ist) l) + | Fold l -> Fold (List.map (glob_constr ist) l) + | Cbv f -> Cbv (glob_flag ist f) + | Lazy f -> Lazy (glob_flag ist f) + | Pattern l -> Pattern (List.map (glob_pattern ist) l) + | (Red _ | Simpl | Hnf as r) -> r + | ExtraRedExpr (s,l) -> ExtraRedExpr (s, List.map (glob_constr ist) l) + +(* Interprets an hypothesis name *) +let glob_hyp_location ist = function + | InHyp id -> + let (loc,_ as id) = skip_metaid id in + InHyp (AI(loc,glob_hyp ist id)) + | InHypType id -> + let (loc,_ as id) = skip_metaid id in + InHypType (AI(loc,glob_hyp ist id)) + +(* Reads a pattern *) +let glob_pattern evc env lfun = function + | Subterm (ido,pc) -> + let lfun = List.map (fun id -> (id, mkVar id)) lfun in + let (metas,_) = interp_constrpattern_gen evc env lfun pc in + metas, Subterm (ido,pc) + | Term pc -> + let lfun = List.map (fun id -> (id, mkVar id)) lfun in + let (metas,_) = interp_constrpattern_gen evc env lfun pc in + metas, Term pc + +let glob_constr_may_eval ist = function + | ConstrEval (r,c) -> ConstrEval (glob_redexp ist r,glob_constr ist c) + | ConstrContext (locid,c) -> + ConstrContext ((loc,glob_hyp ist locid),glob_constr ist c) + | ConstrTypeOf c -> ConstrTypeOf (glob_constr ist c) + | ConstrTerm c -> ConstrTerm (glob_constr ist c) + +(* Reads the hypotheses of a Match Context rule *) +let rec glob_match_context_hyps evc env lfun = function + | (NoHypId mp)::tl -> + let metas1, pat = glob_pattern evc env lfun mp in + let lfun, metas2, hyps = glob_match_context_hyps evc env lfun tl in + lfun, metas1@metas2, (NoHypId pat)::hyps + | (Hyp ((_,s) as locs,mp))::tl -> + let metas1, pat = glob_pattern evc env lfun mp in + let lfun, metas2, hyps = glob_match_context_hyps evc env lfun tl in + s::lfun, metas1@metas2, Hyp (locs,pat)::hyps + | [] -> lfun, [], [] + +(* Utilities *) +let rec filter_some = function + | None :: l -> filter_some l + | Some a :: l -> a :: filter_some l + | [] -> [] + +let extract_names lrc = + List.fold_right + (fun ((loc,name),_) l -> + if List.mem name l then + user_err_loc + (loc, "glob_tactic", str "This variable is bound several times"); + name::l) + lrc [] + +let extract_let_names lrc = + List.fold_right + (fun ((loc,name),_,_) l -> + if List.mem name l then + user_err_loc + (loc, "glob_tactic", str "This variable is bound several times"); + name::l) + lrc [] + +(* Globalizes tactics *) +let rec glob_atomic lf ist = function + (* Basic tactics *) + | TacIntroPattern l -> TacIntroPattern (List.map (glob_intro_pattern lf ist) l) + | TacIntrosUntil hyp -> TacIntrosUntil (glob_quantified_hypothesis ist hyp) + | TacIntroMove (ido,ido') -> + TacIntroMove (option_app (glob_ident lf ist) ido, + option_app (fun (loc,_ as x) -> (loc,glob_hyp ist x)) ido') + | TacAssumption -> TacAssumption + | TacExact c -> TacExact (glob_constr ist c) + | TacApply cb -> TacApply (glob_constr_with_bindings ist cb) + | TacElim (cb,cbo) -> + TacElim (glob_constr_with_bindings ist cb, + option_app (glob_constr_with_bindings ist) cbo) + | TacElimType c -> TacElimType (glob_constr ist c) + | TacCase cb -> TacCase (glob_constr_with_bindings ist cb) + | TacCaseType c -> TacCaseType (glob_constr ist c) + | TacFix (idopt,n) -> TacFix (option_app (glob_ident lf ist) idopt,n) + | TacMutualFix (id,n,l) -> + let f (id,n,c) = (glob_ident lf ist id,n,glob_constr ist c) in + TacMutualFix (glob_ident lf ist id, n, List.map f l) + | TacCofix idopt -> TacCofix (option_app (glob_ident lf ist) idopt) + | TacMutualCofix (id,l) -> + let f (id,c) = (glob_ident lf ist id,glob_constr ist c) in + TacMutualCofix (glob_ident lf ist id, List.map f l) + | TacCut c -> TacCut (glob_constr ist c) + | TacTrueCut (ido,c) -> + TacTrueCut (option_app (glob_ident lf ist) ido, glob_constr ist c) + | TacForward (b,na,c) -> TacForward (b,glob_name lf ist na,glob_constr ist c) + | TacGeneralize cl -> TacGeneralize (List.map (glob_constr ist) cl) + | TacGeneralizeDep c -> TacGeneralizeDep (glob_constr ist c) + | TacLetTac (id,c,clp) -> + TacLetTac (id,glob_constr ist c,glob_clause_pattern ist clp) + | TacInstantiate (n,c) -> TacInstantiate (n,glob_constr ist c) + + (* Automation tactics *) + | TacTrivial l -> TacTrivial l + | TacAuto (n,l) -> TacAuto (n,l) + | TacAutoTDB n -> TacAutoTDB n + | TacDestructHyp (b,(loc,_ as id)) -> TacDestructHyp(b,(loc,glob_hyp ist id)) + | TacDestructConcl -> TacDestructConcl + | TacSuperAuto (n,l,b1,b2) -> TacSuperAuto (n,l,b1,b2) + | TacDAuto (n,p) -> TacDAuto (n,p) + + (* Derived basic tactics *) + | TacOldInduction h -> TacOldInduction (glob_quantified_hypothesis ist h) + | TacNewInduction c -> TacNewInduction (glob_induction_arg ist c) + | TacOldDestruct h -> TacOldDestruct (glob_quantified_hypothesis ist h) + | TacNewDestruct c -> TacNewDestruct (glob_induction_arg ist c) + | TacDoubleInduction (h1,h2) -> + let h1 = glob_quantified_hypothesis ist h1 in + let h2 = glob_quantified_hypothesis ist h2 in + TacDoubleInduction (h1,h2) + | TacDecomposeAnd c -> TacDecomposeAnd (glob_constr ist c) + | TacDecomposeOr c -> TacDecomposeOr (glob_constr ist c) + | TacDecompose (l,c) -> + let l = List.map (glob_qualid_or_metanum ist) l in + TacDecompose (l,glob_constr ist c) + | TacSpecialize (n,l) -> TacSpecialize (n,glob_constr_with_bindings ist l) + | TacLApply c -> TacLApply (glob_constr ist c) + + (* Context management *) + | TacClear l -> TacClear (List.map (glob_hyp_or_metanum ist) l) + | TacClearBody l -> TacClearBody (List.map (glob_hyp_or_metanum ist) l) + | TacMove (dep,id1,id2) -> TacMove (dep,glob_lochyp ist id1,glob_lochyp ist id2) + | TacRename (id1,id2) -> TacRename (glob_lochyp ist id1, glob_lochyp ist id2) + + (* Constructors *) + | TacLeft bl -> TacLeft (glob_bindings ist bl) + | TacRight bl -> TacRight (glob_bindings ist bl) + | TacSplit bl -> TacSplit (glob_bindings ist bl) + | TacAnyConstructor t -> TacAnyConstructor (option_app (glob_tactic ist) t) + | TacConstructor (n,bl) -> TacConstructor (n, glob_bindings ist bl) + + (* Conversion *) + | TacReduce (r,cl) -> + TacReduce (glob_redexp ist r, List.map (glob_hyp_location ist) cl) + | TacChange (c,cl) -> + TacChange (glob_constr ist c, List.map (glob_hyp_location ist) cl) + + (* Equivalence relations *) + | TacReflexivity -> TacReflexivity + | TacSymmetry -> TacSymmetry + | TacTransitivity c -> TacTransitivity (glob_constr ist c) + + (* For extensions *) + | TacExtend (opn,l) -> + let _ = lookup_tactic opn in + TacExtend (opn,List.map (glob_genarg ist) l) + | TacAlias (_,l,body) -> failwith "TODO" + +and glob_tactic ist tac = snd (glob_tactic_seq ist tac) + +and glob_tactic_seq (lfun,lmeta as ist) = function + | TacAtom (loc,t) -> + let lf = ref lfun in + let t = glob_atomic lf ist t in + !lf, TacAtom (loc, t) + | TacFun tacfun -> lfun, TacFun (glob_tactic_fun ist tacfun) + | TacFunRec (name,tacfun) -> + lfun, TacFunRec (name,glob_tactic_fun ((snd name)::lfun,lmeta) tacfun) + | TacLetRecIn (lrc,u) -> + let names = extract_names lrc in + let ist = (names@lfun,lmeta) in + let lrc = List.map (fun (n,b) -> (n,glob_tactic_fun ist b)) lrc in + lfun, TacLetRecIn (lrc,glob_tactic ist u) + | TacLetIn (l,u) -> + let l = List.map (fun (n,c,b) -> (n,option_app (glob_constr_may_eval ist) c,glob_tacarg ist b)) l in + let ist' = ((extract_let_names l)@lfun,lmeta) in + lfun, TacLetIn (l,glob_tactic ist' u) + | TacLetCut l -> + let f (n,c,t) = (n,glob_constr_may_eval ist c,glob_tacarg ist t) in + lfun, TacLetCut (List.map f l) + | TacMatchContext (lr,lmr) -> + lfun, TacMatchContext(lr, glob_match_rule ist lmr) + | TacMatch (c,lmr) -> + lfun, TacMatch (glob_constr_may_eval ist c,glob_match_rule ist lmr) + | TacId -> lfun, TacId + | TacFail n as x -> lfun, x + | TacProgress tac -> lfun, TacProgress (glob_tactic ist tac) + | TacAbstract (tac,s) -> lfun, TacAbstract (glob_tactic ist tac,s) + | TacThen (t1,t2) -> + let lfun', t1 = glob_tactic_seq ist t1 in + let lfun'', t2 = glob_tactic_seq (lfun',lmeta) t2 in + lfun'', TacThen (t1,t2) + | TacThens (t,tl) -> + let lfun', t = glob_tactic_seq ist t in + (* Que faire en cas de (tac complexe avec Match et Thens; tac2) ?? *) + lfun', TacThens (t, List.map (glob_tactic (lfun',lmeta)) tl) + | TacDo (n,tac) -> lfun, TacDo (n,glob_tactic ist tac) + | TacTry tac -> lfun, TacTry (glob_tactic ist tac) + | TacInfo tac -> lfun, TacInfo (glob_tactic ist tac) + | TacRepeat tac -> lfun, TacRepeat (glob_tactic ist tac) + | TacOrelse (tac1,tac2) -> + lfun, TacOrelse (glob_tactic ist tac1,glob_tactic ist tac2) + | TacFirst l -> lfun, TacFirst (List.map (glob_tactic ist) l) + | TacSolve l -> lfun, TacSolve (List.map (glob_tactic ist) l) + | TacArg a -> lfun, TacArg (glob_tacarg ist a) + +and glob_tactic_fun (lfun,lmeta) (var,body) = + let lfun' = List.rev_append (filter_some var) lfun in + (var,glob_tactic (lfun',lmeta) body) + +and glob_tacarg ist = function + | TacVoid -> TacVoid + | Reference r -> Reference (glob_ltac_reference ist r) + | Integer n -> Integer n + | ConstrMayEval c -> ConstrMayEval (glob_constr_may_eval ist c) + | MetaNumArg (loc,n) -> MetaNumArg (loc,glob_metanum ist loc n) + | MetaIdArg (loc,_) -> error_syntactic_metavariables_not_allowed loc + | TacCall (loc,f,l) -> + TacCall (loc,glob_tacarg ist f,List.map (glob_tacarg ist) l) + | Tacexp t -> Tacexp (glob_tactic ist t) + | TacDynamic(_,t) as x -> + (match tag t with + | "tactic"|"value"|"constr" -> x + | s -> anomaly_loc (loc, "Tacinterp.val_interp", + str "Unknown dynamic: <" ++ str s ++ str ">")) + +(* Reads the rules of a Match Context or a Match *) +and glob_match_rule (lfun,lmeta as ist) = function + | (All tc)::tl -> + (All (glob_tactic ist tc))::(glob_match_rule ist tl) + | (Pat (rl,mp,tc))::tl -> + let env = Global.env() in + let lfun',metas1,hyps = glob_match_context_hyps Evd.empty env lfun rl in + let metas2,pat = glob_pattern Evd.empty env lfun mp in + let metas = list_uniquize (metas1@metas2@lmeta) in + (Pat (hyps,pat,glob_tactic (lfun',metas) tc))::(glob_match_rule ist tl) + | [] -> [] + +and glob_genarg ist x = + match genarg_tag x with + | BoolArgType -> in_gen rawwit_bool (out_gen rawwit_bool x) + | IntArgType -> in_gen rawwit_int (out_gen rawwit_int x) + | IntOrVarArgType -> + let f = function + | ArgVar (loc,id) -> ArgVar (loc,glob_hyp ist (loc,id)) + | ArgArg n as x -> x in + in_gen rawwit_int_or_var (f (out_gen rawwit_int_or_var x)) + | StringArgType -> + in_gen rawwit_string (out_gen rawwit_string x) + | PreIdentArgType -> + in_gen rawwit_pre_ident (out_gen rawwit_pre_ident x) + | IdentArgType -> + in_gen rawwit_ident (glob_hyp ist (dummy_loc,out_gen rawwit_ident x)) + | QualidArgType -> + let (loc,qid) = out_gen rawwit_qualid x in + in_gen rawwit_qualid (loc,glob_ltac_qualid ist (loc,qid)) + | ConstrArgType -> + in_gen rawwit_constr (glob_constr ist (out_gen rawwit_constr x)) + | ConstrMayEvalArgType -> + in_gen rawwit_constr_may_eval (glob_constr_may_eval ist (out_gen rawwit_constr_may_eval x)) + | QuantHypArgType -> + in_gen rawwit_quant_hyp + (glob_quantified_hypothesis ist (out_gen rawwit_quant_hyp x)) + | RedExprArgType -> + in_gen rawwit_red_expr (glob_redexp ist (out_gen rawwit_red_expr x)) + | TacticArgType -> + in_gen rawwit_tactic (glob_tactic ist (out_gen rawwit_tactic x)) + | CastedOpenConstrArgType -> + in_gen rawwit_casted_open_constr + (glob_constr ist (out_gen rawwit_casted_open_constr x)) + | ConstrWithBindingsArgType -> + in_gen rawwit_constr_with_bindings + (glob_constr_with_bindings ist (out_gen rawwit_constr_with_bindings x)) + | List0ArgType _ -> app_list0 (glob_genarg ist) x + | List1ArgType _ -> app_list1 (glob_genarg ist) x + | OptArgType _ -> app_opt (glob_genarg ist) x + | PairArgType _ -> app_pair (glob_genarg ist) (glob_genarg ist) x + | ExtraArgType s -> x + + +(************* END GLOBALIZE ************) + +(* Reads the head of Fun *) +let read_fun ast = + let rec read_fun_rec = function + | Node(_,"VOID",[])::tl -> None::(read_fun_rec tl) + | Nvar(_,s)::tl -> (Some s)::(read_fun_rec tl) + | [] -> [] + | _ -> + anomalylabstrm "Tacinterp.read_fun_rec" (str "Fun not well formed") + in + match ast with + | Node(_,"FUNVAR",l) -> read_fun_rec l + | _ -> + anomalylabstrm "Tacinterp.read_fun" (str "Fun not well formed") + +(* Reads the clauses of a Rec *) +let rec read_rec_clauses = function + | [] -> [] + | Node(_,"RECCLAUSE",[Nvar(_,name);it;body])::tl -> + (name,it,body)::(read_rec_clauses tl) + |_ -> + anomalylabstrm "Tacinterp.read_rec_clauses" + (str "Rec not well formed") + +(* Associates variables with values and gives the remaining variables and + values *) +let head_with_value (lvar,lval) = + let rec head_with_value_rec lacc = function + | ([],[]) -> (lacc,[],[]) + | (vr::tvr,ve::tve) -> + (match vr with + | None -> head_with_value_rec lacc (tvr,tve) + | Some v -> head_with_value_rec ((v,ve)::lacc) (tvr,tve)) + | (vr,[]) -> (lacc,vr,[]) + | ([],ve) -> (lacc,[],ve) + in + head_with_value_rec [] (lvar,lval) + +(* Gives a context couple if there is a context identifier *) +let give_context ctxt = function + | None -> [] + | Some id -> [id,VConstr_context ctxt] + +(* Reads a pattern *) +let read_pattern evc env lfun = function + | Subterm (ido,pc) -> + Subterm (ido,snd (interp_constrpattern_gen evc env lfun pc)) + | Term pc -> + Term (snd (interp_constrpattern_gen evc env lfun pc)) + +(* Reads the hypotheses of a Match Context rule *) +let rec read_match_context_hyps evc env lfun lidh = function + | (NoHypId mp)::tl -> + (NoHypId (read_pattern evc env lfun mp)):: + (read_match_context_hyps evc env lfun lidh tl) + | (Hyp ((loc,id) as locid,mp))::tl -> + if List.mem id lidh then + user_err_loc (loc,"Tacinterp.read_match_context_hyps", + str ("Hypothesis pattern-matching variable "^(string_of_id id)^ + " used twice in the same pattern")) + else + (Hyp (locid,read_pattern evc env lfun mp)):: + (read_match_context_hyps evc env lfun (id::lidh) tl) + | [] -> [] + +(* Reads the rules of a Match Context or a Match *) +let rec read_match_rule evc env lfun = function + | (All tc)::tl -> (All tc)::(read_match_rule evc env lfun tl) + | (Pat (rl,mp,tc))::tl -> + (Pat (read_match_context_hyps evc env lfun [] rl, + read_pattern evc env lfun mp,tc)) + ::(read_match_rule evc env lfun tl) + | [] -> [] + +(* For Match Context and Match *) +exception No_match +exception Not_coherent_metas + +let is_match_catchable = function + | No_match | FailError _ -> true + | e -> Logic.catchable_exception e + +(* Verifies if the matched list is coherent with respect to lcm *) +let rec verify_metas_coherence ist lcm = function + | (num,csr)::tl -> + if (List.for_all + (fun (a,b) -> + if a=num then + Reductionops.is_conv ist.env ist.evc b csr + else + true) lcm) then + (num,csr)::(verify_metas_coherence ist lcm tl) + else + raise Not_coherent_metas + | [] -> [] + +(* Tries to match a pattern and a constr *) +let apply_matching pat csr = + try + (Pattern.matches pat csr) + with + PatternMatchingFailure -> raise No_match + +(* Tries to match one hypothesis pattern with a list of hypotheses *) +let apply_one_mhyp_context ist lmatch mhyp lhyps noccopt = + let get_pattern = function + | Hyp (_,pat) -> pat + | NoHypId pat -> pat + and get_id_couple id = function + | Hyp((_,idpat),_) -> [idpat,VIdentifier id] + | NoHypId _ -> [] in + let rec apply_one_mhyp_context_rec mhyp lhyps_acc nocc = function + | (id,hyp)::tl -> + (match (get_pattern mhyp) with + | Term t -> + (try + let lmeta = + verify_metas_coherence ist lmatch (Pattern.matches t hyp) in + (get_id_couple id mhyp,[],lmeta,tl,(id,hyp),None) + with | PatternMatchingFailure | Not_coherent_metas -> + apply_one_mhyp_context_rec mhyp (lhyps_acc@[id,hyp]) 0 tl) + | Subterm (ic,t) -> + (try + let (lm,ctxt) = sub_match nocc t hyp in + let lmeta = verify_metas_coherence ist lmatch lm in + (get_id_couple id mhyp,give_context ctxt + ic,lmeta,tl,(id,hyp),Some nocc) + with + | NextOccurrence _ -> + apply_one_mhyp_context_rec mhyp (lhyps_acc@[id,hyp]) 0 tl + | Not_coherent_metas -> + apply_one_mhyp_context_rec mhyp lhyps_acc (nocc + 1) ((id,hyp)::tl))) + | [] -> raise No_match in + let nocc = + match noccopt with + | None -> 0 + | Some n -> n in + apply_one_mhyp_context_rec mhyp [] nocc lhyps + +(* +let coerce_to_qualid loc = function + | Constr c when isVar c -> make_short_qualid (destVar c) + | Constr c -> + (try qualid_of_sp (sp_of_global (Global.env()) (reference_of_constr c)) + with Not_found -> invalid_arg_loc (loc, "Not a reference")) + | Identifier id -> make_short_qualid id + | Qualid qid -> qid + | _ -> invalid_arg_loc (loc, "Not a reference") +*) + +let constr_to_id loc c = + if isVar c then destVar c + else invalid_arg_loc (loc, "Not an identifier") + +let constr_to_qid loc c = + try qualid_of_sp (sp_of_global (Global.env ()) (reference_of_constr c)) + with _ -> invalid_arg_loc (loc, "Not a global reference") + +(* Check for LetTac *) +let check_clause_pattern ist (l,occl) = + match ist.goalopt with + | Some gl -> + let sign = pf_hyps gl in + let rec check acc = function + | (hyp,l) :: rest -> + let _,hyp = skip_metaid hyp in + if List.mem hyp acc then + error ("Hypothesis "^(string_of_id hyp)^" occurs twice"); + if not (mem_named_context hyp sign) then + error ("No such hypothesis: " ^ (string_of_id hyp)); + (hyp,l)::(check (hyp::acc) rest) + | [] -> [] + in (l,check [] occl) + | None -> error "No goal" + +(* Debug reference *) +let debug = ref DebugOff + +(* Sets the debugger mode *) +let set_debug pos = debug := pos + +(* Gives the state of debug *) +let get_debug () = !debug + +(* Interprets an identifier *) +let eval_ident ist id = + try (unrec (List.assoc id ist.lfun)) + with | Not_found -> +(* + try (vcontext_interp ist (lookup (make_short_qualid id))) + with | Not_found -> +*) +VIdentifier id + +(* Gives the identifier corresponding to an Identifier tactic_arg *) +let id_of_Identifier = function + | VConstr c when isVar c -> destVar c + | VIdentifier id -> id + | _ -> + anomalylabstrm "id_of_Identifier" (str "Not an IDENTIFIER tactic_arg") + +let coerce_to_hypothesis ist = function + | VConstr c when isVar c -> destVar c + | VIdentifier id -> id + | v -> errorlabstrm "coerce_to_hypothesis" + (str "Cannot coerce" ++ spc () ++ pr_value ist.env v ++ spc () ++ + str "to an existing hypothesis") + +let ident_interp ist id = + id_of_Identifier (eval_ident ist id) + +let hyp_interp ist (loc,id) = + coerce_to_hypothesis ist (eval_ident ist id) + +let name_interp ist = function + | Anonymous -> Anonymous + | Name id -> Name (ident_interp ist id) + +let hyp_or_metanum_interp ist = function + | AN (loc,id) -> ident_interp ist id + | MetaNum (loc,n) -> constr_to_id loc (List.assoc n ist.lmatch) + +(* To avoid to move to much simple functions in the big recursive block *) +let forward_vcontext_interp = ref (fun ist v -> failwith "not implemented") + +let interp_pure_qualid ist (loc,qid) = + try (!forward_vcontext_interp ist (lookup qid)) + with | Not_found -> + try VConstr (constr_of_reference (find_reference ist qid)) + with Not_found -> + let (dir,id) = repr_qualid qid in + if dir = empty_dirpath then VIdentifier id + else user_err_loc (loc,"interp_pure_qualid",str "Unknown reference") + +let interp_ltac_reference ist = function + | RIdent (loc,id) -> + (try unrec (List.assoc id ist.lfun) + with | Not_found -> interp_pure_qualid ist (loc,make_short_qualid id)) + | RQualid qid -> interp_pure_qualid ist qid + +(* Interprets a qualified name *) +let eval_qualid ist (loc,qid as locqid) = + let dir, id = repr_qualid qid in + try + if dir = empty_dirpath then unrec (List.assoc id ist.lfun) + else raise Not_found + with | Not_found -> + interp_pure_qualid ist locqid + +let qualid_interp ist qid = + let v = eval_qualid ist qid in + coerce_to_reference ist v + +(* Interprets a qualified name. This can be a metavariable to be injected *) +let qualid_or_metanum_interp ist = function + | AN (loc,qid) -> qid + | MetaNum (loc,n) -> constr_to_qid loc (List.assoc n ist.lmatch) + +let eval_ref_or_metanum ist = function + | AN (loc,qid) -> eval_qualid ist (loc,qid) + | MetaNum (loc,n) -> VConstr (List.assoc n ist.lmatch) + +let interp_evaluable_or_metanum ist c = + let c = eval_ref_or_metanum ist c in + coerce_to_evaluable_ref ist.env c + +let interp_inductive_or_metanum ist c = + let c = eval_ref_or_metanum ist c in + coerce_to_inductive c + +(* Interprets an hypothesis name *) +let interp_hyp_location ist = function + | InHyp id -> InHyp (hyp_interp ist (skip_metaid id)) + | InHypType id -> InHypType (hyp_interp ist (skip_metaid id)) + +let id_opt_interp ist = option_app (ident_interp ist) + +(* Interpretation of constructions *) + + (* Extracted the constr list from lfun *) +let rec constr_list_aux ist = function + | (id,VConstr c)::tl -> (id,c)::(constr_list_aux ist tl) + | (id0,VIdentifier id)::tl -> + (try (id0,(constr_of_id ist id))::(constr_list_aux ist tl) + with | Not_found -> (constr_list_aux ist tl)) + | _::tl -> constr_list_aux ist tl + | [] -> [] + +let constr_list ist = constr_list_aux ist ist.lfun +let interp_constr ocl ist c = + interp_constr_gen ist.evc ist.env (constr_list ist) ist.lmatch c ocl + +let interp_openconstr ist c ocl = + interp_openconstr_gen ist.evc ist.env (constr_list ist) ist.lmatch c ocl + +(* Interprets a constr expression *) +let constr_interp ist c = + let csr = interp_constr None ist c in + begin + db_constr ist.debug ist.env csr; + csr + end + +(* Interprets a constr expression casted by the current goal *) +let cast_constr_interp ist c = + match ist.goalopt with + | Some gl -> + let csr = interp_constr (Some (pf_concl gl)) ist c in + begin + db_constr ist.debug ist.env csr; + csr + end + + | None -> + errorlabstrm "cast_constr_interp" + (str "Cannot cast a constr without goal") + +(* Interprets an open constr expression casted by the current goal *) +let cast_openconstr_interp ist c = + match ist.goalopt with + | Some gl -> interp_openconstr ist c (Some (pf_concl gl)) + | None -> + errorlabstrm "cast_openconstr_interp" + (str "Cannot cast a constr without goal") + +(* Interprets a reduction expression *) +let unfold_interp ist (l,qid) = (l,interp_evaluable_or_metanum ist qid) + +let flag_interp ist red = + { red with rConst = List.map (interp_evaluable_or_metanum ist) red.rConst } + +let pattern_interp ist (l,c) = (l,constr_interp ist c) + +let redexp_interp ist = function + | Unfold l -> Unfold (List.map (unfold_interp ist) l) + | Fold l -> Fold (List.map (constr_interp ist) l) + | Cbv f -> Cbv (flag_interp ist f) + | Lazy f -> Lazy (flag_interp ist f) + | Pattern l -> Pattern (List.map (pattern_interp ist) l) + | (Red _ | Simpl | Hnf as r) -> r + | ExtraRedExpr (s,l) -> ExtraRedExpr (s,List.map (constr_interp ist) l) + +let interp_may_eval f ist = function + | ConstrEval (r,c) -> + let redexp = redexp_interp ist r in + (reduction_of_redexp redexp) ist.env ist.evc (f ist c) + | ConstrContext ((loc,s),c) -> + (try + let ic = f ist c + and ctxt = constr_of_VConstr_context (List.assoc s ist.lfun) in + subst_meta [-1,ic] ctxt + with + | Not_found -> + user_err_loc (loc, "constr_interp", + str "Unbound context identifier" ++ pr_id s)) + | ConstrTypeOf c -> type_of ist.env ist.evc (f ist c) + | ConstrTerm c -> f ist c + +(* Interprets a constr expression possibly to first evaluate *) +let constr_interp_may_eval ist c = + let csr = interp_may_eval (interp_constr None) ist c in + begin + db_constr ist.debug ist.env csr; + csr + end + +let rec interp_intro_pattern ist = function + | IntroOrAndPattern l -> + IntroOrAndPattern (List.map (List.map (interp_intro_pattern ist)) l) + | IntroWildcard -> + IntroWildcard + | IntroIdentifier id -> + IntroIdentifier (ident_interp ist id) + +let interp_quantified_hypothesis ist = function + | AnonHyp n -> AnonHyp n + | NamedHyp id -> + match eval_ident ist id with + | VIdentifier id -> NamedHyp id + | VInteger n -> AnonHyp n + | _ -> invalid_arg_loc + (loc, "Not a reference to an hypothesis: "^(string_of_id id)) + + +let interp_induction_arg ist = function + | ElimOnConstr c -> ElimOnConstr (constr_interp ist c) + | ElimOnAnonHyp n as x -> x + | ElimOnIdent (loc,id) -> + match ist.goalopt with + | None -> error "No goal" + | Some gl -> + if Tactics.is_quantified_hypothesis id gl then ElimOnIdent (loc,id) + else ElimOnConstr + (constr_interp ist (Termast.ast_of_qualid (make_short_qualid id))) + +let binding_interp ist (b,c) = + (interp_quantified_hypothesis ist b,constr_interp ist c) + +let bindings_interp ist = function + | NoBindings -> NoBindings + | ImplicitBindings l -> ImplicitBindings (List.map (constr_interp ist) l) + | ExplicitBindings l -> ExplicitBindings (List.map (binding_interp ist) l) + +let interp_constr_with_bindings ist (c,bl) = + (constr_interp ist c, bindings_interp ist bl) + +(* Interprets a tactic expression *) +let rec val_interp ist ast = + + let value_interp ist = + match ast with + (* Immediate evaluation *) + | TacFun (it,body) -> VFun (ist.lfun,it,body) + | TacFunRec rc -> funrec_interp ist rc + | TacLetRecIn (lrc,u) -> letrec_interp ist lrc u + | TacLetIn (l,u) -> + let addlfun=letin_interp ist l in + val_interp { ist with lfun=addlfun@ist.lfun } u + | TacMatchContext (lr,lmr) -> + (match ist.goalopt with + | None -> VContext (ist,lr,lmr) + | Some g -> match_context_interp ist lr lmr g) + | TacMatch (c,lmr) -> match_interp ist c lmr + | TacArg a -> tacarg_interp ist a + (* Delayed evaluation *) + | t -> VClosure (ist,t) + in + if ist.debug = DebugOn then + match debug_prompt ist.goalopt ast with + | Exit -> VClosure (ist,TacId) + | v -> value_interp {ist with debug=v} + else + value_interp ist + +and eval_tactic ist = function + | TacAtom (loc,t) -> + (try interp_atomic ist t + with e -> Stdpp.raise_with_loc loc e) + | TacFun (it,body) -> assert false + | TacFunRec rc -> assert false + | TacLetRecIn (lrc,u) -> assert false + | TacLetIn (l,u) -> assert false + | TacLetCut l -> letcut_interp ist l + | TacMatchContext _ -> assert false + | TacMatch (c,lmr) -> assert false + | TacId -> tclIDTAC + | TacFail n -> tclFAIL n + | TacProgress tac -> tclPROGRESS (tactic_interp ist tac) + | TacAbstract (tac,s) -> Tactics.tclABSTRACT s (tactic_interp ist tac) + | TacThen (t1,t2) -> tclTHEN (tactic_interp ist t1) (tactic_interp ist t2) + | TacThens (t,tl) -> + tclTHENS (tactic_interp ist t) (List.map (tactic_interp ist) tl) + | TacDo (n,tac) -> tclDO n (tactic_interp ist tac) + | TacTry tac -> tclTRY (tactic_interp ist tac) + | TacInfo tac -> tclINFO (tactic_interp ist tac) + | TacRepeat tac -> tclREPEAT (tactic_interp ist tac) + | TacOrelse (tac1,tac2) -> + tclORELSE (tactic_interp ist tac1) (tactic_interp ist tac2) + | TacFirst l -> tclFIRST (List.map (tactic_interp ist) l) + | TacSolve l -> tclSOLVE (List.map (tactic_interp ist) l) +(* Obsolete ?? + | Node(loc0,"APPTACTIC",[Node(loc1,s,l)]) -> + (Node(loc0,"APP",[Node(loc1,"PRIM-TACTIC",[Node(loc1,s,[])])]@l)) + | Node(_,"PRIMTACTIC",[Node(loc,opn,[])]) -> + VFTactic ([],(interp_atomic opn)) +*) + | TacArg a -> assert false + +and tacarg_interp ist = function + | TacVoid -> VVoid + | Reference r -> interp_ltac_reference ist r + | Integer n -> VInteger n + | ConstrMayEval c -> VConstr (constr_interp_may_eval ist c) + | MetaNumArg (_,n) -> VConstr (List.assoc n ist.lmatch) + | MetaIdArg (loc,_) -> error_syntactic_metavariables_not_allowed loc +(* + | Tacexp t -> VArg (Tacexp ((*tactic_interp ist t,*)t)) +*) + | TacCall (loc,f,l) -> + let fv = tacarg_interp ist f + and largs = List.map (tacarg_interp ist) l in + app_interp ist fv largs loc + | Tacexp t -> val_interp ist t +(* + | Node(loc,s,l) -> + let fv = val_interp ist (Node(loc,"PRIMTACTIC",[Node(loc,s,[])])) + and largs = List.map (val_interp ist) l in + app_interp ist fv largs ast +*) + | TacDynamic(_,t) -> + let tg = (tag t) in + if tg = "tactic" then + let f = (tactic_out t) in val_interp ist (f ist) + else if tg = "value" then + value_out t + else if tg = "constr" then + VConstr (Pretyping.constr_out t) + else + anomaly_loc (loc, "Tacinterp.val_interp", + (str "Unknown dynamic: <" ++ str (Dyn.tag t) ++ str ">")) + +(* Interprets an application node *) +and app_interp ist fv largs loc = + match fv with + | VFTactic(l,f) -> VFTactic(l@largs,f) + | VFun(olfun,var,body) -> + let (newlfun,lvar,lval)=head_with_value (var,largs) in + if lvar=[] then + if lval=[] then + val_interp { ist with lfun=newlfun@olfun } body + else + app_interp ist + (val_interp {ist with lfun=newlfun@olfun } body) lval loc + else + VFun(newlfun@olfun,lvar,body) + | _ -> + user_err_loc (loc, "Tacinterp.app_interp", + (str"Illegal tactic application")) + +(* Gives the tactic corresponding to the tactic value *) +and tactic_of_value vle g = + match vle with + | VClosure (ist,tac) -> eval_tactic ist tac g + | VFTactic (largs,f) -> (f largs g) + | VRTactic res -> res + | VTactic tac -> tac g + | _ -> raise NotTactic + +(* Evaluation with FailError catching *) +and eval_with_fail interp ast goal = + try + (match interp ast with + | VClosure (ist,tac) -> VRTactic (eval_tactic ist tac goal) + | VFTactic (largs,f) -> VRTactic (f largs goal) + | VTactic tac -> VRTactic (tac goal) + | a -> a) + with | FailError lvl -> + if lvl = 0 then + raise No_match + else + raise (FailError (lvl - 1)) + +(* Interprets recursive expressions *) +and funrec_interp ist ((loc,name),(var,body)) = + let ve = ref VVoid in + let newve = VFun((name,VRec ve)::ist.lfun,var,body) in + begin + ve:=newve; + !ve + end + +and letrec_interp ist lrc u = + let lref = Array.to_list (Array.make (List.length lrc) (ref VVoid)) in + let lenv = List.fold_right2 (fun ((loc,name),_) vref l -> (name,VRec vref)::l) + lrc lref [] in + let lve = List.map (fun ((loc,name),(var,body)) -> + (name,VFun(lenv@ist.lfun,var,body))) lrc in + begin + List.iter2 (fun vref (_,ve) -> vref:=ve) lref lve; + val_interp { ist with lfun=lve@ist.lfun } u + end + +(* Interprets the clauses of a LetCutIn *) +and letin_interp ist = function + | [] -> [] + | ((loc,id),None,t)::tl -> (id,tacarg_interp ist t):: (letin_interp ist tl) + | ((loc,id),Some com,tce)::tl -> + let typ = interp_may_eval (interp_constr None) ist com + and tac = tacarg_interp ist tce in + match tac with + | VConstr csr -> + (id,VConstr (mkCast (csr,typ)))::(letin_interp ist tl) + | VIdentifier id -> + (try + (id,VConstr (mkCast (constr_of_id ist id,typ))):: + (letin_interp ist tl) + with | Not_found -> + errorlabstrm "Tacinterp.letin_interp" + (str "Term or tactic expected")) + | _ -> + (try + let t = tactic_of_value tac in + let ndc = + (match ist.goalopt with + | None -> Global.named_context () + | Some g -> pf_hyps g) in + start_proof id (true,NeverDischarge) ndc typ (fun _ _ -> ()); + by t; + let (_,({const_entry_body = pft; const_entry_type = _},_,_)) = + cook_proof () in + delete_proof id; + (id,VConstr (mkCast (pft,typ)))::(letin_interp ist tl) + with | NotTactic -> + delete_proof id; + errorlabstrm "Tacinterp.letin_interp" + (str "Term or tactic expected")) + +(* Interprets the clauses of a LetCut *) +and letcut_interp ist = function + | [] -> tclIDTAC + | (id,com,tce)::tl -> + let typ = constr_interp_may_eval ist com + and tac = tacarg_interp ist tce + and (ndc,ccl) = + match ist.goalopt with + | None -> + errorlabstrm "Tacinterp.letcut_interp" (str + "Do not use Let for toplevel definitions, use Lemma, ... instead") + | Some g -> (pf_hyps g,pf_concl g) in + (match tac with + | VConstr csr -> + let cutt = h_cut typ + and exat = h_exact csr in + tclTHENSV cutt [|tclTHEN (introduction id) + (letcut_interp ist tl);exat|] + +(* let lic = mkLetIn (Name id,csr,typ,ccl) in + let ntac = refine (mkCast (mkMeta (Logic.new_meta ()),lic)) in + tclTHEN ntac (tclTHEN (introduction id) + (letcut_interp ist tl))*) + + | VIdentifier ir -> + (try + let cutt = h_cut typ + and exat = h_exact (constr_of_id ist ir) in + tclTHENSV cutt [| tclTHEN (introduction id) + (letcut_interp ist tl); exat |] + with | Not_found -> + errorlabstrm "Tacinterp.letin_interp" + (str "Term or tactic expected")) + | _ -> + (try + let t = tactic_of_value tac in + start_proof id (false,NeverDischarge) ndc typ (fun _ _ -> ()); + by t; + let (_,({const_entry_body = pft; const_entry_type = _},_,_)) = + cook_proof () in + delete_proof id; + let cutt = h_cut typ + and exat = h_exact pft in + tclTHENSV cutt [|tclTHEN (introduction id) + (letcut_interp ist tl);exat|] + +(* let lic = mkLetIn (Name id,pft,typ,ccl) in + let ntac = refine (mkCast (mkMeta (Logic.new_meta ()),lic)) in + tclTHEN ntac (tclTHEN (introduction id) + (letcut_interp ist tl))*) + with | NotTactic -> + delete_proof id; + errorlabstrm "Tacinterp.letcut_interp" + (str "Term or tactic expected"))) + +(* Interprets the Match Context expressions *) +and match_context_interp ist lr lmr g = +(* let goal = + (match goalopt with + | None -> + errorlabstrm "Tacinterp.apply_match_context" (str + "No goal available") + | Some g -> g) in*) + let rec apply_goal_sub ist goal nocc (id,c) csr mt mhyps hyps = + try + let (lgoal,ctxt) = sub_match nocc c csr in + let lctxt = give_context ctxt id in + if mhyps = [] then + eval_with_fail + (val_interp + { ist with lfun=lctxt@ist.lfun; lmatch=lgoal@ist.lmatch; + goalopt=Some goal}) + mt goal + else + apply_hyps_context {ist with goalopt=Some goal} mt lgoal mhyps hyps + with + | (FailError _) as e -> raise e + | NextOccurrence _ -> raise No_match + | No_match | _ -> + apply_goal_sub ist goal (nocc + 1) (id,c) csr mt mhyps hyps in + let rec apply_match_context ist goal = function + | (All t)::tl -> + (try + eval_with_fail (val_interp {ist with goalopt=Some goal }) t + goal + with No_match | FailError _ -> apply_match_context ist goal tl + | e when Logic.catchable_exception e -> apply_match_context ist goal tl) + | (Pat (mhyps,mgoal,mt))::tl -> + let hyps = make_hyps (pf_hyps goal) in + let hyps = if lr then List.rev hyps else hyps in + let concl = pf_concl goal in + (match mgoal with + | Term mg -> + (try + (let lgoal = apply_matching mg concl in + begin + db_matched_concl ist.debug ist.env concl; + if mhyps = [] then + eval_with_fail (val_interp + {ist with lmatch=lgoal@ist.lmatch; goalopt=Some goal}) mt goal + else + apply_hyps_context { ist with goalopt=Some goal} mt lgoal mhyps + hyps + end) + with e when is_match_catchable e -> apply_match_context ist goal tl) + | Subterm (id,mg) -> + (try + apply_goal_sub ist goal 0 (id,mg) concl mt mhyps hyps + with e when is_match_catchable e -> + apply_match_context ist goal tl)) + | _ -> + errorlabstrm "Tacinterp.apply_match_context" (str + "No matching clauses for Match Context") + + in + apply_match_context ist g + (read_match_rule ist.evc ist.env (constr_list ist) lmr) + +(* Interprets a VContext value *) +and vcontext_interp ist = function + | (VContext (ist',lr,lmr)) as v -> + (match ist.goalopt with + | None -> v + | Some g -> (* Relaunch *) match_context_interp ist' lr lmr g) + | v -> v + +(* Tries to match the hypotheses in a Match Context *) +and apply_hyps_context ist mt lgmatch mhyps hyps = + let rec apply_hyps_context_rec ist mt lfun lmatch mhyps lhyps_mhyp + lhyps_rest noccopt = + let goal = match ist.goalopt with Some g -> g | _ -> assert false in + match mhyps with + | hd::tl -> + let (lid,lc,lm,newlhyps,hyp_match,noccopt) = + apply_one_mhyp_context ist lmatch hd lhyps_mhyp noccopt in + begin + db_matched_hyp ist.debug ist.env hyp_match; + (try + if tl = [] then + eval_with_fail + (val_interp {ist with lfun=lfun@lid@lc@ist.lfun; + lmatch=lmatch@lm@ist.lmatch; + goalopt=Some goal}) + mt goal + else + let nextlhyps = + List.fold_left (fun l e -> if e = hyp_match then l else l@[e]) [] + lhyps_rest in + apply_hyps_context_rec ist mt + (lfun@lid@lc) (lmatch@lm) tl nextlhyps nextlhyps None + with + | (FailError _) as e -> raise e + | e when is_match_catchable e -> + (match noccopt with + | None -> + apply_hyps_context_rec ist mt lfun + lmatch mhyps newlhyps lhyps_rest None + | Some nocc -> + apply_hyps_context_rec ist mt ist.lfun ist.lmatch mhyps + (hyp_match::newlhyps) lhyps_rest (Some (nocc + 1)))) + end + | [] -> + anomalylabstrm "apply_hyps_context_rec" (str + "Empty list should not occur") in + apply_hyps_context_rec ist mt [] lgmatch mhyps hyps hyps None + + (* Interprets extended tactic generic arguments *) +and genarg_interp ist x = + match genarg_tag x with + | BoolArgType -> in_gen wit_bool (out_gen rawwit_bool x) + | IntArgType -> in_gen wit_int (out_gen rawwit_int x) + | IntOrVarArgType -> + let f = function + | ArgVar (loc,id) -> + (match eval_ident ist id with + | VInteger n -> ArgArg n + | _ -> + user_err_loc + (loc,"genarg_interp",str "should be bound to an integer")) + | ArgArg n as x -> x in + in_gen wit_int_or_var (f (out_gen rawwit_int_or_var x)) + | StringArgType -> + in_gen wit_string (out_gen rawwit_string x) + | PreIdentArgType -> + in_gen wit_pre_ident (out_gen rawwit_pre_ident x) + | IdentArgType -> + in_gen wit_ident (ident_interp ist (out_gen rawwit_ident x)) + | QualidArgType -> + in_gen wit_qualid (qualid_interp ist (out_gen rawwit_qualid x)) + | ConstrArgType -> + in_gen wit_constr (constr_interp ist (out_gen rawwit_constr x)) + | ConstrMayEvalArgType -> + in_gen wit_constr_may_eval (constr_interp_may_eval ist (out_gen rawwit_constr_may_eval x)) + | QuantHypArgType -> + in_gen wit_quant_hyp + (interp_quantified_hypothesis ist (out_gen rawwit_quant_hyp x)) + | RedExprArgType -> + in_gen wit_red_expr (redexp_interp ist (out_gen rawwit_red_expr x)) + | TacticArgType -> + in_gen wit_tactic ((*tactic_interp ist*) (out_gen rawwit_tactic x)) + | CastedOpenConstrArgType -> + in_gen wit_casted_open_constr + (cast_openconstr_interp ist (out_gen rawwit_casted_open_constr x)) + | ConstrWithBindingsArgType -> + in_gen wit_constr_with_bindings + (interp_constr_with_bindings ist (out_gen rawwit_constr_with_bindings x)) + | List0ArgType _ -> app_list0 (genarg_interp ist) x + | List1ArgType _ -> app_list1 (genarg_interp ist) x + | OptArgType _ -> app_opt (genarg_interp ist) x + | PairArgType _ -> app_pair (genarg_interp ist) (genarg_interp ist) x + | ExtraArgType s -> lookup_genarg_interp s ist x + +(* Interprets the Match expressions *) +and match_interp ist constr lmr = + let rec apply_sub_match ist nocc (id,c) csr + mt = + match ist.goalopt with + | None -> + (try + let (lm,ctxt) = sub_match nocc c csr in + let lctxt = give_context ctxt id in + val_interp {ist with lfun=lctxt@ist.lfun; lmatch=lm@ist.lmatch} mt + with | NextOccurrence _ -> raise No_match) + | Some g -> + (try + let (lm,ctxt) = sub_match nocc c csr in + let lctxt = give_context ctxt id in + eval_with_fail (val_interp { ist with lfun=lctxt@ist.lfun; + lmatch=lm@ist.lmatch}) + mt g + with + | NextOccurrence n -> raise No_match + | (FailError _) as e -> raise e + | e when is_match_catchable e -> + apply_sub_match ist (nocc + 1) (id,c) csr mt) + in + let rec apply_match ist csr = function + | (All t)::_ -> + (match ist.goalopt with + | None -> + (try val_interp ist t + with e when is_match_catchable e -> apply_match ist csr []) + | Some g -> + (try + eval_with_fail (val_interp ist) t g + with + | (FailError _) as e -> raise e + | e when is_match_catchable e -> + apply_match ist csr [])) + | (Pat ([],mp,mt))::tl -> + (match mp with + | Term c -> + (match ist.goalopt with + | None -> + (try + val_interp + { ist with lmatch=(apply_matching c csr)@ist.lmatch } mt + with e when is_match_catchable e -> apply_match ist csr tl) + | Some g -> + (try + eval_with_fail (val_interp + { ist with lmatch=(apply_matching c csr)@ist.lmatch }) mt g + with + | (FailError _) as e -> raise e + | e when is_match_catchable e -> + apply_match ist csr tl)) + | Subterm (id,c) -> + (try + apply_sub_match ist 0 (id,c) csr mt + with | No_match -> + apply_match ist csr tl)) + | _ -> + errorlabstrm "Tacinterp.apply_match" (str + "No matching clauses for Match") in + let csr = constr_interp_may_eval ist constr + and ilr = read_match_rule ist.evc ist.env (constr_list ist) lmr in + apply_match ist csr ilr + +and tactic_interp ist (ast:raw_tactic_expr) g = + tac_interp ist.lfun ist.lmatch ist.debug ast g + +(* Interprets tactic expressions *) +and tac_interp lfun lmatch debug ast g = + let evc = project g + and env = pf_env g in + let ist = { evc=evc; env=env; lfun=lfun; lmatch=lmatch; + goalopt=Some g; debug=debug } in + try tactic_of_value (val_interp ist ast) g + with | NotTactic -> + errorlabstrm "Tacinterp.tac_interp" (str + "Must be a command or must give a tactic value") + +(* errorlabstrm "Tacinterp.tac_interp" (str + "Interpretation gives a non-tactic value") *) + +(* match (val_interp (evc,env,lfun,lmatch,(Some g),debug) ast) with + | VClosure tac -> (tac g) + | VFTactic (largs,f) -> (f largs g) + | VRTactic res -> res + | _ -> + errorlabstrm "Tacinterp.tac_interp" (str + "Interpretation gives a non-tactic value")*) + +(* Interprets a primitive tactic *) +and interp_atomic ist = function + (* Basic tactics *) + | TacIntroPattern l -> + Elim.h_intro_patterns (List.map (interp_intro_pattern ist) l) + | TacIntrosUntil hyp -> h_intros_until (interp_quantified_hypothesis ist hyp) + | TacIntroMove (ido,ido') -> + h_intro_move (option_app (ident_interp ist) ido) + (option_app (fun x -> ident_interp ist (snd x)) ido') + | TacAssumption -> h_assumption + | TacExact c -> h_exact (cast_constr_interp ist c) + | TacApply cb -> h_apply (interp_constr_with_bindings ist cb) + | TacElim (cb,cbo) -> + h_elim (interp_constr_with_bindings ist cb) + (option_app (interp_constr_with_bindings ist) cbo) + | TacElimType c -> h_elim_type (constr_interp ist c) + | TacCase cb -> h_case (interp_constr_with_bindings ist cb) + | TacCaseType c -> h_case_type (constr_interp ist c) + | TacFix (idopt,n) -> h_fix (id_opt_interp ist idopt) n + | TacMutualFix (id,n,l) -> + let f (id,n,c) = (ident_interp ist id,n,constr_interp ist c) in + h_mutual_fix (ident_interp ist id) n (List.map f l) + | TacCofix idopt -> h_cofix (id_opt_interp ist idopt) + | TacMutualCofix (id,l) -> + let f (id,c) = (ident_interp ist id,constr_interp ist c) in + h_mutual_cofix (ident_interp ist id) (List.map f l) + | TacCut c -> h_cut (constr_interp ist c) + | TacTrueCut (ido,c) -> h_true_cut (id_opt_interp ist ido) (constr_interp ist c) + | TacForward (b,na,c) -> h_forward b (name_interp ist na) (constr_interp ist c) + | TacGeneralize cl -> h_generalize (List.map (constr_interp ist) cl) + | TacGeneralizeDep c -> h_generalize_dep (constr_interp ist c) + | TacLetTac (id,c,clp) -> + let clp = check_clause_pattern ist clp in + h_let_tac (ident_interp ist id) (constr_interp ist c) clp + | TacInstantiate (n,c) -> h_instantiate n (constr_interp ist c) + + (* Automation tactics *) + | TacTrivial l -> Auto.h_trivial l + | TacAuto (n, l) -> Auto.h_auto n l + | TacAutoTDB n -> Dhyp.h_auto_tdb n + | TacDestructHyp (b,id) -> Dhyp.h_destructHyp b (hyp_interp ist id) + | TacDestructConcl -> Dhyp.h_destructConcl + | TacSuperAuto (n,l,b1,b2) -> Auto.h_superauto n l b1 b2 + | TacDAuto (n,p) -> Auto.h_dauto (n,p) + + (* Derived basic tactics *) + | TacOldInduction h -> h_old_induction (interp_quantified_hypothesis ist h) + | TacNewInduction c -> h_new_induction (interp_induction_arg ist c) + | TacOldDestruct h -> h_old_destruct (interp_quantified_hypothesis ist h) + | TacNewDestruct c -> h_new_destruct (interp_induction_arg ist c) + | TacDoubleInduction (h1,h2) -> + let h1 = interp_quantified_hypothesis ist h1 in + let h2 = interp_quantified_hypothesis ist h2 in + Elim.h_double_induction h1 h2 + | TacDecomposeAnd c -> Elim.h_decompose_and (constr_interp ist c) + | TacDecomposeOr c -> Elim.h_decompose_or (constr_interp ist c) + | TacDecompose (l,c) -> + let l = List.map (interp_inductive_or_metanum ist) l in + Elim.h_decompose l (constr_interp ist c) + | TacSpecialize (n,l) -> h_specialize n (interp_constr_with_bindings ist l) + | TacLApply c -> h_lapply (constr_interp ist c) + + (* Context management *) + | TacClear l -> h_clear (List.map (hyp_or_metanum_interp ist) l) + | TacClearBody l -> h_clear_body (List.map (hyp_or_metanum_interp ist) l) + | TacMove (dep,id1,id2) -> + h_move dep (hyp_interp ist id1) (hyp_interp ist id2) + | TacRename (id1,id2) -> + h_rename (hyp_interp ist id1) (hyp_interp ist id2) + + (* Constructors *) + | TacLeft bl -> h_left (bindings_interp ist bl) + | TacRight bl -> h_right (bindings_interp ist bl) + | TacSplit bl -> h_split (bindings_interp ist bl) + | TacAnyConstructor t -> + abstract_tactic (TacAnyConstructor t) + (Tactics.any_constructor (option_app (tactic_interp ist) t)) + | TacConstructor (n,bl) -> + h_constructor (skip_metaid n) (bindings_interp ist bl) + + (* Conversion *) + | TacReduce (r,cl) -> + h_reduce (redexp_interp ist r) (List.map (interp_hyp_location ist) cl) + | TacChange (c,cl) -> + h_change (constr_interp ist c) (List.map (interp_hyp_location ist) cl) + + (* Equivalence relations *) + | TacReflexivity -> h_reflexivity + | TacSymmetry -> h_symmetry + | TacTransitivity c -> h_transitivity (constr_interp ist c) + + (* For extensions *) + | TacExtend (opn,l) -> vernac_tactic (opn,List.map (genarg_interp ist) l) + | TacAlias (_,l,body) -> + let f x = match genarg_tag x with + | IdentArgType -> VIdentifier (ident_interp ist (out_gen rawwit_ident x)) + | QualidArgType -> VConstr (constr_of_reference (qualid_interp ist (out_gen rawwit_qualid x))) + | ConstrArgType -> VConstr (constr_interp ist (out_gen rawwit_constr x)) + | ConstrMayEvalArgType -> + VConstr (constr_interp_may_eval ist (out_gen rawwit_constr_may_eval x)) + | _ -> failwith "This generic type is not supported in alias" in + + tactic_of_value (val_interp { ist with lfun=(List.map (fun (x,c) -> (id_of_string x,f c)) l)@ist.lfun } body) + +let _ = forward_vcontext_interp := vcontext_interp + +(* Interprets tactic arguments *) +let interp_tacarg sign ast = (*unvarg*) (val_interp sign ast) + +(* Initial call for interpretation *) +let interp = fun ast -> tac_interp [] [] !debug ast + +(* Hides interpretation for pretty-print *) +let hide_interp t = abstract_tactic_expr (TacArg (Tacexp t)) (interp t) + +(* For bad tactic calls *) +let bad_tactic_args s = + anomalylabstrm s + (str "Tactic " ++ str s ++ str " called with bad arguments") + +(* Declaration of the TAC-DEFINITION object *) +let add (sp,td) = mactab := Gmap.add sp td !mactab + +let register_tacdef (sp,td) = + let ve = val_interp + {evc=Evd.empty;env=Global.env ();lfun=[]; + lmatch=[]; goalopt=None; debug=get_debug ()} + td in + sp,ve + +let cache_md (_,defs) = + (* Needs a rollback if something goes wrong *) + List.iter (fun (sp,_) -> Nametab.push_tactic_path sp) defs; + List.iter add (List.map register_tacdef defs) + +let (inMD,outMD) = + declare_object ("TAC-DEFINITION", + {cache_function = cache_md; + load_function = (fun _ -> ()); + open_function = cache_md; + export_function = (fun x -> Some x)}) + +(* Adds a definition for tactics in the table *) +let make_absolute_name (loc,id) = + let sp = Lib.make_path id in + if Gmap.mem sp !mactab then + errorlabstrm "Tacinterp.add_tacdef" + (str "There is already a Meta Definition or a Tactic Definition named " + ++ pr_sp sp); + sp + +let add_tacdef tacl = + let lfun = List.map (fun ((loc,id),_) -> id) tacl in + let tacl = List.map (fun (id,tac) -> (make_absolute_name id,tac)) tacl in + List.iter (fun (_,def) -> let _ = glob_tactic (lfun,[]) def in ()) tacl; + let _ = Lib.add_leaf (List.hd lfun) (inMD tacl) in + List.iter + (fun id -> Options.if_verbose msgnl (pr_id id ++ str " is defined")) lfun + +let interp_redexp env evc r = + let ist = + { evc=evc; env=env; lfun=[]; lmatch=[]; goalopt=None; debug=get_debug ()} + in + redexp_interp ist r + +let _ = Auto.set_extern_interp (fun l -> tac_interp [] l (get_debug())) +let _ = Dhyp.set_extern_interp interp diff --git a/tactics/tacinterp.mli b/tactics/tacinterp.mli new file mode 100644 index 000000000..c4017fc88 --- /dev/null +++ b/tactics/tacinterp.mli @@ -0,0 +1,115 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* tactic) + | VRTactic of (goal list sigma * validation) + | VContext of interp_sign * direction_flag + * (pattern_ast,raw_tactic_expr) match_rule list + | VFun of (identifier * value) list * identifier option list *raw_tactic_expr + | VVoid + | VInteger of int + | VIdentifier of identifier + | VConstr of constr + | VConstr_context of constr + | VRec of value ref + +(* Signature for interpretation: val\_interp and interpretation functions *) +and interp_sign = + { evc : Evd.evar_map; + env : Environ.env; + lfun : (identifier * value) list; + lmatch : (int * constr) list; + goalopt : goal sigma option; + debug : debug_info } + +(* Gives the identifier corresponding to an Identifier [tactic_arg] *) +val id_of_Identifier : value -> identifier + +(* Gives the constr corresponding to a Constr [value] *) +val constr_of_VConstr : value -> constr + +(* Transforms an id into a constr if possible *) +val constr_of_id : interp_sign -> identifier -> constr + +(* To embed several objects in Coqast.t *) +val tacticIn : (interp_sign -> raw_tactic_expr) -> raw_tactic_expr +val tacticOut : raw_tactic_expr -> (interp_sign -> raw_tactic_expr) +val valueIn : value -> raw_tactic_arg +val valueOut: raw_tactic_arg -> value +val constrIn : constr -> Coqast.t +val constrOut : Coqast.t -> constr +val loc : Coqast.loc + +(* Sets the debugger mode *) +val set_debug : debug_info -> unit + +(* Gives the state of debug *) +val get_debug : unit -> debug_info + +(* Adds a definition for tactics in the table *) +val add_tacdef : (identifier Util.located * raw_tactic_expr) list -> unit + +(* Adds an interpretation function for extra generic arguments *) +val add_genarg_interp : + string -> + (interp_sign -> raw_generic_argument -> closed_generic_argument) -> unit + +val genarg_interp : + interp_sign -> raw_generic_argument -> closed_generic_argument + +(* Interprets any expression *) +val val_interp : interp_sign -> raw_tactic_expr -> value + +(* +(* Interprets tactic arguments *) +val interp_tacarg : interp_sign -> raw_tactic_expr -> value +*) + +(* Interprets redexp arguments *) +val interp_redexp : Environ.env -> Evd.evar_map -> raw_red_expr + -> Tacred.red_expr + +(* Interprets tactic expressions *) +val tac_interp : (identifier * value) list -> (int * constr) list -> + debug_info -> raw_tactic_expr -> tactic + +(* Interprets constr expressions *) +val constr_interp : interp_sign -> constr_ast -> constr + +(* Initial call for interpretation *) +val interp : raw_tactic_expr -> tactic + +(* Hides interpretation for pretty-print *) +val hide_interp : raw_tactic_expr -> tactic + +(* Adds an interpretation function *) +val interp_add : string * (interp_sign -> Coqast.t -> value) -> unit + +(* Adds a possible existing interpretation function *) +val overwriting_interp_add : string * (interp_sign -> Coqast.t -> value) -> + unit + + diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 2285850c0..f154ef372 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -24,16 +24,19 @@ open Declare open Evd open Pfedit open Tacred +open Rawterm open Tacmach open Proof_trees open Proof_type open Logic open Evar_refiner open Clenv +open Refiner open Tacticals open Hipattern open Coqlib open Nametab +open Tacexpr exception Bound @@ -54,12 +57,14 @@ let rec nb_prod x = (* General functions *) (****************************************) +(* let get_pairs_from_bindings = let pair_from_binding = function | [(Bindings binds)] -> binds | _ -> error "not a binding list!" in List.map pair_from_binding +*) let string_of_inductive c = try match kind_of_term c with @@ -85,8 +90,10 @@ let rec head_constr_bound t l = let head_constr c = try head_constr_bound c [] with Bound -> error "Bound head variable" +(* let bad_tactic_args s l = raise (RefinerError (BadTacticArgs (s,l))) +*) (******************************************) (* Primitive tactics *) @@ -100,54 +107,27 @@ let refine = Tacmach.refine let convert_concl = Tacmach.convert_concl let convert_hyp = Tacmach.convert_hyp let thin = Tacmach.thin -let thin_body = Tacmach.thin_body -let move_hyp = Tacmach.move_hyp +let thin_body = Tacmach.thin_body + +(* Moving hypotheses *) +let move_hyp = Tacmach.move_hyp + +(* Renaming hypotheses *) let rename_hyp = Tacmach.rename_hyp -let mutual_fix = Tacmach.mutual_fix -let fix f n = mutual_fix f n [] - -let fix_noname n = - let id = Pfedit.get_current_proof_name () in - fix id n - -let dyn_mutual_fix argsl gl = - match argsl with - | [Integer n] -> fix_noname n gl - | [Identifier id;Integer n] -> fix id n gl - | ((Identifier id)::(Integer n)::lfix) -> - let rec decomp lar = function - | (Fixexp (id,n,ar)::rest) -> - decomp ((id,n,pf_interp_constr gl ar)::lar) rest - | [] -> (List.rev lar) - | _ -> bad_tactic_args "mutual_fix" argsl - in - let lar = decomp [] lfix in - mutual_fix id n lar gl - | l -> bad_tactic_args "mutual_fix" l +(* Refine as a fixpoint *) +let mutual_fix = Tacmach.mutual_fix +let fix ido n = match ido with + | None -> mutual_fix (Pfedit.get_current_proof_name ()) n [] + | Some id -> mutual_fix id n [] + +(* Refine as a cofixpoint *) let mutual_cofix = Tacmach.mutual_cofix -let cofix f = mutual_cofix f [] - -let cofix_noname n = - let id = Pfedit.get_current_proof_name () in - cofix id n - -let dyn_mutual_cofix argsl gl = - match argsl with - | [] -> cofix_noname gl - | [(Identifier id)] -> cofix id gl - | ((Identifier id)::lcofix) -> - let rec decomp lar = function - | (Cofixexp (id,ar)::rest) -> - decomp ((id,pf_interp_constr gl ar)::lar) rest - | [] -> List.rev lar - | _ -> bad_tactic_args "mutual_cofix" argsl - in - let lar = decomp [] lcofix in - mutual_cofix id lar gl - | l -> bad_tactic_args "mutual_cofix" l +let cofix = function + | None -> mutual_cofix (Pfedit.get_current_proof_name ()) [] + | Some id -> mutual_cofix id [] (**************************************************************) (* Reduction and conversion tactics *) @@ -213,7 +193,7 @@ let change_option t = function | Some id -> reduct_in_hyp (change_hyp_and_check t) id | None -> reduct_in_concl (change_concl_and_check t) -(* Pour usage interne (le niveau User est pris en compte par dyn_reduce) *) +(* Pour usage interne (le niveau User est pris en compte par reduce) *) let red_in_concl = reduct_in_concl red_product let red_in_hyp = reduct_in_hyp red_product let red_option = reduct_option red_product @@ -231,12 +211,7 @@ let unfold_in_hyp loccname = reduct_in_hyp (unfoldn loccname) let unfold_option loccname = reduct_option (unfoldn loccname) let pattern_option l = reduct_option (pattern_occs l) -let dyn_change = function - | [Constr c; Clause cl] -> - (fun goal -> -(* let c = Astterm.interp_type (project goal) (pf_env goal) com in*) - in_combinator (change_in_concl c) (change_in_hyp c) cl goal) - | l -> bad_tactic_args "change" l +let change c = in_combinator (change_in_concl c) (change_in_hyp c) (* A function which reduces accordingly to a reduction expression, as the command Eval does. *) @@ -244,10 +219,6 @@ let dyn_change = function let reduce redexp cl goal = redin_combinator (reduction_of_redexp redexp) cl goal -let dyn_reduce = function - | [Redexp redexp; Clause cl] -> (fun goal -> reduce redexp cl goal) - | l -> bad_tactic_args "reduce" l - (* Unfolding occurrences of a constant *) let unfold_constr = function @@ -360,93 +331,59 @@ let intros_replacing ids gls = (* User-level introduction tactics *) -let dyn_intro = function - | [] -> intro_gen (IntroAvoid []) None true - | [Identifier id] -> intro_gen (IntroMustBe id) None true - | l -> bad_tactic_args "intro" l - -let dyn_intro_move = function - | [Identifier id2] -> intro_gen (IntroAvoid []) (Some id2) true - | [Identifier id; Identifier id2] -> - intro_gen (IntroMustBe id) (Some id2) true - | l -> bad_tactic_args "intro_move" l - -let rec intros_until s g = - match pf_lookup_name_as_renamed (pf_env g) (pf_concl g) s with - | Some depth -> tclDO depth intro g - | None -> - try - ((tclTHEN (reduce (Red true) []) (intros_until s)) g) - with Redelimination -> - errorlabstrm "Intros" - (str ("No hypothesis "^(string_of_id s)^" in current goal") ++ - str " even after head-reduction") - -let rec intros_until_n_gen red n g = - match pf_lookup_index_as_renamed (pf_env g) (pf_concl g) n with - | Some depth -> tclDO depth intro g +let intro_move idopt idopt' = match idopt with + | None -> intro_gen (IntroAvoid []) idopt' true + | Some id -> intro_gen (IntroMustBe id) idopt' true + +let pf_lookup_hypothesis_as_renamed env ccl = function + | AnonHyp n -> pf_lookup_index_as_renamed env ccl n + | NamedHyp id -> pf_lookup_name_as_renamed env ccl id + +let pf_lookup_hypothesis_as_renamed_gen red h gl = + let env = pf_env gl in + let rec aux ccl = + match pf_lookup_hypothesis_as_renamed env ccl h with + | None when red -> aux (reduction_of_redexp (Red true) env Evd.empty ccl) + | x -> x + in + try aux (pf_concl gl) + with Redelimination -> None + +let is_quantified_hypothesis id g = + match pf_lookup_hypothesis_as_renamed_gen true (NamedHyp id) g with + | Some _ -> true + | None -> false + +let msg_quantified_hypothesis = function + | NamedHyp id -> + str "hypothesis " ++ pr_id id + | AnonHyp n -> + int n ++ str (match n with 1 -> "st" | 2 -> "nd" | _ -> "th") ++ + str " non dependent hypothesis" + +let depth_of_quantified_hypothesis red h gl = + match pf_lookup_hypothesis_as_renamed_gen red h gl with + | Some depth -> depth | None -> - if red then - try - ((tclTHEN (reduce (Red true) []) (intros_until_n_gen red n)) g) - with Redelimination -> - errorlabstrm "Intros" - (str ("No "^(string_of_int n)) ++ - str (match n with 1 -> "st" | 2 -> "nd" | _ -> "th") ++ - str " non dependent hypothesis in current goal" ++ - str " even after head-reduction") - else - errorlabstrm "Intros" (str "No such hypothesis in current goal") + errorlabstrm "lookup_quantified_hypothesis" + (str "No " ++ msg_quantified_hypothesis h ++ + str " in current goal" ++ + if red then str " even after head-reduction" else mt ()) + +let intros_until_gen red h g = + tclDO (depth_of_quantified_hypothesis red h g) intro g +let intros_until_id id = intros_until_gen true (NamedHyp id) +let intros_until_n_gen red n = intros_until_gen red (AnonHyp n) + +let intros_until = intros_until_gen true let intros_until_n = intros_until_n_gen true let intros_until_n_wored = intros_until_n_gen false -let dyn_intros_until = function - | [Identifier id] -> intros_until id - | [Integer n] -> intros_until_n n - | l -> bad_tactic_args "Intros until" l - -let tactic_try_intros_until tac = function - | Identifier id -> - tclTHEN (tclTRY (intros_until id)) (tac id) - | Integer n -> - tclTHEN (intros_until_n n) - (fun gl -> let id,_,_ = pf_last_hyp gl in tac id gl) - | c -> bad_tactic_args "tactic_try_intros_until" [c] - -let hide_ident_or_numarg_tactic s tac = - let tacfun = function - | [Identifier id] -> tclTHEN (tclTRY (intros_until id)) (tac id) - | [Integer n] -> - tclTHEN (intros_until_n n) - (fun gl -> let id,_,_ = pf_last_hyp gl in tac id gl) - | _ -> assert false in - add_tactic s tacfun; - fun id -> vernac_tactic(s,[Identifier id]) - - -(* Obsolete, remplace par intros_unitl_n ? -let intros_do n g = - let depth = - let rec lookup all nodep c = match kind_of_term c with - | Prod (name,_,c') -> - (match name with - | Name(s') -> - if dependent (mkRel 1) c' then - lookup (all+1) nodep c' - else if nodep = n then - all - else - lookup (all+1) (nodep+1) c' - | Anonymous -> - if nodep=n then all else lookup (all+1) (nodep+1) c') - | Cast (c,_) -> lookup all nodep c - | _ -> error "No such hypothesis in current goal" - in - lookup 1 1 (pf_concl g) - in - tclDO depth intro g -*) +let try_intros_until tac = function + | NamedHyp id -> tclTHEN (tclTRY (intros_until_id id)) (tac id) + | AnonHyp n -> tclTHEN (intros_until_n n) (onLastHyp tac) + let rec intros_move = function | [] -> tclIDTAC | (hyp,destopt) :: rest -> @@ -508,8 +445,7 @@ let bring_hyps hyps = (* Resolution with missing arguments *) - -let apply_with_bindings (c,lbind) gl = +let apply_with_bindings (c,lbind) gl = let apply = match kind_of_term c with | Lambda _ -> res_pf_cast @@ -539,11 +475,11 @@ let apply_with_bindings (c,lbind) gl = apply kONT clause gl -let apply c = apply_with_bindings (c,[]) -let apply_com = tactic_com (fun c -> apply_with_bindings (c,[])) +let apply c = apply_with_bindings (c,NoBindings) +let apply_com = tactic_com (fun c -> apply_with_bindings (c,NoBindings)) let apply_list = function - | c::l -> apply_with_bindings (c,List.map (fun com ->(Com,com)) l) + | c::l -> apply_with_bindings (c,ImplicitBindings l) | _ -> assert false (* Resolution with no reduction on the type *) @@ -557,85 +493,68 @@ let apply_without_reduce_com = tactic_com apply_without_reduce let refinew_scheme kONT clause gl = res_pf kONT clause gl -let dyn_apply l = - match l with - | [Command com; Bindings binds] -> - tactic_com_bind_list apply_with_bindings (com,binds) - | [Constr c; Cbindings binds] -> - apply_with_bindings (c,binds) - | l -> - bad_tactic_args "apply" l +(* A useful resolution tactic which, if c:A->B, transforms |- C into + |- B -> C and |- A (which is realized by Cut B;[Idtac|Apply c] -(* A useful resolution tactic, equivalent to Cut type_of_c;[Idtac|Apply c] *) + ------------------- + Gamma |- c : A -> B Gamma |- ?2 : A + ---------------------------------------- + Gamma |- B Gamma |- ?1 : B -> C + ----------------------------------------------------- + Gamma |- ? : C + *) let cut_and_apply c gl = let goal_constr = pf_concl gl in match kind_of_term (pf_hnf_constr gl (pf_type_of gl c)) with | Prod (_,c1,c2) when not (dependent (mkRel 1) c2) -> - tclTHENS - (apply_type (mkProd (Anonymous,c2,goal_constr)) - [mkMeta (new_meta())]) - [tclIDTAC;apply_term c [mkMeta (new_meta())]] gl + tclTHENLAST + (apply_type (mkProd (Anonymous,c2,goal_constr)) [mkMeta(new_meta())]) + (apply_term c [mkMeta (new_meta())]) gl | _ -> error "Imp_elim needs a non-dependent product" -let dyn_cut_and_apply = function - | [Command com] -> tactic_com cut_and_apply com - | [Constr c] -> cut_and_apply c - | l -> bad_tactic_args "cut_and_apply" l - (**************************) (* Cut tactics *) (**************************) -let true_cut id c gl = - match kind_of_term (hnf_type_of gl c) with - | Sort _ -> internal_cut id c gl - | _ -> error "Not a proposition or a type" - -let true_cut_anon c gl = +let true_cut idopt c gl = match kind_of_term (hnf_type_of gl c) with | Sort s -> - let d = match s with Prop _ -> "H" | Type _ -> "X" in - let id = next_name_away_with_default d Anonymous (pf_ids_of_hyps gl) in - internal_cut id c gl + let id = + match idopt with + | None -> + let d = match s with Prop _ -> "H" | Type _ -> "X" in + next_name_away_with_default d Anonymous (pf_ids_of_hyps gl) + | Some id -> id + in + internal_cut id c gl | _ -> error "Not a proposition or a type" -let dyn_true_cut = function - | [Command com] -> tactic_com_sort true_cut_anon com - | [Constr c] -> true_cut_anon c - | [Command com; Identifier id] -> tactic_com_sort (true_cut id) com - | [Constr c; Identifier id] -> true_cut id c - | l -> bad_tactic_args "true_cut" l - let cut c gl = match kind_of_term (hnf_type_of gl c) with | Sort _ -> let id=next_name_away_with_default "H" Anonymous (pf_ids_of_hyps gl) in let t = mkProd (Anonymous, c, pf_concl gl) in - tclTHENS + tclTHENFIRST (internal_cut_rev id c) - [tclTHEN (apply_type t [mkVar id]) (thin [id]); - tclIDTAC] gl + (tclTHEN (apply_type t [mkVar id]) (thin [id])) + gl | _ -> error "Not a proposition or a type" -let dyn_cut = function - | [Command com] -> tactic_com_sort cut com - | [Constr c] -> cut c - | l -> bad_tactic_args "cut" l - -let cut_intro t = (tclTHENS (cut t) [intro;tclIDTAC]) +let cut_intro t = tclTHENFIRST (cut t) intro let cut_replacing id t = - (tclTHENS (cut t) - [(tclORELSE (intro_replacing id) - (tclORELSE (intro_erasing id) - (intro_using id))); - tclIDTAC]) + tclTHENFIRST + (cut t) + (tclORELSE + (intro_replacing id) + (tclORELSE (intro_erasing id) + (intro_using id))) let cut_in_parallel l = let rec prec = function | [] -> tclIDTAC - | h::t -> (tclTHENS (cut h) ([prec t;tclIDTAC])) + | h::t -> tclTHENFIRST (cut h) (prec t) in prec (List.rev l) @@ -693,13 +612,6 @@ let generalize lconstr gl = let newcl = List.fold_right (generalize_goal gl) lconstr (pf_concl gl) in apply_type newcl lconstr gl -let dyn_generalize = - fun argsl -> generalize (List.map Tacinterp.constr_of_Constr argsl) - -let dyn_generalize_dep = function - | [Constr csr] -> generalize_dep csr - | l -> bad_tactic_args "dyn_generalize_dep" l - (* Faudra-t-il une version avec plusieurs args de generalize_dep ? Cela peut-tre troublant de faire "Generalize Dependent H n" dans "n:nat; H:n=n |- P(n)" et d'chouer parce que H a disparu aprs la @@ -736,21 +648,35 @@ let quantify lconstr = the left of each x1, ...). *) -let letin_abstract id c (occ_ccl,occ_hyps) gl = - let everywhere = (occ_ccl = None) & (occ_hyps = []) in +let occurrences_of_hyp id = function + | None, [] -> (* Everywhere *) Some [] + | _, occ_hyps -> try Some (List.assoc id occ_hyps) with Not_found -> None + +let occurrences_of_goal = function + | None, [] -> (* Everywhere *) Some [] + | Some gocc as x, _ -> x + | None, _ -> None + +let everywhere (occ_ccl,occ_hyps) = (occ_ccl = None) & (occ_hyps = []) + +let letin_abstract id c occs gl = let env = pf_env gl in let compute_dependency _ (hyp,_,_ as d) ctxt = let d' = try - let occ = if everywhere then [] else List.assoc hyp occ_hyps in - let newdecl = subst_term_occ_decl env occ c d in - if d = newdecl then - if not everywhere then raise (RefinerError (DoesNotOccurIn (c,hyp))) - else raise Not_found - else - (subst1_decl (mkVar id) newdecl, true) - with Not_found -> - (d,List.exists (fun((id,_,_),dep) -> dep && occur_var_in_decl env id d) ctxt) + match occurrences_of_hyp hyp occs with + | None -> raise Not_found + | Some occ -> + let newdecl = subst_term_occ_decl env occ c d in + if d = newdecl then + if not (everywhere occs) + then raise (RefinerError (DoesNotOccurIn (c,hyp))) + else raise Not_found + else + (subst1_decl (mkVar id) newdecl, true) + with Not_found -> + (d,List.exists + (fun ((id,_,_),dep) -> dep && occur_var_in_decl env id d) ctxt) in d'::ctxt in let ctxt' = fold_named_context compute_dependency env ~init:[] in @@ -758,27 +684,7 @@ let letin_abstract id c (occ_ccl,occ_hyps) gl = if b then ((d::depdecls,(hyp,lhyp)::marks), lhyp) else (accu, Some hyp) in let (depdecls,marks),_ = List.fold_left compute_marks (([],[]),None) ctxt' in -(* - let abstract ((depdecls,marks as accu),lhyp) (hyp,_,_ as d) = - try - let occ = if everywhere then [] else List.assoc hyp occ_hyps in - let newdecl = subst_term_occ_decl env occ c d in - if d = newdecl then - if not everywhere then raise (RefinerError (DoesNotOccurIn (c,hyp))) - else - if List.exists (fun (id,_,_) -> occur_var id d) - (accu, Some hyp) - else - let newdecl = subst1_decl (mkVar id) newdecl in - ((newdecl::depdecls,(hyp,lhyp)::marks), lhyp) - with Not_found -> - (accu, Some hyp) - in - let (depdecls,marks),_ = - fold_named_context_reverse abstract ~init:(([],[]),None) env in(* *) -*) - let occ_ccl = if everywhere then Some [] else occ_ccl in - let ccl = match occ_ccl with + let ccl = match occurrences_of_goal occs with | None -> pf_concl gl | Some occ -> subst1 (mkVar id) (subst_term_occ env occ c (pf_concl gl)) in @@ -810,7 +716,7 @@ let letin_tac with_eq name c occs gl = if with_eq then tclIDTAC else thin_body [id]; intros_move marks ] gl -let check_hypotheses_occurrences_list env occl = +let check_hypotheses_occurrences_list env (_,occl) = let rec check acc = function | (hyp,_) :: rest -> if List.mem hyp acc then @@ -821,27 +727,9 @@ let check_hypotheses_occurrences_list env occl = | [] -> () in check [] occl -let dyn_lettac args gl = match args with - | [Identifier id; Command com; Letpatterns (o,l)] -> - check_hypotheses_occurrences_list (pf_env gl) l; - letin_tac true (Name id) (pf_interp_constr gl com) (o,l) gl - | [Identifier id; Constr c; Letpatterns (o,l)] -> - check_hypotheses_occurrences_list (pf_env gl) l; - letin_tac true (Name id) c (o,l) gl - | l -> bad_tactic_args "letin" l - let nowhere = (Some [],[]) -let dyn_forward args gl = match args with - | [Quoted_string s; Command com; Identifier id] -> - letin_tac (s="KeepBody") (Name id) (pf_interp_constr gl com) nowhere gl - | [Quoted_string s; Constr c; Identifier id] -> - letin_tac (s="KeepBody") (Name id) c nowhere gl - | [Quoted_string s; Constr c] -> - letin_tac (s="KeepBody") Anonymous c nowhere gl - | [Quoted_string s; Command c] -> - letin_tac (s="KeepBody") Anonymous (pf_interp_constr gl c) nowhere gl - | l -> bad_tactic_args "forward" l +let forward b na c = letin_tac b na c nowhere (********************************************************************) (* Exact tactics *) @@ -857,23 +745,10 @@ let exact_check c gl = let exact_no_check = refine -let dyn_exact_no_check cc gl = match cc with - | [Constr c] -> exact_no_check c gl - | [Command com] -> - let evc = (project gl) in - let concl = (pf_concl gl) in - let c = Astterm.interp_casted_constr evc (pf_env gl) com concl in - refine c gl - | l -> bad_tactic_args "exact_no_check" l - -let dyn_exact_check cc gl = match cc with - | [Constr c] -> exact_check c gl - | [Command com] -> - let evc = (project gl) in - let concl = (pf_concl gl) in - let c = Astterm.interp_casted_constr evc (pf_env gl) com concl in - refine c gl - | l -> bad_tactic_args "exact_check" l +let exact_proof c gl = + (* on experimente la synthese d'ise dans exact *) + let c = Astterm.interp_casted_constr (project gl) (pf_env gl) c (pf_concl gl) + in refine c gl let (assumption : tactic) = fun gl -> let concl = pf_concl gl in @@ -885,11 +760,6 @@ let (assumption : tactic) = fun gl -> in arec (pf_hyps gl) -let dyn_assumption = function - | [] -> assumption - | l -> bad_tactic_args "assumption" l - - (*****************************************************************) (* Modification of a local context *) (*****************************************************************) @@ -899,22 +769,10 @@ let dyn_assumption = function * subsequently used in other hypotheses or in the conclusion of the * goal. *) -let clear ids gl = thin ids gl -let dyn_clear = function - | [Clause ids] -> - if ids=[] then tclIDTAC - else - let out = function InHyp id -> id | _ -> assert false in - clear (List.map out ids) - | l -> bad_tactic_args "clear" l +let clear ids gl = (* avant seul dyn_clear n'echouait pas en [] *) + if ids=[] then tclIDTAC gl else thin ids gl let clear_body = thin_body -let dyn_clear_body = function - | [Clause ids] -> - let out = function InHyp id -> id | _ -> assert false in - clear_body (List.map out ids) - | l -> bad_tactic_args "clear_body" l - (* Takes a list of booleans, and introduces all the variables * quantified in the goal which are associated with a value @@ -940,44 +798,10 @@ let new_hyp mopt c lbind g = | Some m -> if m < nargs then list_firstn m tstack else tstack | None -> tstack) in - (tclTHENL (tclTHEN (kONT clause.hook) + (tclTHENLAST (tclTHEN (kONT clause.hook) (cut (pf_type_of g cut_pf))) ((tclORELSE (apply cut_pf) (exact_no_check cut_pf)))) g -let dyn_new_hyp argsl gl = - match argsl with - | [Integer n; Command com; Bindings binds] -> - tactic_bind_list - (new_hyp (Some n) - (pf_interp_constr gl com)) - binds gl - | [Command com; Bindings binds] -> - tactic_bind_list - (new_hyp None - (pf_interp_constr gl com)) - binds gl - | [Integer n; Constr c; Cbindings binds] -> - new_hyp (Some n) c binds gl - | [Constr c; Cbindings binds] -> - new_hyp None c binds gl - | l -> bad_tactic_args "new_hyp" l - -(* Moving hypotheses *) - -let dyn_move = function - | [Identifier idfrom; Identifier idto] -> move_hyp false idfrom idto - | l -> bad_tactic_args "move" l - -let dyn_move_dep = function - | [Identifier idfrom; Identifier idto] -> move_hyp true idfrom idto - | l -> bad_tactic_args "move_dep" l - -(* Renaming hypotheses *) - -let dyn_rename = function - | [Identifier idfrom; Identifier idto] -> rename_hyp idfrom idto - | l -> bad_tactic_args "rename" l - (************************) (* Introduction tactics *) (************************) @@ -1002,63 +826,28 @@ let constructor_tac boundopt i lbind gl = let one_constructor i = constructor_tac None i -let any_constructor gl = - let cl = pf_concl gl in - let (mind,redcl) = pf_reduce_to_quantified_ind gl cl in - let nconstr = - Array.length (snd (Global.lookup_inductive mind)).mind_consnames - and sigma = project gl in - if nconstr = 0 then error "The type has no constructors"; - tclFIRST (List.map (fun i -> one_constructor i []) - (interval 1 nconstr)) gl - (* Try to apply the constructor of the inductive definition followed by a tactic t given as an argument. Should be generalize in Constructor (Fun c : I -> tactic) *) -let tclConstrThen t gl = - let mind = fst (pf_reduce_to_quantified_ind gl (pf_concl gl)) - in let lconstr = - (snd (Global.lookup_inductive mind)).mind_consnames - in let nconstr = Array.length lconstr - in +let any_constructor tacopt gl = + let t = match tacopt with None -> tclIDTAC | Some t -> t in + let mind = fst (pf_reduce_to_quantified_ind gl (pf_concl gl)) in + let nconstr = + Array.length (snd (Global.lookup_inductive mind)).mind_consnames in if nconstr = 0 then error "The type has no constructors"; - tclFIRST (List.map (fun i -> (tclTHEN (one_constructor i []) t)) + tclFIRST (List.map (fun i -> tclTHEN (one_constructor i NoBindings) t) (interval 1 nconstr)) gl -let dyn_constructor = function - | [Integer i; Bindings binds] -> tactic_bind_list (one_constructor i) binds - | [Integer i; Cbindings binds] -> (one_constructor i) binds - | [Tac (tac,_)] -> tclConstrThen tac - | [] -> any_constructor - | l -> bad_tactic_args "constructor" l - - let left = constructor_tac (Some 2) 1 -let simplest_left = left [] - -let dyn_left = function - | [Cbindings binds] -> left binds - | [Bindings binds] -> tactic_bind_list left binds - | l -> bad_tactic_args "left" l +let simplest_left = left NoBindings let right = constructor_tac (Some 2) 2 -let simplest_right = right [] - -let dyn_right = function - | [Cbindings binds] -> right binds - | [Bindings binds] -> tactic_bind_list right binds - | l -> bad_tactic_args "right" l - +let simplest_right = right NoBindings let split = constructor_tac (Some 1) 1 -let simplest_split = split [] - -let dyn_split = function - | [Cbindings binds] -> split binds - | [Bindings binds] -> tactic_bind_list split binds - | l -> bad_tactic_args "split" l +let simplest_split = split NoBindings (********************************************) (* Elimination tactics *) @@ -1112,17 +901,33 @@ let general_elim (c,lbindc) (elimc,lbindelimc) gl = (* Elimination tactic with bindings but using the default elimination * constant associated with the type. *) -let default_elim (c,lbindc) gl = +let find_eliminator c gl = let env = pf_env gl in let (ind,t) = reduce_to_quantified_ind env (project gl) (pf_type_of gl c) in let s = elimination_sort_of_goal gl in - let elimc = Indrec.lookup_eliminator ind s in - general_elim (c,lbindc) (elimc,[]) gl - + try Indrec.lookup_eliminator ind s + with Not_found -> + let dir, base = repr_path (path_of_inductive env ind) in + let id = Indrec.make_elimination_ident base s in + errorlabstrm "default_elim" + (str "Cannot find the elimination combinator :" ++ + pr_id id ++ spc () ++ + str "The elimination of the inductive definition :" ++ + pr_id base ++ spc () ++ str "on sort " ++ + spc () ++ print_sort (new_sort_in_family s) ++ + str " is probably not allowed") + +let default_elim (c,lbindc) gl = + general_elim (c,lbindc) (find_eliminator c gl,NoBindings) gl + +let elim (c,lbindc) elim gl = + match elim with + | Some (elimc,lbindelimc) -> general_elim (c,lbindc) (elimc,lbindelimc) gl + | None -> general_elim (c,lbindc) (find_eliminator c gl,NoBindings) gl (* The simplest elimination tactic, with no substitutions at all. *) -let simplest_elim c = default_elim (c,[]) +let simplest_elim c = default_elim (c,NoBindings) (* Elimination in hypothesis *) @@ -1383,11 +1188,6 @@ let find_atomic_param_of_ind mind indtyp = Others solutions are welcome *) -(* -type hyp_status = -let hyps_map -*) - exception Shunt of identifier option let cook_sign hyp0 indvars env = @@ -1438,44 +1238,12 @@ let cook_sign hyp0 indvars env = let statuslists = (!lstatus,List.rev !rstatus) in (statuslists, lhyp0, !indhyps, !decldeps) - -(* Vieille version en une seule passe grace l'ordre suprieur mais - trop difficile comprendre - -let cook_sign hyp0 indvars sign = - let finaldeps = ref ([],[]) in - let indhyps = ref [] in - let hyp0succ = ref None in - let cook_init (hdeps,tdeps) rhyp before = - finaldeps := (List.rev hdeps, List.rev tdeps); - (None, []) in - let cook_hyp compute_rhyp hyp typ ((hdeps,tdeps) as deps) = - fun rhyp before -> - match () with - _ when (List.mem hyp indvars) - -> let result = compute_rhyp deps rhyp before in - indhyps := hyp::!indhyps; result - | _ when hyp = hyp0 - -> let (lhyp,statl) = compute_rhyp deps rhyp true in - hyp0succ := lhyp; (None (* fake value *),statl) - | _ when (List.exists (fun id -> occur_var id typ) (hyp0::indvars) - or List.exists (fun id -> occur_var id typ) hdeps) - -> let deps' = (hyp::hdeps, typ::tdeps) in - let (lhyp,statl) = compute_rhyp deps' rhyp before in - let hyp = if before then lhyp else rhyp in - (lhyp,(DEPENDENT (before,hyp,hyp))::statl) - | _ -> - let (_,statl) = compute_rhyp deps (Some hyp) before - in (Some hyp, statl) - in let (_,statuslist) = it_sign cook_hyp cook_init sign ([],[]) None false in - (statuslist, !hyp0succ, !indhyps, !finaldeps) -*) - -let induction_tac varname typ (elimc,elimt) gl = +let induction_tac varname typ (elimc,elimt,lbindelimc) gl = let c = mkVar varname in let (wc,kONT) = startWalk gl in - let indclause = make_clenv_binding wc (c,typ) [] in - let elimclause = make_clenv_binding wc (mkCast (elimc,elimt),elimt) [] in + let indclause = make_clenv_binding wc (c,typ) NoBindings in + let elimclause = + make_clenv_binding wc (mkCast (elimc,elimt),elimt) lbindelimc in elimination_clause_scheme kONT elimclause indclause gl let is_indhyp p n t = @@ -1510,9 +1278,9 @@ let compute_elim_signature_and_roughly_check elimt mind = | Prod (_,t,c) -> (check_branch n t lra.(n)) :: (check_elim c (n+1)) | _ -> error "Not an eliminator: some constructor case is lacking" in let _,elimt3 = decompose_prod_n npred elimt2 in - check_elim elimt3 0 + Array.of_list (check_elim elimt3 0) -let induction_from_context isrec style hyp0 gl = +let induction_from_context isrec style elim hyp0 gl = (*test suivant sans doute inutile car refait par le letin_tac*) if List.mem hyp0 (ids_of_named_context (Global.named_context())) then errorlabstrm "induction" @@ -1520,12 +1288,17 @@ let induction_from_context isrec style hyp0 gl = let tmptyp0 = pf_get_hyp_typ gl hyp0 in let env = pf_env gl in let (mind,typ0) = pf_reduce_to_quantified_ind gl tmptyp0 in - let indvars = find_atomic_param_of_ind mind (snd (decompose_prod typ0)) in - let elimc = - if isrec then Indrec.lookup_eliminator mind (elimination_sort_of_goal gl) - else Indrec.make_case_gen env (project gl) mind (elimination_sort_of_goal gl) - in + let elimc,lbindelimc = match elim with + | None -> + let s = elimination_sort_of_goal gl in + (if isrec then Indrec.lookup_eliminator mind s + else Indrec.make_case_gen env (project gl) mind s), + NoBindings + | Some elim -> + (* Not really robust: no control on the form of the combinator *) + elim in let elimt = pf_type_of gl elimc in + let indvars = find_atomic_param_of_ind mind (snd (decompose_prod typ0)) in let (statlists,lhyp0,indhyps,deps) = cook_sign hyp0 indvars env in let tmpcl = it_mkNamedProd_or_LetIn (pf_concl gl) deps in let lr = compute_elim_signature_and_roughly_check elimt mind in @@ -1538,32 +1311,35 @@ let induction_from_context isrec style hyp0 gl = eux qui ouvrent de nouveaux buts arrivent en premier dans la liste des sous-buts du fait qu'ils sont le plus gauche dans le combinateur engendr par make_case_gen (un "Cases (hyp0 ?) of - ...") et on ne peut plus appliquer tclTHENST aprs; en revanche, + ...") et on ne peut plus appliquer tclTHENSI aprs; en revanche, comme lookup_eliminator renvoie un combinateur de la forme "ind_rec ... (hyp0 ?)", les buts correspondant des arguments de - hyp0 sont maintenant la fin et tclTHENS marche !!! *) + hyp0 sont maintenant la fin et tclTHENSI marche !!! *) +(* if not isrec && nb_prod typ0 <> 0 && lr <> [] (* passe-droit *) then error "Cases analysis on a functional term not implemented"; - +*) tclTHENLIST [ apply_type tmpcl args; thin dephyps; - tclTHENST + (if isrec then tclTHENFIRSTn else tclTHENLASTn) (tclTHEN - (induction_tac hyp0 typ0 (elimc,elimt)) + (induction_tac hyp0 typ0 (elimc,elimt,lbindelimc)) (thin (hyp0::indhyps))) - (List.map + (Array.map (induct_discharge style mind statlists hyp0 lhyp0 (List.rev dephyps)) lr) - tclIDTAC ] + ] gl let induction_with_atomization_of_ind_arg isrec hyp0 = tclTHEN (atomize_param_of_ind hyp0) - (induction_from_context isrec false hyp0) + (induction_from_context isrec false None hyp0) -let new_induct isrec c gl = +(* This is Induction since V7 ("natural" induction both in quantified + premisses and introduced ones) *) +let new_induct_gen isrec c gl = match kind_of_term c with | Var id when not (mem_named_context id (Global.named_context())) -> induction_with_atomization_of_ind_arg isrec id gl @@ -1575,60 +1351,34 @@ let new_induct isrec c gl = (letin_tac true (Name id) c (None,[])) (induction_with_atomization_of_ind_arg isrec id) gl -(* -let new_induct_nodep isrec n = - tclTHEN (intros_until_n n) (induction_with_atomization_of_ind_arg isrec None) -*) +let new_induct_destruct isrec = function + | ElimOnConstr c -> new_induct_gen isrec c + | ElimOnAnonHyp n -> + tclTHEN (intros_until_n n) (tclLAST_HYP (new_induct_gen isrec)) + (* Identifier apart because id can be quantified in goal and not typable *) + | ElimOnIdent (_,id) -> + tclTHEN (tclTRY (intros_until_id id)) (new_induct_gen isrec (mkVar id)) + +let new_induct = new_induct_destruct true +let new_destruct = new_induct_destruct false (* The registered tactic, which calls the default elimination * if no elimination constant is provided. *) - -let dyn_elim = function - | [Constr mp; Cbindings mpbinds] -> - default_elim (mp,mpbinds) - | [Command mp; Bindings mpbinds] -> - tactic_com_bind_list default_elim (mp,mpbinds) - | [Command mp; Bindings mpbinds; Command elimc; Bindings elimcbinds] -> - let funpair2funlist f = (function [x;y] -> f x y | _ -> assert false) in - tactic_com_bind_list_list - (funpair2funlist general_elim) - [(mp,mpbinds);(elimc,elimcbinds)] - | [Constr mp; Cbindings mpbinds; Constr elimc; Cbindings elimcbinds] -> - general_elim (mp,mpbinds) (elimc,elimcbinds) - | l -> bad_tactic_args "elim" l (* Induction tactics *) (* This was Induction before 6.3 (induction only in quantified premisses) *) -let raw_induct s = tclTHEN (intros_until s) (tclLAST_HYP simplest_elim) +let raw_induct s = tclTHEN (intros_until_id s) (tclLAST_HYP simplest_elim) let raw_induct_nodep n = tclTHEN (intros_until_n n) (tclLAST_HYP simplest_elim) (* This was Induction in 6.3 (hybrid form) *) -let old_induct s =tclORELSE (raw_induct s) (induction_from_context true true s) +let old_induct_id s = + tclORELSE (raw_induct s) (induction_from_context true true None s) let old_induct_nodep = raw_induct_nodep -(* This is Induction since V7 ("natural" induction both in quantified - premisses and introduced ones) *) -let dyn_new_induct = function - | [(Command c)] -> tactic_com (new_induct true) c - | [(Constr x)] -> new_induct true x - (* Identifier apart because id can be quantified in goal and not typable *) - | [Integer _ | Identifier _ as arg] -> - tactic_try_intros_until (fun id -> new_induct true (mkVar id)) arg - | l -> bad_tactic_args "induct" l - -(* This was Induction before 6.3 (induction only in quantified premisses) -let dyn_raw_induct = function - | [Identifier x] -> raw_induct x - | [Integer n] -> raw_induct_nodep n - | l -> bad_tactic_args "raw_induct" l -*) - -(* This was Induction in 6.3 (hybrid form) *) -let dyn_old_induct = function - | [(Identifier n)] -> old_induct n - | [Integer n] -> raw_induct_nodep n - | l -> bad_tactic_args "raw_induct" l +let old_induct = function + | NamedHyp id -> old_induct_id id + | AnonHyp n -> old_induct_nodep n (* Case analysis tactics *) @@ -1640,37 +1390,20 @@ let general_case_analysis (c,lbindc) gl = let case = if occur_term c (pf_concl gl) then Indrec.make_case_dep else Indrec.make_case_gen in let elim = case env sigma mind sort in - general_elim (c,lbindc) (elim,[]) gl + general_elim (c,lbindc) (elim,NoBindings) gl -let simplest_case c = general_case_analysis (c,[]) - -let dyn_case =function - | [Constr mp; Cbindings mpbinds] -> - general_case_analysis (mp,mpbinds) - | [Command mp; Bindings mpbinds] -> - tactic_com_bind_list general_case_analysis (mp,mpbinds) - | l -> bad_tactic_args "case" l - +let simplest_case c = general_case_analysis (c,NoBindings) (* Destruction tactics *) -let destruct s = (tclTHEN (intros_until s) (tclLAST_HYP simplest_case)) -let destruct_nodep n = (tclTHEN (intros_until_n n) (tclLAST_HYP simplest_case)) +let old_destruct_id s = + (tclTHEN (intros_until_id s) (tclLAST_HYP simplest_case)) +let old_destruct_nodep n = + (tclTHEN (intros_until_n n) (tclLAST_HYP simplest_case)) -let dyn_new_destruct = function - | [(Command c)] -> tactic_com (new_induct false) c - | [(Constr x)] -> new_induct false x - (* Identifier apart because id can be quantified in goal and not typable *) - | [Integer _ | Identifier _ as arg] -> - tactic_try_intros_until (fun id -> new_induct false (mkVar id)) arg - | l -> bad_tactic_args "induct" l - -let dyn_old_destruct = function - | [Identifier x] -> destruct x - | [Integer n] -> destruct_nodep n - | l -> bad_tactic_args "destruct" l - -let dyn_destruct = dyn_old_destruct +let old_destruct = function + | NamedHyp id -> old_destruct_id id + | AnonHyp n -> old_destruct_nodep n (* * Eliminations giving the type instead of the proof. @@ -1695,22 +1428,12 @@ let elim_type t gl = let elimc = Indrec.lookup_eliminator ind (elimination_sort_of_goal gl) in elim_scheme_type elimc t gl -let dyn_elim_type = function - | [Constr c] -> elim_type c - | [Command com] -> tactic_com_sort elim_type com - | l -> bad_tactic_args "elim_type" l - let case_type t gl = let (ind,t) = pf_reduce_to_atomic_ind gl t in let env = pf_env gl in let elimc = Indrec.make_case_gen env (project gl) ind (elimination_sort_of_goal gl) in elim_scheme_type elimc t gl -let dyn_case_type = function - | [Constr c] -> case_type c - | [Command com] -> tactic_com case_type com - | l -> bad_tactic_args "case_type" l - (* Some eliminations frequently used *) @@ -1750,14 +1473,15 @@ let orE id gl = let dorE b cls gl = match cls with | (Some id) -> orE id gl - | None -> (if b then right else left) [] gl + | None -> (if b then right else left) NoBindings gl let impE id gl = let t = pf_get_hyp_typ gl id in if is_imp_term (pf_hnf_constr gl t) then let (dom, _, rng) = destProd (pf_hnf_constr gl t) in - (tclTHENS (cut_intro rng) - [tclIDTAC;apply_term (mkVar id) [mkMeta (new_meta())]]) gl + tclTHENLAST + (cut_intro rng) + (apply_term (mkVar id) [mkMeta (new_meta())]) gl else errorlabstrm "impE" (str("Tactic impE expects "^(string_of_id id)^ @@ -1772,55 +1496,15 @@ let dImp cls gl = (* Tactics related with logic connectives *) (************************************************) -(* Contradiction *) - -let contradiction_on_hyp id gl = - let hyp = pf_get_hyp_typ gl id in - if is_empty_type hyp then - simplest_elim (mkVar id) gl - else - error "Not a contradiction" - -(* Absurd *) -let absurd c gls = - (tclTHENS - (tclTHEN (elim_type (build_coq_False ())) (cut c)) - ([(tclTHENS - (cut (applist(build_coq_not (),[c]))) - ([(tclTHEN intros - ((fun gl -> - let ida = pf_nth_hyp_id gl 1 - and idna = pf_nth_hyp_id gl 2 in - exact_no_check (applist(mkVar idna,[mkVar ida])) gl))); - tclIDTAC])); - tclIDTAC])) gls - -let dyn_absurd = function - | [Constr c] -> absurd c - | [Command com] -> tactic_com_sort absurd com - | l -> bad_tactic_args "absurd" l - -let contradiction gls = - tclTHENLIST [ intros; elim_type (build_coq_False ()); assumption ] gls - -let dyn_contradiction = function - | [] -> contradiction - | l -> bad_tactic_args "contradiction" l - -(* Relfexivity tactics *) +(* Reflexivity tactics *) let reflexivity gl = match match_with_equation (pf_concl gl) with | None -> error "The conclusion is not a substitutive equation" - | Some (hdcncl,args) -> one_constructor 1 [] gl + | Some (hdcncl,args) -> one_constructor 1 NoBindings gl let intros_reflexivity = (tclTHEN intros reflexivity) -let dyn_reflexivity = function - | [] -> intros_reflexivity - | _ -> errorlabstrm "Tactics.reflexivity" - (str "Tactic applied to bad arguments!") - (* Symmetry tactics *) (* This tactic first tries to apply a constant named sym_eq, where eq @@ -1842,19 +1526,16 @@ let symmetry gl = | [c1;c2] -> mkApp (hdcncl, [| c2; c1 |]) | _ -> assert false in - (tclTHENS (cut symc) - [ tclTHENLIST [ intro; - tclLAST_HYP simplest_case; - one_constructor 1 [] ]; - tclIDTAC ]) gl + tclTHENLAST (cut symc) + (tclTHENLIST + [ intro; + tclLAST_HYP simplest_case; + one_constructor 1 NoBindings ]) + gl end let intros_symmetry = (tclTHEN intros symmetry) -let dyn_symmetry = function - | [] -> intros_symmetry - | l -> bad_tactic_args "symmetry" l - (* Transitivity tactics *) (* This tactic first tries to apply a constant named trans_eq, where eq @@ -1886,22 +1567,16 @@ let transitivity t gl = | [c1;c2] -> mkApp (hdcncl, [| t; c2 |]) | _ -> assert false in - (tclTHENS (cut eq2) - [tclTHENS (cut eq1) - [ tclTHENLIST [ tclDO 2 intro; - tclLAST_HYP simplest_case; - assumption ]; - tclIDTAC]; - tclIDTAC])gl + tclTHENFIRST (cut eq2) + (tclTHENFIRST (cut eq1) + (tclTHENLIST + [ tclDO 2 intro; + tclLAST_HYP simplest_case; + assumption ])) gl end let intros_transitivity n = tclTHEN intros (transitivity n) -let dyn_transitivity = function - | [Constr n] -> intros_transitivity n - | [Command n] -> tactic_com intros_transitivity n - | l -> bad_tactic_args "transitivity" l - (* tactical to save as name a subproof such that the generalisation of the current goal, abstracted with respect to the local signature, is solved by tac *) @@ -1922,8 +1597,8 @@ let abstract_subproof name tac gls = in if occur_existential concl then error "Abstract cannot handle existentials"; let lemme = - start_proof na NeverDischarge current_sign concl; - let _,(const,strength) = + start_proof na (false,Nametab.NeverDischarge) current_sign concl (fun _ _ -> ()); + let _,(const,(_,strength),_) = try by (tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac)); let r = cook_proof () in @@ -1947,15 +1622,3 @@ let tclABSTRACT name_op tac gls = | None -> add_suffix (get_current_proof_name ()) "_subproof" in abstract_subproof s tac gls - -let dyn_tclABSTRACT = - hide_tactic "ABSTRACT" - (function - | [Tac (tac,_)] -> - tclABSTRACT None tac - | [Identifier s; Tac (tac,_)] -> - tclABSTRACT (Some s) tac - | _ -> invalid_arg "tclABSTRACT") - - - diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 76a21ba83..14d479362 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -21,8 +21,9 @@ open Evar_refiner open Clenv open Tacred open Tacticals +open Tacexpr open Nametab -(*i*) +open Rawterm (* Main tactics. *) @@ -34,7 +35,7 @@ val type_clenv_binding : named_context sigma -> val string_of_inductive : constr -> string val head_constr : constr -> constr list val head_constr_bound : constr -> constr list -> constr list -val bad_tactic_args : string -> tactic_arg list -> 'a +val is_quantified_hypothesis : identifier -> goal sigma -> bool exception Bound @@ -47,12 +48,9 @@ val convert_hyp : named_declaration -> tactic val thin : identifier list -> tactic val mutual_fix : identifier -> int -> (identifier * int * constr) list -> tactic -val fix : identifier -> int -> tactic +val fix : identifier option -> int -> tactic val mutual_cofix : identifier -> (identifier * constr) list -> tactic -val cofix : identifier -> tactic - -val dyn_mutual_fix : tactic_arg list -> tactic -val dyn_mutual_cofix : tactic_arg list -> tactic +val cofix : identifier option -> tactic (*s Introduction tactics. *) @@ -62,9 +60,7 @@ val fresh_id : identifier list -> identifier -> goal sigma -> identifier val intro : tactic val introf : tactic val intro_force : bool -> tactic -val dyn_intro : tactic_arg list -> tactic - -val dyn_intro_move : tactic_arg list -> tactic +val intro_move : identifier option -> identifier option -> tactic val intro_replacing : identifier -> tactic val intro_using : identifier -> tactic @@ -75,42 +71,31 @@ val intros_replacing : identifier list -> tactic val intros : tactic -(*i Obsolete, subsumed by Elim.dyn_intro_pattern -val dyn_intros_using : tactic_arg list -> tactic -i*) +(* [depth_of_quantified_hypothesis b h g] returns the index of [h] in + the conclusion of goal [g], up to head-reduction if [b] is [true] *) +val depth_of_quantified_hypothesis : + bool -> quantified_hypothesis -> goal sigma -> int -val intros_until : identifier -> tactic -val intros_until_n : int -> tactic val intros_until_n_wored : int -> tactic -val dyn_intros_until : tactic_arg list -> tactic +val intros_until : quantified_hypothesis -> tactic val intros_clearing : bool list -> tactic (* Assuming a tactic [tac] depending on an hypothesis identifier, - [tactic_try_intros_until tac arg] first assumes that arg denotes a + [try_intros_until tac arg] first assumes that arg denotes a quantified hypothesis (denoted by name or by index) and try to introduce it in context before to apply [tac], otherwise assume the hypothesis is already in context and directly apply [tac] *) -val tactic_try_intros_until : (identifier,tactic_arg) parse_combinator - -(* Assuming a tactic [tac] depending on an hypothesis identifier, - [hide_ident_or_numarg_tactic str tac] registers a tactic which - compose [tac] with "Intros Until" and returns a tactic which - behaves as [tac] (without implicit "Intros until") but hiding the - implementation under the name [str] *) -val hide_ident_or_numarg_tactic : identifier hide_combinator +val try_intros_until : + (identifier -> tactic) -> quantified_hypothesis -> tactic (*s Exact tactics. *) val assumption : tactic -val dyn_assumption : tactic_arg list -> tactic - val exact_no_check : constr -> tactic -val dyn_exact_no_check : tactic_arg list -> tactic - val exact_check : constr -> tactic -val dyn_exact_check : tactic_arg list -> tactic +val exact_proof : Coqast.t -> tactic (*s Reduction tactics. *) @@ -143,28 +128,21 @@ val unfold_option : (int list * Closure.evaluable_global_reference) list -> hyp_location option -> tactic val reduce : red_expr -> hyp_location list -> tactic -val dyn_reduce : tactic_arg list -> tactic -val dyn_change : tactic_arg list -> tactic +val change : constr -> hyp_location list -> tactic val unfold_constr : global_reference -> tactic -val pattern_option : - (int list * constr * constr) list -> hyp_location option -> tactic +val pattern_option : (int list * constr) list -> hyp_location option -> tactic (*s Modification of the local context. *) val clear : identifier list -> tactic -val dyn_clear : tactic_arg list -> tactic val clear_body : identifier list -> tactic -val dyn_clear_body : tactic_arg list -> tactic val new_hyp : int option ->constr -> constr substitution -> tactic -val dyn_new_hyp : tactic_arg list -> tactic - -val dyn_move : tactic_arg list -> tactic -val dyn_move_dep : tactic_arg list -> tactic -val dyn_rename : tactic_arg list -> tactic +val move_hyp : bool -> identifier -> identifier -> tactic +val rename_hyp : identifier -> identifier -> tactic (*s Resolution tactics. *) @@ -175,47 +153,37 @@ val bring_hyps : named_context -> tactic val apply : constr -> tactic val apply_without_reduce : constr -> tactic val apply_list : constr list -> tactic -val apply_with_bindings : (constr * constr substitution) -> tactic -val dyn_apply : tactic_arg list -> tactic +val apply_with_bindings : constr with_bindings -> tactic val cut_and_apply : constr -> tactic -val dyn_cut_and_apply : tactic_arg list -> tactic (*s Elimination tactics. *) -val general_elim : constr * constr substitution -> - constr * constr substitution -> tactic -val default_elim : constr * constr substitution -> tactic -val simplest_elim : constr -> tactic -val dyn_elim : tactic_arg list -> tactic +val general_elim : constr with_bindings -> constr with_bindings -> tactic +val default_elim : constr with_bindings -> tactic +val simplest_elim : constr -> tactic +val elim : constr with_bindings -> constr with_bindings option -> tactic +val general_elim_in : identifier -> constr * constr substitution -> + constr * constr substitution -> tactic +val old_induct : quantified_hypothesis -> tactic val general_elim_in : identifier -> constr * constr substitution -> constr * constr substitution -> tactic -val old_induct : identifier -> tactic -val old_induct_nodep : int -> tactic -val dyn_old_induct : tactic_arg list -> tactic -val dyn_new_induct : tactic_arg list -> tactic +val new_induct : constr induction_arg -> tactic (*s Case analysis tactics. *) -val general_case_analysis : constr * constr substitution -> tactic +val general_case_analysis : constr with_bindings -> tactic val simplest_case : constr -> tactic -val dyn_case : tactic_arg list -> tactic -val destruct : identifier -> tactic -val destruct_nodep : int -> tactic -val dyn_destruct : tactic_arg list -> tactic -val dyn_new_destruct : tactic_arg list -> tactic +val old_destruct : quantified_hypothesis -> tactic +val new_destruct : constr induction_arg -> tactic (*s Eliminations giving the type instead of the proof. *) val case_type : constr -> tactic -val dyn_case_type : tactic_arg list -> tactic - val elim_type : constr -> tactic -val dyn_elim_type : tactic_arg list -> tactic - (*s Some eliminations which are frequently used. *) @@ -232,8 +200,7 @@ val dorE : bool -> clause ->tactic val constructor_tac : int option -> int -> constr substitution -> tactic val one_constructor : int -> constr substitution -> tactic -val any_constructor : tactic -val tclConstrThen : tactic -> tactic +val any_constructor : tactic option -> tactic val left : constr substitution -> tactic val simplest_left : tactic val right : constr substitution -> tactic @@ -241,45 +208,27 @@ val simplest_right : tactic val split : constr substitution -> tactic val simplest_split : tactic -val dyn_constructor : tactic_arg list -> tactic -val dyn_left : tactic_arg list -> tactic -val dyn_right : tactic_arg list -> tactic -val dyn_split : tactic_arg list -> tactic - (*s Logical connective tactics. *) -val absurd : constr -> tactic -val dyn_absurd : tactic_arg list -> tactic - -val contradiction_on_hyp : identifier -> tactic -val contradiction : tactic -val dyn_contradiction : tactic_arg list -> tactic - -val reflexivity : tactic +val reflexivity : tactic val intros_reflexivity : tactic -val dyn_reflexivity : tactic_arg list -> tactic - + val symmetry : tactic val intros_symmetry : tactic -val dyn_symmetry : tactic_arg list -> tactic val transitivity : constr -> tactic val intros_transitivity : constr -> tactic -val dyn_transitivity : tactic_arg list -> tactic val cut : constr -> tactic val cut_intro : constr -> tactic val cut_replacing : identifier -> constr -> tactic val cut_in_parallel : constr list -> tactic -val true_cut : identifier -> constr -> tactic -val true_cut_anon : constr -> tactic -val dyn_cut : tactic_arg list -> tactic -val dyn_true_cut : tactic_arg list -> tactic -val dyn_lettac : tactic_arg list -> tactic -val dyn_forward : tactic_arg list -> tactic +val true_cut : identifier option -> constr -> tactic +val letin_tac : bool -> name -> constr -> + identifier clause_pattern -> tactic +val forward : bool -> name -> constr -> tactic val generalize : constr list -> tactic -val dyn_generalize : tactic_arg list -> tactic -val dyn_generalize_dep : tactic_arg list -> tactic +val generalize_dep : constr -> tactic val tclABSTRACT : identifier option -> tactic -> tactic diff --git a/tactics/wcclausenv.ml b/tactics/wcclausenv.ml index 007a3aa99..78f3890c5 100644 --- a/tactics/wcclausenv.ml +++ b/tactics/wcclausenv.ml @@ -18,6 +18,7 @@ open Sign open Reductionops open Environ open Logic +open Refiner open Tacmach open Evd open Proof_trees @@ -99,10 +100,12 @@ let clenv_constrain_with_bindings bl clause = matchrec clause bl (* What follows is part of the contents of the former file tactics3.ml *) - +(* 2/2002: replaced THEN_i by THENSLAST to solve a bug in + Tacticals.general_elim when the eliminator has missing bindings *) + let elim_res_pf_THEN_i kONT clenv tac gls = let clenv' = (clenv_unique_resolver true clenv gls) in - tclTHEN_i (clenv_refine kONT clenv') (tac clenv') gls + tclTHENLASTn (clenv_refine kONT clenv') (tac clenv') gls let rec build_args acc ce p_0 p_1 = match kind_of_term p_0, p_1 with diff --git a/tactics/wcclausenv.mli b/tactics/wcclausenv.mli index 26c46e89a..45b955326 100644 --- a/tactics/wcclausenv.mli +++ b/tactics/wcclausenv.mli @@ -45,6 +45,6 @@ val add_prods_sign : **i*) val elim_res_pf_THEN_i : - (wc -> tactic) -> wc clausenv -> (wc clausenv -> int -> tactic) -> tactic + (wc -> tactic) -> wc clausenv -> (wc clausenv -> tactic array) -> tactic val applyUsing : constr -> tactic diff --git a/toplevel/cerrors.ml b/toplevel/cerrors.ml index 7f20f90c7..1fbc3a90e 100644 --- a/toplevel/cerrors.ml +++ b/toplevel/cerrors.ml @@ -84,7 +84,7 @@ let rec explain_exn_default = function | Nametab.GlobalizationConstantError q -> hov 0 (str "Error:" ++ spc () ++ str "No constant of this name:" ++ spc () ++ Nametab.pr_qualid q) - | Tacmach.FailError i -> + | Refiner.FailError i -> hov 0 (str "Error: Fail tactic always fails (level " ++ int i ++ str").") | Stdpp.Exc_located (loc,exc) -> diff --git a/toplevel/class.ml b/toplevel/class.ml index 3096b6067..8bac0812e 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -329,6 +329,16 @@ let try_add_new_identity_coercion id stre ~source ~target = let try_add_new_coercion_with_source ref stre ~source = try_add_new_coercion_core ref stre (Some source) None false +let add_coercion_hook stre ref = + try_add_new_coercion ref stre; + Options.if_verbose message + (string_of_qualid (shortest_qualid_of_global (Global.env()) ref) + ^ " is now a coercion") + +let add_subclass_hook stre ref = + let cl = class_of_ref ref in + try_add_new_coercion_subclass cl stre + (* try_add_new_class : global_reference -> strength -> unit *) let class_of_global = function diff --git a/toplevel/class.mli b/toplevel/class.mli index f140351ce..388f664e4 100644 --- a/toplevel/class.mli +++ b/toplevel/class.mli @@ -44,6 +44,10 @@ val try_add_new_coercion_with_source : global_reference -> strength -> val try_add_new_identity_coercion : identifier -> strength -> source:cl_typ -> target:cl_typ -> unit +val add_coercion_hook : Proof_type.declaration_hook + +val add_subclass_hook : Proof_type.declaration_hook + (* [try_add_new_class ref] declares [ref] as a new class; usually, this is done implicitely by [try_add_new_coercion]'s functions *) val try_add_new_class : global_reference -> strength -> unit diff --git a/toplevel/command.ml b/toplevel/command.ml index 9c1192112..234910af0 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -19,6 +19,7 @@ open Environ open Reduction open Tacred open Declare +open Nametab open Names open Nameops open Coqast @@ -74,8 +75,8 @@ let declare_global_definition ident ce n local = if_verbose message ((string_of_id ident) ^ " is defined"); ConstRef sp -let definition_body_red red_option ident (local,n) com comtypeopt = - let ce = constant_entry_of_com (com,comtypeopt,false) in +let declare_definition red_option ident (local,n) c typopt = + let ce = constant_entry_of_com (c,typopt,false) in let ce' = red_constant_entry ce red_option in match n with | NeverDischarge -> declare_global_definition ident ce' n local @@ -95,42 +96,36 @@ let definition_body_red red_option ident (local,n) com comtypeopt = anomalylabstrm "Command.definition_body_red" (str "Strength NotDeclare not for Definition, only for Let") -let definition_body = definition_body_red None - -let syntax_definition ident com = - let c = interp_rawconstr Evd.empty (Global.env()) com in +let syntax_definition ident c = + let c = interp_rawconstr Evd.empty (Global.env()) c in Syntax_def.declare_syntactic_definition ident c; if_verbose message ((string_of_id ident) ^ " is now a syntax macro") -(* 2| Variable definitions *) +(* 2| Variable/Hypothesis/Parameter/Axiom declarations *) -let parameter_def_var ident c = - let c = interp_type Evd.empty (Global.env()) c in - let sp = declare_constant ident (ParameterEntry c, NeverDischarge) in - if_verbose message ((string_of_id ident) ^ " is assumed"); - sp - -let declare_global_assumption ident c = - let sp = parameter_def_var ident c in - msg_warning (pr_id ident ++ str" is declared as a parameter" ++ - str" because it is at a global level"); - ConstRef sp +let assumption_message id = + if_verbose message ((string_of_id id) ^ " is assumed") -let hypothesis_def_var is_refining ident n c = +let declare_assumption ident n c = + let c = interp_type Evd.empty (Global.env()) c in match n with - | NeverDischarge -> declare_global_assumption ident c + | NeverDischarge -> + let _ = declare_constant ident (ParameterEntry c, NeverDischarge) in + assumption_message ident | DischargeAt (disch_sp,_) -> if Lib.is_section_p disch_sp then begin - let t = interp_type Evd.empty (Global.env()) c in - let sp = declare_variable ident (Lib.cwd(),SectionLocalAssum t,n) in - if_verbose message ((string_of_id ident) ^ " is assumed"); - if is_refining then + let _ = declare_variable ident (Lib.cwd(),SectionLocalAssum c,n) in + assumption_message ident; + if is_verbose () & Pfedit.refining () then msgerrnl (str"Warning: Variable " ++ pr_id ident ++ - str" is not visible from current goals"); - VarRef ident + str" is not visible from current goals") end else - declare_global_assumption ident c + let _ = declare_constant ident (ParameterEntry c, NeverDischarge) in + assumption_message ident; + if_verbose + msg_warning (pr_id ident ++ str" is declared as a parameter" ++ + str" because it is at a global level"); | NotDeclare -> anomalylabstrm "Command.hypothesis_def_var" (str "Strength NotDeclare not for Variable, only for Let") @@ -221,9 +216,41 @@ let declare_mutual_with_eliminations mie = Indrec.declare_eliminations sp; sp -let build_mutual lparams lnamearconstrs finite = - let mie = interp_mutual lparams lnamearconstrs finite in - let _ = declare_mutual_with_eliminations mie in () +let eq_la (id,ast) (id',ast') = id = id' & alpha_eq(ast,ast') + +let extract_coe lc = + List.fold_right + (fun (id,addcoe,t) (l1,l2) -> + ((if addcoe then id::l1 else l1), (id,t)::l2)) lc ([],[]) + +let extract_coe_la_lc = function + | [] -> anomaly "Vernacentries: empty list of inductive types" + | (id,la,ar,lc)::rest -> + let rec check = function + | [] -> [],[] + | (id,la',ar,lc)::rest -> + if (List.length la = List.length la') && + (List.for_all2 eq_la la la') + then + let mcoes, mspec = check rest in + let coes, lc' = extract_coe lc in + (coes::mcoes,(id,ar,lc')::mspec) + else + error ("Parameters should be syntactically the same "^ + "for each inductive type") + in + let mcoes, mspec = check rest in + let coes, lc' = extract_coe lc in + (coes,la,(id,ar,lc'):: mspec) + +let build_mutual lind finite = + let (coes,lparams,lnamearconstructs) = extract_coe_la_lc lind in + let mie = interp_mutual lparams lnamearconstructs finite in + let _ = declare_mutual_with_eliminations mie in + List.iter + (fun id -> + Class.try_add_new_coercion + (locate (make_short_qualid id)) NeverDischarge) coes (* try to find non recursive definitions *) @@ -389,13 +416,6 @@ let build_corecursive lnameardef = in () -let inductive_of_ident qid = - match Nametab.global dummy_loc qid with - | IndRef ind -> ind - | ref -> errorlabstrm "inductive_of_ident" - (pr_id (id_of_global (Global.env()) ref) ++ - spc () ++ str "is not an inductive type") - let build_scheme lnamedepindsort = let lrecnames = List.map (fun (f,_,_,_) -> f) lnamedepindsort and sigma = Evd.empty @@ -403,7 +423,7 @@ let build_scheme lnamedepindsort = let lrecspec = List.map (fun (_,dep,indid,sort) -> - let ind = inductive_of_ident indid in + let ind = Nametab.global_inductive indid in let (mib,mip) = Global.lookup_inductive ind in (ind,mib,mip,dep,interp_elimination_sort sort)) lnamedepindsort @@ -420,7 +440,7 @@ let build_scheme lnamedepindsort = let lrecref = List.fold_right2 declare listdecl lrecnames [] in if_verbose ppnl (recursive_message (Array.of_list lrecref)) -let start_proof_com sopt stre com = +let start_proof_com sopt (local,stre) com hook = let env = Global.env () in let sign = Global.named_context () in let id = match sopt with @@ -435,28 +455,28 @@ let start_proof_com sopt stre com = in let c = interp_type Evd.empty env com in let _ = Typeops.infer_type env c in - Pfedit.start_proof id stre sign c + Pfedit.start_proof id (local,stre) sign c hook let apply_tac_not_declare id pft = function | None -> error "Type of Let missing" | Some typ -> - let cutt = vernac_tactic ("Cut",[Constr typ]) - and exat = vernac_tactic ("Exact",[Constr pft]) in + let cutt = Hiddentac.h_cut typ + and exat = Hiddentac.h_exact pft in Pfedit.delete_current_proof (); - Pfedit.by (tclTHENS cutt [introduction id;exat]) + Pfedit.by (Refiner.tclTHENSV cutt [|introduction id;exat|]) -let save id const strength = +let save id const strength hook = let {const_entry_body = pft; const_entry_type = tpo; const_entry_opaque = opacity } = const in begin match strength with | DischargeAt (disch_sp,_) when Lib.is_section_p disch_sp && not opacity -> let c = SectionLocalDef (pft, tpo) in - let _ = declare_variable id (Lib.cwd(), c, strength) - in () + let _ = declare_variable id (Lib.cwd(), c, strength) in + hook strength (VarRef id) | NeverDischarge | DischargeAt _ -> - let _ = declare_constant id (ConstantEntry const,strength) - in () + let ref = ConstRef (declare_constant id (ConstantEntry const,strength)) in + hook strength ref | NotDeclare -> apply_tac_not_declare id pft tpo end; if not (strength = NotDeclare) then @@ -466,9 +486,9 @@ let save id const strength = end let save_named opacity = - let id,(const,strength) = Pfedit.cook_proof () in + let id,(const,(local,strength),hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in - save id const strength + save id const strength hook let check_anonymity id save_ident = if atompart_of_id id <> "Unnamed_thm" then @@ -478,16 +498,16 @@ let check_anonymity id save_ident = *) let save_anonymous opacity save_ident = - let id,(const,strength) = Pfedit.cook_proof () in + let id,(const,(local,strength),hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in check_anonymity id save_ident; - save save_ident const strength + save save_ident const strength hook let save_anonymous_with_strength strength opacity save_ident = - let id,(const,_) = Pfedit.cook_proof () in + let id,(const,_,hook) = Pfedit.cook_proof () in let const = { const with const_entry_opaque = opacity } in check_anonymity id save_ident; - save save_ident const strength + save save_ident const strength hook let get_current_context () = try Pfedit.get_current_goal_context () diff --git a/toplevel/command.mli b/toplevel/command.mli index baa983788..474fe9202 100644 --- a/toplevel/command.mli +++ b/toplevel/command.mli @@ -9,11 +9,11 @@ (*i $Id$ i*) (*i*) +open Util open Names open Term -open Declare -open Library open Nametab +open Library (*i*) (*s Declaration functions. The following functions take ASTs, @@ -21,6 +21,7 @@ open Nametab functions of [Declare]; they return an absolute reference to the defined object *) +(*i val constant_entry_of_com : Coqast.t * Coqast.t option * bool -> Safe_typing.constant_entry @@ -31,24 +32,23 @@ val declare_global_definition : val definition_body : identifier -> bool * strength -> Coqast.t -> Coqast.t option -> global_reference +i*) -val definition_body_red : Tacred.red_expr option -> identifier +val declare_definition : Tacred.red_expr option -> identifier -> bool * strength -> Coqast.t -> Coqast.t option -> global_reference val syntax_definition : identifier -> Coqast.t -> unit (*i -val abstraction_definition : identifier -> int array -> Coqast.t -> unit -i*) - val hypothesis_def_var : bool -> identifier -> strength -> Coqast.t -> global_reference val parameter_def_var : identifier -> Coqast.t -> constant +i*) + +val declare_assumption : identifier -> strength -> Coqast.t -> unit -val build_mutual : - (identifier * Coqast.t) list -> - (identifier * Coqast.t * (identifier * Coqast.t) list) list -> bool -> unit +val build_mutual : Vernacexpr.inductive_expr list -> bool -> unit val declare_mutual_with_eliminations : Indtypes.mutual_inductive_entry -> section_path @@ -59,9 +59,9 @@ val build_recursive : val build_corecursive : (identifier * Coqast.t * Coqast.t) list -> unit -val build_scheme : (identifier * bool * Nametab.qualid * Coqast.t) list -> unit +val build_scheme : (identifier * bool * Nametab.qualid located * Coqast.t) list -> unit -val start_proof_com : identifier option -> strength -> Coqast.t -> unit +val start_proof_com : identifier option -> bool * strength -> Coqast.t -> Proof_type.declaration_hook -> unit (*s [save_named b] saves the current completed proof under the name it was started; boolean [b] tells if the theorem is declared opaque; it diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index c70bab141..a62c1fae4 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -67,10 +67,7 @@ let load_vernac_obj () = let require_list = ref ([] : string list) let add_require s = require_list := s :: !require_list let require () = - List.iter - (fun s -> - let qid = Nametab.make_short_qualid (id_of_string (Filename.basename s)) in - Library.require_module_from_file None qid s false) + List.iter (fun s -> Library.require_module_from_file None None s false) (List.rev !require_list) let compile_list = ref ([] : (bool * string) list) diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml index 9ee3a5621..14e4feff5 100644 --- a/toplevel/discharge.ml +++ b/toplevel/discharge.ml @@ -22,6 +22,7 @@ open Cooking open Typeops open Libobject open Lib +open Nametab open Declare open Impargs open Classops diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 4041d311c..c2892a74b 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -381,10 +381,6 @@ let explain_refiner_cannot_generalize ty = let explain_refiner_not_well_typed c = str"The term " ++ prterm c ++ str" is not well-typed" -let explain_refiner_bad_tactic_args s l = - str "Internal tactic " ++ str s ++ str " cannot be applied to " ++ - Tacmach.pr_tactic (s,l) - let explain_intro_needs_product () = str "Introduction tactics needs products" @@ -404,7 +400,6 @@ let explain_refiner_error = function | CannotUnify (m,n) -> explain_refiner_cannot_unify m n | CannotGeneralize ty -> explain_refiner_cannot_generalize ty | NotWellTyped c -> explain_refiner_not_well_typed c - | BadTacticArgs (s,l) -> explain_refiner_bad_tactic_args s l | IntroNeedsProduct -> explain_intro_needs_product () | DoesNotOccurIn (c,hyp) -> explain_does_not_occur_in c hyp | NonLinearProof c -> explain_non_linear_proof c diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml index 487c87bd4..cabdee95c 100644 --- a/toplevel/metasyntax.ml +++ b/toplevel/metasyntax.ml @@ -12,33 +12,41 @@ open Pp open Util open Coqast open Ast -open Pcoq open Extend open Esyntax open Libobject open Library open Summary +open Astterm +open Vernacexpr +open Pcoq (************************* **** PRETTY-PRINTING **** *************************) -(* Done here to get parsing/g_*.ml4 non dependent from kernel *) -let constr_parser_with_glob = map_entry Astterm.globalize_constr Constr.constr -let tactic_parser_with_glob = map_entry Astterm.globalize_ast Tactic.tactic -let vernac_parser_with_glob = map_entry Astterm.globalize_ast Vernac_.vernac +let globalize_typed_ast t = + let sign = Global.named_context () in + match t with + | Ast.PureAstNode t -> Ast.PureAstNode (globalize_constr t) + | _ -> (* TODO *) t (* This updates default parsers for Grammar actions and Syntax *) (* patterns by inserting globalization *) -let _ = update_constr_parser constr_parser_with_glob -let _ = update_tactic_parser tactic_parser_with_glob -let _ = update_vernac_parser vernac_parser_with_glob +(* Done here to get parsing/g_*.ml4 non dependent from kernel *) +let _ = Pcoq.set_globalizer globalize_typed_ast (* This installs default quotations parsers to escape the ast parser *) (* "constr" is used by default in quotations found in the ast parser *) +let constr_parser_with_glob = Pcoq.map_entry Astterm.globalize_constr Constr.constr +let tactic_parser_with_glob = (*map_entry Astterm.globalize_tactic*) Tactic.tactic +let vernac_parser_with_glob = (*map_entry Astterm.globalize_vernac*) Vernac_.vernac + let _ = define_quotation true "constr" constr_parser_with_glob +(* let _ = define_quotation false "tactic" tactic_parser_with_glob let _ = define_quotation false "vernac" vernac_parser_with_glob +*) (* Pretty-printer state summary *) let _ = @@ -105,16 +113,30 @@ let (inGrammar, outGrammar) = cache_function = cache_grammar; export_function = (fun x -> Some x)}) -let add_grammar_obj univ al = - Lib.add_anonymous_leaf (inGrammar (Extend.interp_grammar_command univ al)) +let gram_define_entry (u,_ as univ) ((ntl,nt),et,assoc,rl) = + let etyp = match et with None -> entry_type_from_name u | Some e -> e in + create_entry_if_new univ nt etyp; + let etyp = match etyp with + | AstListType -> ETastl + | GenAstType Genarg.ConstrArgType -> ETast + | PureAstType -> ETast + | _ -> error "Cannot arbitrarily extend non ast entries" in + (nt, etyp, assoc, rl) + +let add_grammar_obj univ l = + let u = create_univ_if_new univ in + let entryl = List.map (gram_define_entry u) l in + let g = interp_grammar_command univ get_entry_type entryl in + Lib.add_anonymous_leaf (inGrammar (Egrammar.AstGrammar g)) + +let add_tactic_grammar g = + Lib.add_anonymous_leaf (inGrammar (Egrammar.TacticGrammar g)) (* printing grammar entries *) let print_grammar univ entry = let u = get_univ univ in let te = get_entry u entry in - match te with - | Ast e -> Gram.Entry.print e - | ListAst e -> Gram.Entry.print e + Gram.Entry.print (object_of_typed_entry te) (* Infix, distfix *) @@ -150,135 +172,154 @@ let (inInfix, outInfix) = cache_function = cache_infix; export_function = (fun x -> Some x)}) +(* Build the syntax and grammar rules *) + +type symbol = + | Terminal of string + | NonTerminal of (int * parenRelation) * string + let prec_assoc = function - | Some(Gramext.RightA) -> (":L",":E") - | Some(Gramext.LeftA) -> (":E",":L") - | Some(Gramext.NonA) -> (":L",":L") - | None -> (":E",":L") (* LEFTA by default *) + | Some(Gramext.RightA) -> (L,E) + | Some(Gramext.LeftA) -> (E,L) + | Some(Gramext.NonA) -> (L,L) + | None -> (E,L) (* LEFTA by default *) let constr_tab = [| "constr0"; "constr1"; "constr2"; "constr3"; "lassoc_constr4"; "constr5"; "constr6"; "constr7"; "constr8"; "constr9"; "constr10" |] -let constr_rule n p = if p = ":E" then constr_tab.(n) else constr_tab.(n-1) +let constr_rule (n,p) = if p = E then constr_tab.(n) else constr_tab.(n-1) -let nonterm_meta nt var = NonTerm(NtShort nt, ETast, Some var) +let nonterm_meta nt var = + NonTerm(ProdPrimitive ("constr",nt), Some (var,ETast)) -(* infix_syntax_entry : int -> string -> string -> grammar_entry - * takes precedence, infix op and prefix op and returns - * the corresponding Syntax entry *) +let meta_pattern m = Pmeta(m,Tany) -let infix_syntax_entry assoc n inf prefname astpref = - let (lp,rp) = match assoc with - | Some(Gramext.RightA) -> (Extend.L,Extend.E) - | Some(Gramext.LeftA) -> (Extend.E,Extend.L) - | Some(Gramext.NonA) -> (Extend.L,Extend.L) - | None -> (Extend.E,Extend.L) (* LEFTA by default *) - in - let inf = - (* Not necessary but increases legibility (e.g. for "=_S") *) - if is_letter (inf.[String.length inf -1]) then " "^inf^" " else inf - in - [{Extend.syn_id = prefname; - Extend.syn_prec = n,0,0; - Extend.syn_astpat = - Pnode - ("APPLIST", - Pcons - (Pquote astpref, - Pcons (Pmeta ("$e1", Tany), Pcons (Pmeta ("$e2", Tany), Pnil)))); - Extend.syn_hunks = - [Extend.UNP_BOX - (Extend.PpHOVB 1, - [Extend.PH (Pmeta ("$e1", Tany), None, lp); - Extend.UNP_BRK (0, 1); Extend.RO inf; - Extend.PH (Pmeta ("$e2", Tany), None, rp)])]}] - -(* infix_gram_entry : int -> string -> string -> grammar_entry - * takes precedence, infix op and prefix op and returns - * the corresponding Grammar entry *) - -let gram_infix assoc n infl prefname astpref = - let (lp,rp) = prec_assoc assoc in - let action = - Aast(Pnode("APPLIST", - Pcons(Pquote astpref, - Pcons(Pmeta("$a",Tany), - Pcons(Pmeta("$b",Tany),Pnil))))) - in +let collect_metas sl = + List.fold_right + (fun it metatl -> match it with + | NonTerminal (_,m) -> Pcons(meta_pattern m, metatl) + | _ -> metatl) + sl Pnil + +let make_pattern astf symbols = + Pnode("APPLIST",Pcons(Pquote astf, collect_metas symbols)) + +let make_hunks symbols = + List.fold_right + (fun it l -> match it with + | NonTerminal ((_,lp),m) -> PH (meta_pattern m, None, lp) :: l + | Terminal s -> + let n,s = + if is_letter (s.[String.length s -1]) or is_letter (s.[0]) + then 1,s^" " else 0,s + in + UNP_BRK (n, 1) :: RO s :: l) + symbols [] + +let make_production = + List.map (function + | NonTerminal (lp,m) -> nonterm_meta (constr_rule lp) m + | Terminal s -> Term ("",s)) + +let make_syntax_rule n prefname symbols astpref = + [{syn_id = prefname; + syn_prec = n,0,0; + syn_astpat = make_pattern astpref symbols; + syn_hunks = [UNP_BOX (PpHOVB 1, make_hunks symbols)]}] + +let make_infix_syntax_rule n prefname symbols astpref ref = + let hunk = match symbols with + | [NonTerminal ((_,lp),e1);Terminal s;NonTerminal ((_,rp),e2)] -> + UNP_INFIX (ref,e1,e2,(lp,rp)) + | _ -> error "Don't know to handle this infix rule" in + [{syn_id = prefname; + syn_prec = n,0,0; + syn_astpat = make_pattern astpref symbols; + syn_hunks = [hunk]}] + + +let make_grammar_rule n fname symbols astpref = + let prod = make_production symbols in + let action = Act (PureAstPat (make_pattern astpref symbols)) in + Egrammar.AstGrammar { gc_univ = "constr"; gc_entries = - [ { ge_name = constr_rule n ":E"; + [ { ge_name = constr_rule (n, E); ge_type = ETast; gl_assoc = None; gl_rules = - [ { gr_name = prefname; - gr_production = - (nonterm_meta (constr_rule n lp) "$a") - ::(List.map (fun s -> Term("", s)) infl) - @[nonterm_meta (constr_rule n rp) "$b"]; + [ { gr_name = fname; + gr_production = prod; gr_action = action} ] } ] } -let add_infix assoc n inf pr = +let make_infix_symbols assoc n sl = + let (lp,rp) = prec_assoc assoc in + NonTerminal ((n,lp),"$a")::(List.map (fun s -> Terminal s) sl) + @[NonTerminal ((n,rp),"$b")] + + +let cache_infix2 (_,(ref,inf)) = + let sp = match ref with + | Nametab.TrueGlobal r -> Nametab.sp_of_global (Global.env()) r + | Nametab.SyntacticDef sp -> sp in + declare_infix_symbol sp inf + +let (inInfix2, outInfix2) = + declare_object + ("INFIX2", + { load_function = (fun _ -> ()); + open_function = cache_infix2; + cache_function = cache_infix2; + export_function = (fun x -> Some x)}) + +let add_infix assoc n inf (loc,qid) = + let ref = + try + Nametab.extended_locate qid + with Not_found -> + error ("Unknown reference: "^(Nametab.string_of_qualid qid)) in + let pr = Astterm.ast_of_extended_ref_loc loc ref in (* check the precedence *) if n<1 or n>10 then errorlabstrm "Metasyntax.infix_grammar_entry" (str"Precedence must be between 1 and 10."); +(* Pourquoi ?? if (assoc<>None) & (n<6 or n>9) then errorlabstrm "Vernacentries.infix_grammar_entry" (str"Associativity Precedence must be 6,7,8 or 9."); - (* check the grammar entry *) +*) let prefname = inf^"_infix" in - let gram_rule = gram_infix assoc n (split inf) prefname pr in - (* check the syntax entry *) - let syntax_rule = infix_syntax_entry assoc n inf prefname pr in - Lib.add_anonymous_leaf (inInfix(gram_rule,syntax_rule)) + let symbols = make_infix_symbols assoc n (split inf) in + let gram_rule = make_grammar_rule n prefname symbols pr in + let syntax_rule = make_infix_syntax_rule n prefname symbols pr ref in + Lib.add_anonymous_leaf (inInfix(gram_rule,syntax_rule)); + Lib.add_anonymous_leaf (inInfix2(ref,inf)) (* Distfix *) (* Distfix LEFTA 7 "/ _ / _ /" app.*) +(* TODO add boxes information in the expression *) -let rec make_symbols c_first c_next c_last new_var = function - | [] -> [] - | ["_"] -> [nonterm_meta c_last ("$e"^(string_of_int new_var))] +let create_meta n = "$e"^(string_of_int n) + +let rec find_symbols c_first c_next c_last new_var = function + | [] -> [] + | ["_"] -> [NonTerminal (c_last, create_meta new_var)] | ("_"::sl) -> - let symb = nonterm_meta c_first ("$e"^(string_of_int new_var)) in - symb :: make_symbols c_next c_next c_last (new_var+1) sl + let symb = NonTerminal (c_first, create_meta new_var) in + symb :: find_symbols c_next c_next c_last (new_var+1) sl | s :: sl -> - let symb = Term("", s) in - symb :: make_symbols c_next c_next c_last new_var sl - -let collect_metas sl = - List.fold_right - (fun it metatl -> - match it with - | NonTerm(_,_,Some m) -> Pcons(Pmeta(m,Tany),metatl) - | _ -> metatl) - sl Pnil + let symb = Terminal s in + symb :: find_symbols c_next c_next c_last new_var sl -let distfix assoc n sl fname astf = - let (lp,rp) = prec_assoc assoc in - let symbols = - make_symbols (constr_rule n lp) constr_tab.(8) (constr_rule n rp) 0 sl in - let action = - Aast(Pnode("APPLIST",Pcons(Pquote astf, collect_metas symbols))) - in - { gc_univ = "constr"; - gc_entries = - [ { ge_name = constr_rule n ":E"; - ge_type = ETast; - gl_assoc = None; - gl_rules = - [ { gr_name = fname; - gr_production = symbols; - gr_action = action} ] - } - ] - } - -let add_distfix a n df astf = +let add_distfix assoc n df astf = let fname = df^"_distfix" in - Lib.add_anonymous_leaf (inInfix(distfix a n (split df) fname astf, [])) + let (lp,rp) = prec_assoc assoc in + let symbols = find_symbols (n,lp) (8,E) (n,rp) 0 (split df) in + let gram_rule = make_grammar_rule n fname symbols astf in + let syntax_rule = make_syntax_rule n fname symbols astf in + Lib.add_anonymous_leaf (inInfix(gram_rule,syntax_rule)) diff --git a/toplevel/metasyntax.mli b/toplevel/metasyntax.mli index 34fbbef84..1eeec61b0 100644 --- a/toplevel/metasyntax.mli +++ b/toplevel/metasyntax.mli @@ -8,15 +8,21 @@ (*i $Id$ i*) +(*i*) +open Extend +open Vernacexpr +(*i*) + (* Adding grammar and pretty-printing objects in the environment *) -val add_syntax_obj : string -> Coqast.t list -> unit +val add_syntax_obj : string -> syntax_entry_ast list -> unit -val add_grammar_obj : string -> Coqast.t list -> unit +val add_grammar_obj : string -> grammar_entry_ast list -> unit val add_token_obj : string -> unit +val add_tactic_grammar : (string * (string * grammar_production list) * Tacexpr.raw_tactic_expr) list -> unit val add_infix : - Gramext.g_assoc option -> int -> string -> Coqast.t -> unit + Gramext.g_assoc option -> int -> string -> Nametab.qualid Util.located -> unit val add_distfix : Gramext.g_assoc option -> int -> string -> Coqast.t -> unit diff --git a/toplevel/protectedtoplevel.ml b/toplevel/protectedtoplevel.ml index 0e28a8373..7a6c1f2ae 100644 --- a/toplevel/protectedtoplevel.ml +++ b/toplevel/protectedtoplevel.ml @@ -159,11 +159,11 @@ let protected_loop input_chan = try while true do parse_one_command_group input_chan done with - | Vernacinterp.Drop -> raise Vernacinterp.Drop - | Vernacinterp.Quit -> exit 0 + | Vernacexpr.Drop -> raise Vernacexpr.Drop + | Vernacexpr.Quit -> exit 0 | End_of_file -> exit 0 - | DuringCommandInterp(loc, Vernacinterp.Quit) -> raise Vernacinterp.Quit - | DuringCommandInterp(loc, Vernacinterp.Drop) -> raise Vernacinterp.Drop + | DuringCommandInterp(loc, Vernacexpr.Quit) -> raise Vernacexpr.Quit + | DuringCommandInterp(loc, Vernacexpr.Drop) -> raise Vernacexpr.Drop | DuringCommandInterp(loc, e) -> explain_and_restart e | e -> explain_and_restart e in diff --git a/toplevel/record.ml b/toplevel/record.ml index df6d4ec67..c7907b167 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -17,6 +17,7 @@ open Termops open Environ open Declarations open Declare +open Nametab open Coqast open Astterm open Command @@ -29,16 +30,23 @@ open Type_errors (********** definition d'un record (structure) **************) let occur_fields id fs = - List.exists (fun (_,_,a) -> Ast.occur_var_ast id a) fs + List.exists + (function + | Vernacexpr.AssumExpr (_,a) -> Ast.occur_var_ast id a + | Vernacexpr.DefExpr (_,a,_) -> Ast.occur_var_ast id a) + fs let name_of id = if id = wildcard then Anonymous else Name id -let interp_decl env (id,assum,t) = - if assum then - (name_of id,None,interp_type Evd.empty env t) - else - let j = judgment_of_rawconstr Evd.empty env t in - (Name id,Some j.uj_val, j.uj_type) +let interp_decl sigma env = function + | Vernacexpr.AssumExpr(id,t) -> (name_of id,None,interp_type Evd.empty env t) + | Vernacexpr.DefExpr(id,c,t) -> + let c = match t with + | None -> c + | Some t -> Ast.ope("CAST",[c; t]) + in + let j = judgment_of_rawconstr Evd.empty env c in + (Name id,Some j.uj_val, j.uj_type) let build_decl_entry sigma env (id,t) = (id,Typeops.LocalAssum (interp_type Evd.empty env t)) @@ -48,22 +56,14 @@ let typecheck_params_and_fields ps fs = let env1,newps = List.fold_left (fun (env,newps) (id,t) -> - let decl = interp_decl env (id,true,t) in + let decl = interp_decl Evd.empty env (Vernacexpr.AssumExpr (id,t)) in (push_rel decl env,decl::newps)) (env0,[]) ps in -(* - let env2,newfs = - List.fold_left - (fun (env,newfs) d -> - let decl = interp_decl env d in - (push_named decl env, decl::newfs)) - (env1,[]) fs -*) let env2,newfs = List.fold_left (fun (env,newfs) d -> - let decl = interp_decl env d in + let decl = interp_decl Evd.empty env d in (push_rel decl env, decl::newfs)) (env1,[]) fs in @@ -213,7 +213,8 @@ let definition_structure (is_coe,idstruc,ps,cfs,idbuild,s) = let coers,fs = List.split cfs in let nparams = List.length ps in let idps = List.map fst ps in - let allnames = idstruc::idps@(List.map (fun (id,_,_) -> id) fs) in + let extract_name = function Vernacexpr.AssumExpr(id,_) -> id | Vernacexpr.DefExpr (id,_,_) -> id in + let allnames = idstruc::idps@(List.map extract_name fs) in if not (list_distinct allnames) then error "Two objects have the same name"; if occur_fields idstruc fs then error "A record cannot be recursive"; (* Now, younger decl in params and fields is on top *) diff --git a/toplevel/record.mli b/toplevel/record.mli index f511dadd4..1419b9c3c 100644 --- a/toplevel/record.mli +++ b/toplevel/record.mli @@ -22,5 +22,5 @@ val declare_projections : inductive -> bool list -> rel_context -> constant option list val definition_structure : - bool * identifier * (identifier * Coqast.t) list * - (bool * (identifier * bool * Coqast.t)) list * identifier * sorts -> unit + bool * identifier * (identifier * Genarg.constr_ast) list * + (bool * Vernacexpr.local_decl_expr) list * identifier * sorts -> unit diff --git a/toplevel/recordobj.ml b/toplevel/recordobj.ml index 50f2ef83b..92c8e5524 100755 --- a/toplevel/recordobj.ml +++ b/toplevel/recordobj.ml @@ -72,3 +72,5 @@ let objdef_declare ref = List.iter (fun (refi,c,argj) -> add_new_objdef ((refi,c),v,lt,params,argj)) comp + +let add_object_hook _ = objdef_declare diff --git a/toplevel/recordobj.mli b/toplevel/recordobj.mli index 10354968f..590482653 100755 --- a/toplevel/recordobj.mli +++ b/toplevel/recordobj.mli @@ -8,4 +8,7 @@ (* $Id$ *) -val objdef_declare : Nametab.global_reference -> unit +open Nametab + +val objdef_declare : global_reference -> unit +val add_object_hook : Proof_type.declaration_hook diff --git a/toplevel/toplevel.ml b/toplevel/toplevel.ml index 29c0e6055..589bc9ad6 100644 --- a/toplevel/toplevel.ml +++ b/toplevel/toplevel.ml @@ -226,13 +226,13 @@ let print_toplevel_error exc = match exc with | End_of_input -> msgerrnl (mt ()); pp_flush(); exit 0 - | Vernacinterp.Drop -> (* Last chance *) - if Mltop.is_ocaml_top() then raise Vernacinterp.Drop; + | Vernacexpr.Drop -> (* Last chance *) + if Mltop.is_ocaml_top() then raise Vernacexpr.Drop; (str"Error: There is no ML toplevel." ++ fnl ()) - | Vernacinterp.ProtectedLoop -> - raise Vernacinterp.ProtectedLoop - | Vernacinterp.Quit -> - raise Vernacinterp.Quit + | Vernacexpr.ProtectedLoop -> + raise Vernacexpr.ProtectedLoop + | Vernacexpr.Quit -> + raise Vernacexpr.Quit | _ -> (if is_pervasive_exn exc then (mt ()) else locstrm) ++ Cerrors.explain_exn exc @@ -305,12 +305,12 @@ let rec coq_switch b = end else protected_loop stdin with - | Vernacinterp.Drop -> () - | Vernacinterp.ProtectedLoop -> + | Vernacexpr.Drop -> () + | Vernacexpr.ProtectedLoop -> Lib.declare_initial_state(); coq_switch false | End_of_input -> msgerrnl (mt ()); pp_flush(); exit 0 - | Vernacinterp.Quit -> exit 0 + | Vernacexpr.Quit -> exit 0 | e -> msgerrnl (str"Anomaly: toplevel loop. Please report."); coq_switch b diff --git a/toplevel/toplevel.mli b/toplevel/toplevel.mli index 77f2ce89d..fbde39707 100644 --- a/toplevel/toplevel.mli +++ b/toplevel/toplevel.mli @@ -23,7 +23,7 @@ type input_buffer = { mutable str : string; (* buffer of already read characters *) mutable len : int; (* number of chars in the buffer *) mutable bols : int list; (* offsets in str of begining of lines *) - mutable tokens : Gram.parsable; (* stream of tokens *) + mutable tokens : Pcoq.Gram.parsable; (* stream of tokens *) mutable start : int } (* stream count of the first char of the buffer *) (* The input buffer of stdin. *) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index f42ae213b..e28821268 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -15,6 +15,7 @@ open Util open Options open System open Coqast +open Vernacexpr open Vernacinterp (* The functions in this module may raise (unexplainable!) exceptions. @@ -88,7 +89,7 @@ exception End_of_input let parse_phrase (po, verbch) = match Pcoq.Gram.Entry.parse Pcoq.main_entry po with - | Some com -> verbose_phrase verbch (Ast.loc com); com + | Some (loc,_ as com) -> verbose_phrase verbch loc; com | None -> raise End_of_input (* vernac parses the given stream, executes interpfun on the syntax tree it @@ -97,37 +98,35 @@ let parse_phrase (po, verbch) = let just_parsing = ref false let rec vernac interpfun input = - let com = parse_phrase input in - let rec interp com = - match com with - | Node (_,"LoadFile",[Str(_, verbosely); Str(_,fname)]) -> - let verbosely = (verbosely = "Verbose") in - let _ = read_vernac_file verbosely (make_suffix fname ".v") in - () - - | Node(_,"VernacList",l) -> - List.iter interp l - - | Node(_,"Time",[v]) -> - let tstart = System.get_time() in - interp v; - let tend = System.get_time() in - msgnl (str"Finished transaction in " ++ - System.fmt_time_difference tstart tend) - - | _ -> if not !just_parsing then interpfun com + let (loc,com) = parse_phrase input in + let rec interp = function + | VernacLoad (verbosely, fname) -> + let _ = read_vernac_file verbosely (make_suffix fname ".v") in + () + + | VernacList l -> List.iter (fun (_,v) -> interp v) l + + | VernacTime v -> + let tstart = System.get_time() in + if not !just_parsing then interpfun v; + let tend = System.get_time() in + msgnl (str"Finished transaction in " ++ + System.fmt_time_difference tstart tend) + + | v -> if not !just_parsing then interpfun v + in try interp com with e -> - raise (DuringCommandInterp (Ast.loc com, e)) + raise (DuringCommandInterp (loc, e)) and read_vernac_file verbosely s = let interpfun = if verbosely then - Vernacinterp.interp + Vernacentries.interp else - Options.silently Vernacinterp.interp + Options.silently Vernacentries.interp in let (in_chan, fname, input) = open_file_twice_if verbosely s in try @@ -146,7 +145,7 @@ and read_vernac_file verbosely s = * easier. *) let raw_do_vernac po = - vernac (States.with_heavy_rollback Vernacinterp.interp) (po,None); + vernac (States.with_heavy_rollback Vernacentries.interp) (po,None); Lib.mark_end_of_command() (* Load a vernac file. Errors are annotated with file and location *) diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli index da1aa182d..ad89461f2 100644 --- a/toplevel/vernac.mli +++ b/toplevel/vernac.mli @@ -18,7 +18,8 @@ exception Error_in_file of string * (string * Coqast.loc) * exn (* Read a vernac command on the specified input (parse only). Raises [End_of_file] if EOF (or Ctrl-D) is reached. *) -val parse_phrase : Pcoq.Gram.parsable * in_channel option -> Coqast.t +val parse_phrase : Pcoq.Gram.parsable * in_channel option -> + Coqast.loc * Vernacexpr.vernac_expr (* Reads and executes vernac commands from a stream. The boolean [just_parsing] disables interpretation of commands. *) diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index a72302458..8d8de77b9 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -10,78 +10,59 @@ (* Concrete syntax of the mathematical vernacular MV V2.6 *) -open Declarations open Pp open Util open Options -open System open Names open Nameops open Term open Pfedit open Tacmach open Proof_trees -open Proof_type -open Evar_refiner -open Tacred -open Environ -open Vernacinterp -open Coqast -open Ast open Astterm open Prettyp open Printer open Tacinterp -open Tactic_debug open Command open Goptions -open Declare open Nametab -open Safe_typing +open Vernacexpr -(* Dans join_binders, s'il y a un "?", on perd l'info qu'il est partag *) -let join_binders binders = - List.flatten - (List.map (fun (idl,c) -> List.map (fun id -> (id,c)) idl) binders) +(* Pcoq hooks *) -let add = vinterp_add +type pcoq_hook = { + start_proof : unit -> unit; + solve : int -> unit; + abort : string -> unit; + search : searchable -> dir_path list * bool -> unit; + print_name : qualid located -> unit; + print_check : Environ.unsafe_judgment -> unit; + print_eval : (constr -> constr) -> Environ.env -> Coqast.t -> Environ.unsafe_judgment -> unit; + show_goal : int option -> unit +} -(* equivalent to pr_subgoals but start from the prooftree and - check for uninstantiated existential variables *) +let pcoq = ref None +let set_pcoq_hook f = pcoq := Some f + +(*******************) +(* "Show" commands *) + + (* equivalent to pr_subgoals but start from the prooftree and + check for uninstantiated existential variables *) let pr_subgoals_of_pfts pfts = - let gls = fst (frontier (proof_of_pftreestate pfts)) in + let gls = fst (Refiner.frontier (proof_of_pftreestate pfts)) in let sigma = project (top_goal_of_pftreestate pfts) in pr_subgoals_existential sigma gls -let show_script () = - let pts = get_pftreestate () in - let pf = proof_of_pftreestate pts - and evc = evc_of_pftreestate pts in - msgnl(Refiner.print_script true evc (Global.named_context()) pf) - -let show_prooftree () = - let pts = get_pftreestate () in - let pf = proof_of_pftreestate pts - and evc = evc_of_pftreestate pts in - msg(Refiner.print_proof evc (Global.named_context()) pf) - let show_open_subgoals () = - let pfts = get_pftreestate () in - msg(pr_subgoals_of_pfts pfts) - -let show_nth_open_subgoal n = - let pf = proof_of_pftreestate (get_pftreestate ()) in - msg(pr_subgoal n (fst(frontier pf))) - -let show_open_subgoals_focused () = let pfts = get_pftreestate () in match focus() with | 0 -> msg(pr_subgoals_of_pfts pfts) | n -> let pf = proof_of_pftreestate pfts in - let gls = fst(frontier pf) in + let gls = fst(Refiner.frontier pf) in if n > List.length gls then (make_focus 0; msg(pr_subgoals_of_pfts pfts)) else if List.length gls < 2 then @@ -90,26 +71,119 @@ let show_open_subgoals_focused () = msg (v 0 (int(List.length gls) ++ str" subgoals" ++ cut () ++ pr_subgoal n gls)) +let show_nth_open_subgoal n = + let pf = proof_of_pftreestate (get_pftreestate ()) in + msg(pr_subgoal n (fst(Refiner.frontier pf))) + +let show_proof () = + let pts = get_pftreestate () in + let pf = proof_of_pftreestate pts in + let cursor = cursor_of_pftreestate pts in + let evc = evc_of_pftreestate pts in + let (pfterm,meta_types) = extract_open_pftreestate pts in + msgnl (str"LOC: " ++ + prlist_with_sep pr_spc pr_int (List.rev cursor) ++ fnl () ++ + str"Subgoals" ++ fnl () ++ + prlist (fun (mv,ty) -> int mv ++ str" -> " ++ prtype ty ++ fnl ()) + meta_types + ++ str"Proof: " ++ prterm (Evarutil.nf_evar evc pfterm)) + let show_node () = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts and cursor = cursor_of_pftreestate pts in msg (prlist_with_sep pr_spc pr_int (List.rev cursor) ++ fnl () ++ - prgl pf.goal ++ fnl () ++ - (match pf.ref with + prgl (goal_of_proof pf) ++ fnl () ++ + (match pf.Proof_type.ref with | None -> (str"BY ") | Some(r,spfl) -> - (str"BY " ++ Refiner.pr_rule r ++ fnl () ++ str" " ++ + (str"BY " ++ Refiner.pr_rule r ++ fnl () ++ + str" " ++ hov 0 (prlist_with_sep pr_fnl prgl - (List.map (fun p -> p.goal) spfl)))) ++ + (List.map goal_of_proof spfl)))) ++ fnl ()) +let show_script () = + let pts = get_pftreestate () in + let pf = proof_of_pftreestate pts + and evc = evc_of_pftreestate pts in + msgnl(Refiner.print_script true evc (Global.named_context()) pf) + let show_top_evars () = let pfts = get_pftreestate () in let gls = top_goal_of_pftreestate pfts in let sigma = project gls in msg (pr_evars_int 1 (Evd.non_instantiated sigma)) +let show_prooftree () = + let pts = get_pftreestate () in + let pf = proof_of_pftreestate pts + and evc = evc_of_pftreestate pts in + msg(Refiner.print_proof evc (Global.named_context()) pf) + +let print_subgoals () = if_verbose show_open_subgoals () + + (* Simulate the Intro(s) tactic *) + +let fresh_id_of_name avoid gl = function + Anonymous -> Tactics.fresh_id avoid (id_of_string "H") gl + | Name id -> id + +let rec do_renum avoid gl = function + [] -> mt () + | [n] -> pr_id (fresh_id_of_name avoid gl n) + | n :: l -> + let id = fresh_id_of_name avoid gl n in + pr_id id ++ spc () ++ do_renum (id :: avoid) gl l + +let show_intro all = + let pf = get_pftreestate() in + let gl = nth_goal_of_pftreestate 1 pf in + let l,_= decompose_prod (strip_outer_cast (pf_concl gl)) in + let nl = List.rev_map fst l in + if all then + msgnl (do_renum [] gl nl) + else + try + let n = List.hd nl in + msgnl (pr_id (fresh_id_of_name [] gl n)) + with Failure "hd" -> message "" + +(********************) +(* "Print" commands *) + +let print_path_entry (s,l) = + (str s ++ tbrk (0,2) ++ str (string_of_dirpath l)) + +let print_loadpath () = + let l = Library.get_full_load_path () in + msgnl (Pp.t (str "Physical path: " ++ + tab () ++ str "Logical Path:" ++ fnl () ++ + prlist_with_sep pr_fnl print_path_entry l)) + +let print_modules () = + let opened = Library.opened_modules () + and loaded = Library.loaded_modules () in + let loaded_opened = list_intersect loaded opened + and only_loaded = list_subtract loaded opened in + str"Loaded and imported modules: " ++ + pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++ + str"Loaded and not imported modules: " ++ + pr_vertical_list pr_dirpath only_loaded + +let dump_universes s = + let output = open_out s in + try + Univ.dump_universes output (Global.universes ()); + close_out output; + msg (str ("Universes written to file \""^s^"\".") ++ fnl ()) + with + e -> close_out output; raise e + + +(*********************) +(* "Locate" commands *) + let locate_file f = try let _,file = @@ -119,7 +193,7 @@ let locate_file f = msg (hov 0 (str"Can't find file" ++ spc () ++ str f ++ spc () ++ str"on loadpath" ++ fnl ())) -let print_located_qualid qid = +let print_located_qualid (_,qid) = try let sp = Nametab.sp_of_global (Global.env()) (Nametab.locate qid) in msg (pr_sp sp ++ fnl ()) @@ -127,33 +201,7 @@ let print_located_qualid qid = try msg (pr_sp (Syntax_def.locate_syntactic_definition qid) ++ fnl ()) with Not_found -> - error ((Nametab.string_of_qualid qid) ^ " is not a defined object") - -let print_path_entry (s,l) = - (str s ++ tbrk (0,2) ++ str (string_of_dirpath l)) - -let print_loadpath () = - let l = Library.get_full_load_path () in - msgnl (Pp.t (str "Physical path: " ++ - tab () ++ str "Logical Path:" ++ fnl () ++ - prlist_with_sep pr_fnl print_path_entry l)) - -let get_current_context_of_args = function - | [VARG_NUMBER n] -> get_goal_context n - | [] -> get_current_context () - | _ -> bad_vernac_args "goal_of_args" - -let isfinite = function - | "Inductive" -> true - | "CoInductive" -> false - | _ -> anomaly "Unexpected string" - -(* Locate commands *) -let _ = - add "LocateFile" - (function - | [VARG_STRING f] -> (fun () -> locate_file f) - | _ -> bad_vernac_args "LocateFile") + msg (pr_qualid qid ++ str " is not a defined object") let msg_found_library = function | Library.LibLoaded, fulldir, file -> @@ -161,12 +209,11 @@ let msg_found_library = function str file ++ fnl ()) | Library.LibInPath, fulldir, file -> msg (pr_dirpath fulldir ++ str " is bound to file " ++ str file ++ fnl ()) - -let msg_notfound_library qid = function +let msg_notfound_library loc qid = function | Library.LibUnmappedDir -> let dir = fst (Nametab.repr_qualid qid) in - errorlabstrm "locate_library" - (str "Cannot find a physical path bound to logical path " ++ + user_err_loc (loc,"locate_library", + str "Cannot find a physical path bound to logical path " ++ pr_dirpath dir ++ fnl ()) | Library.LibNotFound -> msg (hov 0 @@ -174,144 +221,566 @@ let msg_notfound_library qid = function spc () ++ Nametab.pr_qualid qid ++ fnl ())) | e -> assert false -let _ = - add "LocateLibrary" - (function - | [VARG_QUALID qid] -> - (fun () -> - try msg_found_library (Library.locate_qualified_library qid) - with e -> msg_notfound_library qid e) - | _ -> bad_vernac_args "LocateLibrary") - -let _ = - add "Locate" - (function - | [VARG_QUALID qid] -> (fun () -> print_located_qualid qid) - | _ -> bad_vernac_args "Locate") - -(* For compatibility *) -let _ = - add "ADDPATH" - (function - | [VARG_STRING dir] -> - (fun () -> Mltop.add_path dir Nameops.default_root_prefix) - | [VARG_STRING dir ; VARG_QUALID alias] -> - let aliasdir,aliasname = Nametab.repr_qualid alias in - (fun () -> Mltop.add_path dir (extend_dirpath aliasdir aliasname)) - | _ -> bad_vernac_args "ADDPATH") - -(* For compatibility *) -let _ = - add "DELPATH" - (function - | [VARG_STRING dir] -> (fun () -> Library.remove_path dir) - | _ -> bad_vernac_args "DELPATH") - -let _ = - add "RECADDPATH" - (function - | [VARG_STRING dir] -> - (fun () -> Mltop.add_rec_path dir Nameops.default_root_prefix) - | [VARG_STRING dir ; VARG_QUALID alias] -> - let aliasdir,aliasname = Nametab.repr_qualid alias in - (fun () ->Mltop.add_rec_path dir (extend_dirpath aliasdir aliasname)) - | _ -> bad_vernac_args "RECADDPATH") - -(* For compatibility *) -let _ = - add "PrintPath" - (function - | [] -> (fun () -> print_loadpath ()) - | _ -> bad_vernac_args "PrintPath") +let print_located_library (loc,qid) = + try msg_found_library (Library.locate_qualified_library qid) + with e -> msg_notfound_library loc qid e -let _ = - add "PWD" - (function - | [] -> (fun () -> print_endline (Sys.getcwd())) - | _ -> bad_vernac_args "PWD") -let _ = - add "DROP" - (function - | [] -> (fun () -> raise Drop) - | _ -> bad_vernac_args "DROP") +(**********) +(* Syntax *) -let _ = - add "PROTECTEDLOOP" - (function - | [] -> (fun () -> raise ProtectedLoop) - | _ -> bad_vernac_args "PROTECTEDLOOP") +let vernac_syntax = Metasyntax.add_syntax_obj -let _ = - add "QUIT" - (function - | [] -> (fun () -> raise Quit) - | _ -> anomaly "To fill the empty holes...") +let vernac_grammar = Metasyntax.add_grammar_obj -let _ = - add "CD" - (function - | [VARG_STRING s] -> - (fun () -> - begin - try Sys.chdir (glob s) - with Sys_error str -> warning ("Cd failed: " ^ str) - end; - if_verbose print_endline (Sys.getcwd())) - | [] -> (fun () -> if_verbose print_endline (Sys.getcwd())) - | _ -> bad_vernac_args "CD") - -(* Managing states *) +let vernac_infix assoc n inf qid = + let sp = sp_of_global (Global.env()) (global qid) in + let dir = repr_dirpath (Nameops.dirpath sp) in +(* + if dir <> [] then begin + let modname = List.hd dir in + let long_inf = (string_of_id modname) ^ "." ^ inf in + Metasyntax.add_infix assoc n long_inf qid + end; +*) + Metasyntax.add_infix assoc n inf qid + +let vernac_distfix assoc n inf qid = + Metasyntax.add_distfix assoc n inf (Astterm.globalize_qualid qid) + + +(***********) +(* Gallina *) + +let interp_assumption = function + | (AssumptionHypothesis | AssumptionVariable) -> Declare.make_strength_0 () + | (AssumptionAxiom | AssumptionParameter) -> Nametab.NeverDischarge + +let interp_definition = function + | Definition -> (false, Nametab.NeverDischarge) + | LocalDefinition -> (true, Declare.make_strength_0 ()) + +let interp_theorem = function + | (Theorem | Lemma | Decl) -> Nametab.NeverDischarge + | Fact -> Declare.make_strength_1 () + | Remark -> Declare.make_strength_0 () + +let interp_goal = function + | StartTheoremProof x -> (false, interp_theorem x) + | StartDefinitionBody x -> interp_definition x + +let vernac_definition kind id red_option c typ_opt hook = + let red_option = match red_option with + | None -> None + | Some r -> + let (evc,env)= Command.get_current_context () in + Some (interp_redexp env evc r) in + let (local,stre as x) = interp_definition kind in + let ref = declare_definition red_option id x c typ_opt in + hook stre ref + +let vernac_start_proof kind sopt t lettop hook = + if not(refining ()) then + if lettop then + errorlabstrm "Vernacentries.StartProof" + (str "Let declarations can only be used in proof editing mode") +(* else if s = None then + error "repeated Goal not permitted in refining mode"*); + start_proof_com sopt (interp_goal kind) t hook; + print_subgoals (); + if !pcoq <> None then (out_some !pcoq).start_proof () + +let vernac_end_proof is_opaque idopt = + if_verbose show_script (); + match idopt with + | None -> save_named is_opaque + | Some (id,None) -> save_anonymous is_opaque id + | Some (id,Some kind) -> + save_anonymous_with_strength (interp_theorem kind) is_opaque id + + (* A stupid macro that should be replaced by ``Exact c. Save.'' all along + the theories [??] *) + +let vernac_exact_proof c = + by (Tactics.exact_proof c); + save_named true + +let vernac_assumption kind l = + let stre = interp_assumption kind in + List.iter (fun (id,c) -> declare_assumption id stre c) l + +let vernac_inductive f indl = build_mutual indl f + +let vernac_fixpoint = build_recursive + +let vernac_cofixpoint = build_corecursive + +let vernac_scheme = build_scheme + + +(**********************) +(* Gallina extensions *) + +let vernac_record iscoe struc binders sort nameopt cfs = + let const = match nameopt with + | None -> add_prefix "Build_" struc + | Some id -> id in + let s = Astterm.interp_sort sort in + Record.definition_structure (iscoe,struc,binders,cfs,const,s) + + (* Sections *) + +let vernac_begin_section id = let _ = Lib.open_section id in () + +let vernac_end_section id = + check_no_pending_proofs (); + Discharge.close_section (is_verbose ()) id + +let is_obsolete_module (_,qid) = + match repr_qualid qid with + | dir, id when dir = empty_dirpath -> + (match string_of_id id with + | ("Refine" | "Inv" | "Equality" | "EAuto" | "AutoRewrite" + | "EqDecide" | "Xml" | "Extraction" | "Tauto" | "Setoid_replace" + | "Elimdep" ) + -> true + | _ -> false) + | _ -> false + +let vernac_require import _ qidl = + try + match import with + | None -> List.iter Library.read_module qidl + | Some b -> Library.require_module None qidl b + with e -> + (* Compatibility message *) + let qidl' = List.filter is_obsolete_module qidl in + if qidl' = qidl then + List.iter + (fun (_,qid) -> + warning ("Module "^(string_of_qualid qid)^ + " is now built-in and shouldn't be required")) qidl + else + raise e + +let vernac_import export qidl = + List.iter (Library.import_module export) qidl + +let vernac_canonical locqid = + Recordobj.objdef_declare (Nametab.global locqid) + +let cl_of_qualid = function + | FunClass -> Classops.CL_FUN + | SortClass -> Classops.CL_SORT + | RefClass (loc,qid) -> Class.class_of_ref (Nametab.global (loc, qid)) + +let vernac_coercion stre (loc,qid as locqid) qids qidt = + let target = cl_of_qualid qidt in + let source = cl_of_qualid qids in + let ref = Nametab.global locqid in + Class.try_add_new_coercion_with_target ref stre source target; + if_verbose message ((Nametab.string_of_qualid qid) ^ " is now a coercion") + +let vernac_identity_coercion stre id qids qidt = + let target = cl_of_qualid qidt in + let source = cl_of_qualid qids in + Class.try_add_new_identity_coercion id stre source target + + +(***********) +(* Solving *) + +let vernac_solve n tcom = + if not (refining ()) then + error "Unknown command of the non proof-editing mode"; + solve_nth n (Tacinterp.hide_interp tcom); + print_subgoals(); + (* in case a strict subtree was completed, + go back to the top of the prooftree *) + if subtree_solved () then + (reset_top_of_tree (); print_subgoals()); + if !pcoq <> None then (out_some !pcoq).solve n + + (* A command which should be a tactic. It has been + added by Christine to patch an error in the design of the proof + machine, and enables to instantiate existential variables when + there are no more goals to solve. It cannot be a tactic since + all tactics fail if there are no further goals to prove. *) + +let vernac_solve_existential = instantiate_nth_evar_com + + +(*****************************) +(* Auxiliary file management *) + +let vernac_require_from export spec id filename = + Library.require_module_from_file spec (Some id) filename export + +let vernac_add_loadpath isrec pdir ldiropt = + let alias = match ldiropt with + | None -> Nameops.default_root_prefix + | Some ldir -> ldir in + (if isrec then Mltop.add_rec_path else Mltop.add_path) pdir alias + +let vernac_remove_loadpath = Library.remove_path + + (* Coq syntax for ML or system commands *) + +let vernac_add_ml_path isrec s = + (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir) (System.glob s) + +let vernac_declare_ml_module l = Mltop.declare_ml_modules l + +let vernac_chdir = function + | None -> print_endline (Sys.getcwd()) + | Some s -> + begin + try Sys.chdir (System.glob s) + with Sys_error str -> warning ("Cd failed: " ^ str) + end; + if_verbose print_endline (Sys.getcwd()) + + +(********************) +(* State management *) let abort_refine f x = if Pfedit.refining() then delete_all_proofs (); f x (* used to be: error "Must save or abort current goal first" *) -let _ = - add "WriteState" - (function - | [VARG_STRING file] -> - (fun () -> abort_refine States.extern_state file) - | [VARG_IDENTIFIER id] -> - let file = string_of_id id in - (fun () -> abort_refine States.extern_state file) - | _ -> bad_vernac_args "SaveState") +let vernac_write_state file = abort_refine States.extern_state file -let _ = - add "RestoreState" - (function - | [VARG_STRING file] -> - (fun () -> abort_refine States.intern_state file) - | [VARG_IDENTIFIER id] -> - let file = string_of_id id in - (fun () -> abort_refine States.intern_state file) - | _ -> bad_vernac_args "LoadState") +let vernac_restore_state file = abort_refine States.intern_state file + +(*************) (* Resetting *) -let _ = - add "ResetName" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> abort_refine Lib.reset_name id) - | _ -> bad_vernac_args "ResetName") +let vernac_reset_name id = abort_refine Lib.reset_name id -let _ = - add "ResetInitial" - (function - | [] -> (fun () -> abort_refine Lib.reset_initial ()) - | _ -> bad_vernac_args "ResetInitial") +let vernac_reset_initial () = abort_refine Lib.reset_initial () -let _ = - add "Back" - (function - | [] -> (fun () -> Lib.back 1) - | [VARG_NUMBER n] -> - if n < 1 then error "argument of Back must be positive" - else (fun () -> Lib.back n) - | _ -> bad_vernac_args "Back") +let vernac_back n = Lib.back n + + +(************) +(* Commands *) +let vernac_declare_tactic_definition _ l = Tacinterp.add_tacdef l + +let vernac_hints = Auto.add_hints + +let vernac_syntactic_definition id c = function + | None -> syntax_definition id c + | Some n -> + let l = list_tabulate (fun _ -> Ast.ope("ISEVAR",[])) n in + let c = Ast.ope ("APPLIST",c :: l) in + syntax_definition id c + +let vernac_declare_implicits locqid = function + | Some imps -> Impargs.declare_manual_implicits (Nametab.global locqid) imps + | None -> Impargs.declare_implicits (Nametab.global locqid) + +let make_silent_if_not_pcoq b = + if !pcoq <> None then + error "Turning on/off silent flag is not supported in Centaur mode" + else make_silent b + +let _ = + declare_bool_option + { optsync = false; + optname = "silent"; + optkey = (PrimaryTable "Silent"); + optread = is_silent; + optwrite = make_silent_if_not_pcoq } + +let _ = + declare_bool_option + { optsync = true; + optname = "implicit arguments"; + optkey = (SecondaryTable ("Implicit","Arguments")); + optread = Impargs.is_implicit_args; + optwrite = Impargs.make_implicit_args } + +let _ = + declare_bool_option + { optsync = true; + optname = "coercion printing"; + optkey = (SecondaryTable ("Printing","Coercions")); + optread = (fun () -> !Termast.print_coercions); + optwrite = (fun b -> Termast.print_coercions := b) } + +let _ = + declare_int_option + { optsync=false; + optkey=PrimaryTable("Undo"); + optname="the undo limit"; + optread=Pfedit.get_undo; + optwrite=Pfedit.set_undo } + +let _ = + declare_int_option + { optsync=false; + optkey=SecondaryTable("Hyps","Limit"); + optname="the hypotheses limit"; + optread=Options.print_hyps_limit; + optwrite=Options.set_print_hyps_limit } + +let _ = + declare_int_option + { optsync=true; + optkey=SecondaryTable("Printing","Depth"); + optname="the printing depth"; + optread=Pp_control.get_depth_boxes; + optwrite=Pp_control.set_depth_boxes } + +let vernac_set_opacity opaq locqid = + match Nametab.global locqid with + | ConstRef sp -> + if opaq then Tacred.set_opaque_const sp + else Tacred.set_transparent_const sp + | VarRef id -> + if opaq then Tacred.set_opaque_var id + else Tacred.set_transparent_var id + | _ -> error "cannot set an inductive type or a constructor as transparent" + +let vernac_set_option key = function + | StringValue s -> set_string_option_value key s + | IntValue n -> set_int_option_value key (Some n) + | BoolValue b -> set_bool_option_value key b + +let vernac_unset_option key = + try set_bool_option_value key false + with _ -> + set_int_option_value key None + +let vernac_add_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#add s + | QualidRefValue locqid -> (get_ref_table key)#add locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_remove_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#remove s + | QualidRefValue locqid -> (get_ref_table key)#remove locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_mem_option key lv = + let f = function + | StringRefValue s -> (get_string_table key)#mem s + | QualidRefValue locqid -> (get_ref_table key)#mem locqid + in + try List.iter f lv with Not_found -> error_undeclared_key key + +let vernac_print_option key = + try (get_ref_table key)#print + with not_found -> + try (get_string_table key)#print + with not_found -> + try print_option_value key + with Not_found -> error_undeclared_key key + +let get_current_context_of_args = function + | Some n -> get_goal_context n + | None -> get_current_context () + +let vernac_check_may_eval redexp glopt rc = + let (evmap,env) = get_current_context_of_args glopt in + let c = interp_constr evmap env rc in + let j = Typeops.typing env c in + match redexp with + | None -> + if !pcoq <> None then (out_some !pcoq).print_check j + else msg (print_judgment env j) + | Some r -> + let redfun = Tacred.reduction_of_redexp (interp_redexp env evmap r) in + if !pcoq <> None + then (out_some !pcoq).print_eval (redfun env evmap) env rc j + else msg (print_eval redfun env j) + + (* The same but avoiding the current goal context if any *) +let vernac_global_check c = + let evmap = Evd.empty in + let env = Global.env() in + let c = interp_constr evmap env c in + let senv = Global.safe_env() in + let j = Safe_typing.typing senv c in + msg (print_safe_judgment env j) + +let vernac_print = function + | PrintTables -> print_tables () + | PrintLocalContext -> msg (print_local_context ()) + | PrintFullContext -> msg (print_full_context_typ ()) + | PrintSectionContext qid -> msg (print_sec_context_typ qid) + | PrintInspect n -> msg (inspect n) + | PrintGrammar (uni,ent) -> Metasyntax.print_grammar uni ent + | PrintLoadPath -> (* For compatibility ? *) print_loadpath () + | PrintModules -> msg (print_modules ()) + | PrintMLLoadPath -> Mltop.print_ml_path () + | PrintMLModules -> Mltop.print_ml_modules () + | PrintName qid -> + if !pcoq <> None then (out_some !pcoq).print_name qid + else msg (print_name qid) + | PrintOpaqueName qid -> msg (print_opaque_name qid) + | PrintGraph -> ppnl (Prettyp.print_graph()) + | PrintClasses -> ppnl (Prettyp.print_classes()) + | PrintCoercions -> ppnl (Prettyp.print_coercions()) + | PrintCoercionPaths (cls,clt) -> + ppnl (Prettyp.print_path_between (cl_of_qualid cls) (cl_of_qualid clt)) + | PrintUniverses None -> pp (Univ.pr_universes (Global.universes ())) + | PrintUniverses (Some s) -> dump_universes s + | PrintHint qid -> Auto.print_hint_ref (Nametab.global qid) + | PrintHintGoal -> Auto.print_applicable_hint () + | PrintHintDbName s -> Auto.print_hint_db_by_name s + | PrintHintDb -> Auto.print_searchtable () + + +let global_loaded_library (loc, qid) = + try Nametab.locate_loaded_library qid + with Not_found -> + user_err_loc (loc, "global_loaded_library", + str "Module/section " ++ pr_qualid qid ++ str " not found") + +let interp_search_restriction = function + | SearchOutside l -> (List.map global_loaded_library l, true) + | SearchInside l -> (List.map global_loaded_library l, false) + +let vernac_search s r = + let r = interp_search_restriction r in + if !pcoq <> None then (out_some !pcoq).search s r else + match s with + | SearchPattern c -> + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + Search.search_pattern pat r + | SearchRewrite c -> + let _,pat = interp_constrpattern Evd.empty (Global.env()) c in + Search.search_rewrite pat r + | SearchHead locqid -> + Search.search_by_head (Nametab.global locqid) r + +let vernac_locate = function + | LocateTerm qid -> print_located_qualid qid + | LocateLibrary qid -> print_located_library qid + | LocateFile f -> locate_file f + + +(********************) +(* Proof management *) + +let vernac_goal = function + | None -> () + | Some c -> + if not (refining()) then begin + start_proof_com None (false,Nametab.NeverDischarge) c (fun _ _ ->()); + print_subgoals () + end else + error "repeated Goal not permitted in refining mode" + +let vernac_abort = function + | None -> + delete_current_proof (); + if_verbose message "Current goal aborted"; + if !pcoq <> None then (out_some !pcoq).abort "" + | Some id -> + delete_proof id; + let s = string_of_id id in + if_verbose message ("Goal "^s^" aborted"); + if !pcoq <> None then (out_some !pcoq).abort s + +let vernac_abort_all () = + if refining() then begin + delete_all_proofs (); + message "Current goals aborted" + end else + error "No proof-editing in progress" + +let vernac_restart () = restart_proof(); show_open_subgoals () + + (* Proof switching *) + +let vernac_suspend = suspend_proof + +let vernac_resume = function + | None -> resume_last_proof () + | Some id -> resume_proof id + +let vernac_undo n = + undo n; + show_open_subgoals () + + (* Est-ce normal que "Focus" ne semble pas se comporter comme "Focus 1" ? *) +let vernac_focus = function + | None -> make_focus 1; show_open_subgoals () + | Some n -> traverse_nth_goal n; show_open_subgoals () + + (* Reset the focus to the top of the tree *) +let vernac_unfocus () = + make_focus 0; reset_top_of_tree (); show_open_subgoals () + +let vernac_go = function + | GoTo n -> Pfedit.traverse n;show_node() + | GoTop -> Pfedit.reset_top_of_tree ();show_node() + | GoNext -> Pfedit.traverse_next_unproven ();show_node() + | GoPrev -> Pfedit.traverse_prev_unproven ();show_node() + +let apply_subproof f occ = + let pts = get_pftreestate() in + let evc = evc_of_pftreestate pts in + let rec aux pts = function + | [] -> pts + | (n::l) -> aux (Tacmach.traverse n pts) occ in + let pts = aux pts (occ@[-1]) in + let pf = proof_of_pftreestate pts in + f evc (Global.named_context()) pf + +let explain_proof occ = msg (apply_subproof (Refiner.print_script true) occ) + +let explain_tree occ = msg (apply_subproof Refiner.print_proof occ) + +let vernac_show = function + | ShowGoal nopt -> + if !pcoq <> None then (out_some !pcoq).show_goal nopt + else (match nopt with + | None -> show_open_subgoals () + | (Some n) -> show_nth_open_subgoal n) + | ShowGoalImplicitly None -> Impargs.implicitly show_open_subgoals () + | ShowGoalImplicitly (Some n) -> Impargs.implicitly show_nth_open_subgoal n + | ShowProof -> show_proof () + | ShowNode -> show_node () + | ShowScript -> show_script () + | ShowExistentials -> show_top_evars () + | ShowTree -> show_prooftree () + | ShowProofNames -> + msgnl (prlist_with_sep pr_spc pr_id (Pfedit.get_all_proof_names())) + | ShowIntros all -> show_intro all + | ExplainProof occ -> explain_proof occ + | ExplainTree occ -> explain_tree occ + +let vernac_check_guard () = + let pts = get_pftreestate () in + let pf = proof_of_pftreestate pts + and cursor = cursor_of_pftreestate pts in + let (pfterm,_) = extract_open_pftreestate pts in + let message = + try + Inductiveops.control_only_guard (Evarutil.evar_env (goal_of_proof pf)) + pfterm; + (str "The condition holds up to here") + with UserError(_,s) -> + (str ("Condition violated : ") ++s) + in + msgnl message + +let vernac_debug b = + set_debug (if b then Tactic_debug.DebugOn else Tactic_debug.DebugOff) + + +(**************************) +(* Not supported commands *) (*** let _ = add "ResetSection" @@ -320,12 +789,10 @@ let _ = (fun () -> reset_section (string_of_id id)) | _ -> bad_vernac_args "ResetSection") -***) - (* Modules *) let _ = - add "BeginModule" + vinterp_add "BeginModule" (function | [VARG_IDENTIFIER id] -> let s = string_of_id id in @@ -336,7 +803,7 @@ let _ = | _ -> bad_vernac_args "BeginModule") let _ = - add "WriteModule" + vinterp_add "WriteModule" (function | [VARG_IDENTIFIER id] -> let mid = Lib.end_module id in @@ -347,1103 +814,29 @@ let _ = | _ -> bad_vernac_args "WriteModule") let _ = - add "ReadModule" - (function l -> - let l = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "ReadModule") l in - fun () -> List.iter Library.read_module l) - -let _ = - add "ImportModule" - (function l -> - let l = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "ImportModule") l in - fun () -> List.iter (Library.import_module false) l) - -let _ = - add "ExportModule" - (function l -> - let l = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "ExportModule") l in - fun () -> List.iter (Library.import_module true) l) - -let _ = - add "PrintModules" - (function - | [] -> - (fun () -> - let opened = Library.opened_modules () - and loaded = Library.loaded_modules () in - let loaded_opened = list_intersect loaded opened - and only_loaded = list_subtract loaded opened in - msg (str"Loaded and imported modules: " ++ - pr_vertical_list pr_dirpath loaded_opened ++ fnl () ++ - str"Loaded and not imported modules: " ++ - pr_vertical_list pr_dirpath only_loaded)) - | _ -> bad_vernac_args "PrintModules") - -(* Sections *) - -let _ = - add "BeginSection" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> let _ = Lib.open_section id in ()) - | _ -> bad_vernac_args "BeginSection") - -let _ = - add "EndSection" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> - check_no_pending_proofs (); - Discharge.close_section (is_verbose ()) id) - | _ -> bad_vernac_args "EndSection") - -(* Proof switching *) - -let _ = - add "GOAL" - (function - | [VARG_CONSTR com] -> - (fun () -> - if not (refining()) then begin - start_proof_com None Declare.NeverDischarge com; - if_verbose show_open_subgoals () - end else - error "repeated Goal not permitted in refining mode") - | [] -> (fun () -> ()) - | _ -> bad_vernac_args "GOAL") - -let _ = add "Comments" (function _ -> (fun () -> ())) - -let _ = - add "ABORT" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> - delete_proof id; - message ("Goal "^(string_of_id id)^" aborted")) - | [] -> (fun () -> - delete_current_proof (); - message "Current goal aborted") - | _ -> bad_vernac_args "ABORT") - -let _ = - add "ABORTALL" - (function - | [] -> (fun () -> - if refining() then begin - delete_all_proofs (); - message "Current goals aborted" - end else - error "No proof-editing in progress") - - | _ -> bad_vernac_args "ABORTALL") - -let _ = - add "RESTART" - (function - | [] -> (fun () -> (restart_proof();show_open_subgoals ())) - | _ -> bad_vernac_args "RESTART") - -let _ = - add "SUSPEND" - (function - | [] -> (fun () -> suspend_proof ()) - | _ -> bad_vernac_args "SUSPEND") - -let _ = - add "RESUME" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> resume_proof id) - | [] -> (fun () -> resume_last_proof ()) - | _ -> bad_vernac_args "RESUME") - -let _ = - add "FOCUS" - (function - | [] -> - (fun () -> make_focus 1; show_open_subgoals_focused ()) - | [VARG_NUMBER n] -> - (fun () -> traverse_nth_goal n; show_open_subgoals_focused ()) - | _ -> bad_vernac_args "FOCUS") - -(* Reset the focus to the top of the tree *) - -let _ = - add "UNFOCUS" - (function - | [] -> - (fun () -> - make_focus 0; reset_top_of_tree (); show_open_subgoals ()) - | _ -> bad_vernac_args "UNFOCUS") - -let _ = declare_bool_option - {optsync = true; - optname = "implicit arguments"; - optkey = (SecondaryTable ("Implicit","Arguments")); - optread = Impargs.is_implicit_args; - optwrite = Impargs.make_implicit_args } - -let _ = - add "IMPLICIT_ARGS_ON" - (function - | [] -> (fun () -> set_bool_option_value - (SecondaryTable ("Implicit","Arguments")) true) - | _ -> bad_vernac_args "IMPLICIT_ARGS_ON") - -let _ = - add "IMPLICIT_ARGS_OFF" - (function - | [] -> (fun () -> set_bool_option_value - (SecondaryTable ("Implicit","Arguments")) false) - | _ -> bad_vernac_args "IMPLICIT_ARGS_OFF") - -let _ = - add "TEST_IMPLICIT_ARGS" - (function - | [] -> - (fun () -> - if Impargs.is_implicit_args () then - message "Implicit arguments mode is set" - else - message "Implicit arguments mode is unset") - | _ -> bad_vernac_args "IMPLICIT_ARGS_OFF") - -let coercion_of_qualid loc qid = - let ref = Nametab.global loc qid in - let coe = Classops.coe_of_reference ref in - if not (Classops.coercion_exists coe) then - errorlabstrm "try_add_coercion" - (Printer.pr_global ref ++ str" is not a coercion"); - coe - -let _ = declare_bool_option - {optsync = true; - optname = "coercion printing"; - optkey = (SecondaryTable ("Printing","Coercions")); - optread = (fun () -> !Termast.print_coercions); - optwrite = (fun b -> Termast.print_coercions := b) } - -let number_list = - List.map (function VARG_NUMBER n -> n | _ -> bad_vernac_args "Number list") - -let _ = - add "IMPLICITS" - (function - | (VARG_STRING "") :: (VARG_QUALID qid) :: l -> - (fun () -> - let imps = number_list l in - Impargs.declare_manual_implicits - (Nametab.global dummy_loc qid) imps) - | [VARG_STRING "Auto"; VARG_QUALID qid] -> - (fun () -> - Impargs.declare_implicits (Nametab.global dummy_loc qid)) - | _ -> bad_vernac_args "IMPLICITS") - -let interp_definition_kind = function - | "THEOREM" -> (NeverDischarge, true) - | "LEMMA" -> (NeverDischarge, true) - | "FACT" -> (make_strength_1 (), true) - | "REMARK" -> (make_strength_0 (), true) - | "DEFINITION" -> (NeverDischarge, false) - | "LET" -> (make_strength_1 (), false) - | "LETTOP" -> (NeverDischarge, false) - | "LOCAL" -> (make_strength_0 (), false) - | _ -> anomaly "Unexpected definition kind" - -let _ = - add "SaveNamed" - (function - | [] -> - (fun () -> if_verbose show_script (); save_named true) - | _ -> bad_vernac_args "SaveNamed") - -let _ = - add "DefinedNamed" - (function - | [] -> - (fun () -> if_verbose show_script (); save_named false) - | _ -> bad_vernac_args "DefinedNamed") - -let _ = - add "DefinedAnonymous" - (function - | [VARG_IDENTIFIER id] -> - (fun () -> - if_verbose show_script (); - save_anonymous false id) - | _ -> bad_vernac_args "DefinedAnonymous") - -let _ = - add "SaveAnonymous" - (function - | [VARG_STRING kind; VARG_IDENTIFIER id] -> - (fun () -> - let (strength, opacity) = interp_definition_kind kind in - if_verbose show_script (); - save_anonymous_with_strength strength opacity id) - | [VARG_IDENTIFIER id] -> - (fun () -> - if_verbose show_script (); - save_anonymous true id) - | _ -> bad_vernac_args "SaveAnonymous") - -let _ = - add "TRANSPARENT" - (fun id_list () -> - List.iter - (function - | VARG_QUALID qid -> - (match Nametab.global dummy_loc qid with - | ConstRef sp -> Tacred.set_transparent_const sp - | VarRef id -> Tacred.set_transparent_var id - | _ -> error - "cannot set an inductive type or a constructor as transparent") - | _ -> bad_vernac_args "TRANSPARENT") - id_list) - -let _ = - add "OPAQUE" - (fun id_list () -> - List.iter - (function - | VARG_QUALID qid -> - (match Nametab.global dummy_loc qid with - | ConstRef sp -> Tacred.set_opaque_const sp - | VarRef id -> Tacred.set_opaque_var id - | _ -> error - "cannot set an inductive type or a constructor as opaque") - | _ -> bad_vernac_args "OPAQUE") - id_list) - -let _ = - add "UNDO" - (function - | [VARG_NUMBER n] -> - (fun () -> (undo n;show_open_subgoals_focused())) - | _ -> bad_vernac_args "UNDO") - -let _ = - add "SHOW" - (function - | [] -> (fun () -> show_open_subgoals ()) - | [VARG_NUMBER n] -> (fun () -> show_nth_open_subgoal n) - | _ -> bad_vernac_args "SHOW") - -let _ = - add "SHOWIMPL" - (function - | [] -> (fun () -> Impargs.implicitely show_open_subgoals ()) - | [VARG_NUMBER n] -> - (fun () -> Impargs.implicitely show_nth_open_subgoal n) - | _ -> bad_vernac_args "SHOWIMPL") - -let _ = - add "ShowNode" - (function - | [] -> (fun () -> show_node()) - | _ -> bad_vernac_args "ShowNode") - -let _ = - add "ShowEx" - (function - | [] -> (fun () -> show_top_evars ()) - | _ -> bad_vernac_args "ShowEx") - -let _ = - add "Go" - (function - | [VARG_NUMBER n] -> - (fun () -> (Pfedit.traverse n;show_node())) - | [VARG_STRING"top"] -> - (fun () -> (Pfedit.reset_top_of_tree ();show_node())) - | [VARG_STRING"next"] -> - (fun () -> (Pfedit.traverse_next_unproven ();show_node())) - | [VARG_STRING"prev"] -> - (fun () -> (Pfedit.traverse_prev_unproven ();show_node())) - | _ -> bad_vernac_args "Go") - -let _ = - add "ShowProof" - (function - | [] -> - (fun () -> - let pts = get_pftreestate () in - let pf = proof_of_pftreestate pts in - let cursor = cursor_of_pftreestate pts in - let evc = evc_of_pftreestate pts in - let (pfterm,meta_types) = extract_open_pftreestate pts in - msgnl (str"LOC: " ++ - prlist_with_sep pr_spc pr_int (List.rev cursor) ++ fnl () ++ - str"Subgoals" ++ fnl () ++ - prlist - (fun (mv,ty) -> - (int mv ++ str" -> " ++ prtype ty ++ fnl ())) - meta_types ++ - str"Proof: " ++ prterm (Evarutil.nf_evar evc pfterm))) - | _ -> bad_vernac_args "ShowProof") - -let _ = - add "CheckGuard" - (function - | [] -> - (fun () -> - let pts = get_pftreestate () in - let pf = proof_of_pftreestate pts - and cursor = cursor_of_pftreestate pts in - let (pfterm,_) = extract_open_pftreestate pts in - let message = - try - Inductiveops.control_only_guard (Evarutil.evar_env pf.goal) - pfterm; - (str "The condition holds up to here") - with UserError(_,s) -> - (str ("Condition violated : ") ++s) - in - msgnl message) - | _ -> bad_vernac_args "CheckGuard") - -let _ = - add "ShowTree" - (function - | [] -> (fun () -> show_prooftree()) - | _ -> bad_vernac_args "ShowTree") - -let _ = - add "ShowScript" - (function - | [] -> (fun () -> show_script()) - | _ -> bad_vernac_args "ShowScript") - -let _ = - add "ExplainProof" - (function l -> - let l = number_list l in - fun () -> - let pts = get_pftreestate() in - let evc = evc_of_pftreestate pts in - let rec aux pts = function - | [] -> pts - | (n::l) -> aux (Tacmach.traverse n pts) l in - let pts = aux pts (l@[-1]) in - let pf = proof_of_pftreestate pts in - msg (Refiner.print_script true evc (Global.named_context()) pf)) - -let _ = - add "ExplainProofTree" - (function l -> - let l = number_list l in - fun () -> - let pts = get_pftreestate() in - let evc = evc_of_pftreestate pts in - let rec aux pts = function - | [] -> pts - | (n::l) -> aux (Tacmach.traverse n pts) l in - let pts = aux pts (l@[-1]) in - let pf = proof_of_pftreestate pts in - msg (Refiner.print_proof evc (Global.named_context()) pf)) - -let _ = - add "ShowProofs" - (function [] -> - (fun () -> - let l = Pfedit.get_all_proof_names() in - msgnl (prlist_with_sep pr_spc pr_id l)) - | _ -> bad_vernac_args "ShowProofs") - -let _ = - add "PrintAll" - (function - | [] -> (fun () -> msg(print_full_context_typ ())) - | _ -> bad_vernac_args "PrintAll") - -let _ = - add "PRINT" - (function - | [] -> (fun () -> msg(print_local_context())) - | _ -> bad_vernac_args "PRINT") - -(* Pris en compte dans PrintOption *) - -let _ = - add "PrintId" - (function - | [VARG_QUALID qid] -> (fun () -> msg(print_name qid)) - | _ -> bad_vernac_args "PrintId") - -let _ = - add "PrintOpaqueId" - (function - | [VARG_QUALID qid] -> (fun () -> msg(print_opaque_name qid)) - | _ -> bad_vernac_args "PrintOpaqueId") - -let _ = - add "PrintSec" - (function - | [VARG_QUALID qid] -> (fun () -> msg(print_sec_context_typ qid)) - | _ -> bad_vernac_args "PrintSec") - -let _ = declare_bool_option - {optsync = false; - optname = "silent"; - optkey = (PrimaryTable "Silent"); - optread = is_silent; - optwrite = make_silent } - -let _ = - add "BeginSilent" - (function - | [] -> (fun () -> set_bool_option_value (PrimaryTable "Silent") true) - | _ -> bad_vernac_args "BeginSilent") - -let _ = - add "EndSilent" - (function - | [] -> (fun () -> set_bool_option_value (PrimaryTable "Silent") false) - | _ -> bad_vernac_args "EndSilent") - -let _ = - add "DebugOn" - (function - | [] -> (fun () -> set_debug DebugOn) - | _ -> bad_vernac_args "DebugOn") - -let _ = - add "DebugOff" - (function - | [] -> (fun () -> set_debug DebugOff) - | _ -> bad_vernac_args "DebugOff") - -let _ = - add "StartProof" - (function - | [VARG_STRING kind;VARG_IDENTIFIER s;VARG_CONSTR com] -> - let stre = fst (interp_definition_kind kind) in - fun () -> - begin - if (kind = "LETTOP") && not(refining ()) then - errorlabstrm "Vernacentries.StartProof" (str - "Let declarations can only be used in proof editing mode"); - start_proof_com (Some s) stre com; - if_verbose show_open_subgoals () - end - | _ -> bad_vernac_args "StartProof") - -let _ = - add "TheoremProof" - (function - | [VARG_STRING kind; VARG_IDENTIFIER id; - VARG_CONSTR com; VARG_VARGLIST coml] -> - let calls = List.map - (function - | (VCALL(na,l)) -> (na,l) - | _ -> bad_vernac_args "") - coml - in - let (stre,opacity) = interp_definition_kind kind in - (fun () -> - try - States.with_heavy_rollback - (fun () -> - start_proof_com (Some id) stre com; - if_verbose show_open_subgoals (); - List.iter Vernacinterp.call calls; - if_verbose show_script (); - if not (kind = "LETTOP") then - save_named opacity - else - let csr = interp_type Evd.empty (Global.env ()) com - and (_,({const_entry_body = pft},_)) = cook_proof () in - let cutt = vernac_tactic ("Cut",[Constr csr]) - and exat = vernac_tactic ("Exact",[Constr pft]) in - delete_proof id; - by (tclTHENS cutt [introduction id;exat])) - () - with e -> - if (is_unsafe "proof") && not (kind = "LETTOP") then begin - msgnl (str "Warning: checking of theorem " ++ pr_id id ++ - spc () ++ str "failed" ++ - str "... converting to Axiom"); - delete_proof id; - let _ = parameter_def_var id com in () - end else - errorlabstrm "Vernacentries.TheoremProof" - (str "checking of theorem " ++ pr_id id ++ spc () ++ - str "failed... aborting")) - | _ -> bad_vernac_args "TheoremProof") - -let _ = - add "DEFINITION" - (function - | (VARG_STRING kind :: VARG_IDENTIFIER id :: VARG_CONSTR c :: rest) -> - let typ_opt,red_option = match rest with - | [] -> None, None - | [VARG_CONSTR t] -> Some t, None - | [VARG_TACTIC_ARG (Redexp redexp)] -> - None, Some redexp - | [VARG_CONSTR t; VARG_TACTIC_ARG (Redexp redexp)] -> - Some t, Some redexp - | _ -> bad_vernac_args "DEFINITION" - in - let local,stre,coe,objdef,idcoe = match kind with - | "DEFINITION" -> false,NeverDischarge, false,false,false - | "COERCION" -> false,NeverDischarge, true, false,false - | "LCOERCION" -> true, make_strength_0(),true, false,false - | "LET" -> true, make_strength_2(),false,false,false - | "LOCAL" -> true, make_strength_0(),false,false,false - | "OBJECT" -> false,NeverDischarge, false,true, false - | "LOBJECT" -> true, make_strength_0(),false,true, false - | "OBJCOERCION" -> false,NeverDischarge, true, true, false - | "LOBJCOERCION" -> true,make_strength_0(),true,true, false - | "SUBCLASS" -> false,NeverDischarge, false,false,true - | "LSUBCLASS" -> true, make_strength_0(),false,false,true - | _ -> anomaly "Unexpected string" - in - fun () -> - let ref = - definition_body_red red_option id (local,stre) c typ_opt in - if coe then begin - Class.try_add_new_coercion ref stre; - if_verbose message ((string_of_id id) ^ " is now a coercion") - end; - if idcoe then begin - let cl = Class.class_of_ref ref in - Class.try_add_new_coercion_subclass cl stre - end; - if objdef then Recordobj.objdef_declare ref - | _ -> bad_vernac_args "DEFINITION") - -let _ = - add "CANONICAL" - (function - | [VARG_QUALID qid] -> - fun () -> - let ref = Nametab.global dummy_loc qid in - Recordobj.objdef_declare ref - | _ -> bad_vernac_args "CANONICAL") - -let _ = - add "VARIABLE" - (function - | [VARG_STRING kind; VARG_BINDERLIST slcl] -> - if List.length slcl = 1 & List.length (fst (List.hd slcl)) = 1 then - (match kind with - | "VARIABLES" -> warning "Many variables are expected" - | "HYPOTHESES" -> warning "Many hypotheses are expected" - | "COERCIONS" -> warning "Many hypotheses are expected" - | _ -> ()); - let coe = match kind with - | "COERCION" -> true - | "COERCIONS" -> true - | _ -> false - in - let stre = make_strength_0() in - fun () -> - List.iter - (fun (sl,c) -> - List.iter - (fun id -> - let ref = hypothesis_def_var (refining()) id stre c in - if coe then Class.try_add_new_coercion ref stre) - sl) - slcl - | _ -> bad_vernac_args "VARIABLE") - -let _ = - add "PARAMETER" - (function - | [VARG_STRING kind; VARG_BINDERLIST slcl] -> - if List.length slcl = 1 & List.length (fst (List.hd slcl)) = 1 & - kind = "PARAMETERS" then warning "Many parameters are expected"; - fun () -> - List.iter - (fun (sl,c) -> - List.iter - (fun s -> - let _ = parameter_def_var s c in ()) - sl) - slcl - | _ -> bad_vernac_args "PARAMETER") - -let _ = - add "Eval" - (function - | VARG_TACTIC_ARG (Redexp redexp) :: VARG_CONSTR c :: g -> - let (evmap,sign) = get_current_context_of_args g in - let redfun = print_eval (reduction_of_redexp redexp) sign in - fun () -> msg (redfun (judgment_of_rawconstr evmap sign c)) - | _ -> bad_vernac_args "Eval") - -let _ = - add "Check" - (function - | VARG_STRING "PRINTTYPE" :: VARG_CONSTR c :: _ -> - (fun () -> - let evmap = Evd.empty in - let env = Global.env() in - let c = interp_constr evmap env c in - let senv = Global.safe_env() in - let j = Safe_typing.typing senv c in - msg (print_safe_judgment env j)) - | VARG_STRING "CHECK" :: VARG_CONSTR c :: g -> - (fun () -> - let (evmap, env) = get_current_context_of_args g in - let c = interp_constr evmap env c in - let (j,cst) = Typeops.infer env c in - let _ = Environ.add_constraints cst env in - msg (print_judgment env j)) - | _ -> bad_vernac_args "Check") - - -let extract_qualid = function - | VARG_QUALID qid -> - (try Nametab.locate_loaded_library qid - with Not_found -> - error ("Module/section "^(Nametab.string_of_qualid qid)^" not found")) - | _ -> bad_vernac_args "extract_qualid" - -let inside_outside = function - | (VARG_STRING "outside") :: l -> List.map extract_qualid l, true - | (VARG_STRING "inside") :: l -> List.map extract_qualid l, false - | [] -> [], true - | _ -> bad_vernac_args "inside/outside" - -let _ = - add "SearchRewrite" - (function - | (VARG_CONSTR c) :: l -> - (fun () -> - let _,pat = interp_constrpattern Evd.empty (Global.env()) c in - Search.search_rewrite pat (inside_outside l)) - | _ -> bad_vernac_args "SearchRewrite") - -let _ = - add "SearchPattern" - (function - | (VARG_CONSTR c) :: l -> - (fun () -> - let _,pat = interp_constrpattern Evd.empty (Global.env()) c in - Search.search_pattern pat (inside_outside l)) - | _ -> bad_vernac_args "SearchPattern") - -let _ = - add "SEARCH" - (function - | (VARG_QUALID qid) :: l -> - (fun () -> - let ref = Nametab.global dummy_loc qid in - Search.search_by_head ref (inside_outside l)) - | _ -> bad_vernac_args "SEARCH") - -let _ = - add "INSPECT" - (function - | [VARG_NUMBER n] -> (fun () -> msg(inspect n)) - | _ -> bad_vernac_args "INSPECT") - -let _ = - add "SETUNDO" - (function - | [VARG_NUMBER n] -> (fun () -> set_undo n) - | _ -> bad_vernac_args "SETUNDO") - -let _ = - add "UNSETUNDO" - (function - | [] -> (fun () -> reset_undo ()) - | _ -> bad_vernac_args "UNSETUNDO") - -let _ = - add "SETHYPSLIMIT" - (function - | [VARG_NUMBER n] -> (fun () -> set_print_hyps_limit n) - | _ -> bad_vernac_args "SETHYPSLIMIT") - -let _ = - add "UNSETHYPSLIMIT" - (function - | [] -> (fun () -> unset_print_hyps_limit ()) - | _ -> bad_vernac_args "UNSETHYPSLIMIT") - -let _ = - add "ONEINDUCTIVE" - (function - | [ VARG_STRING f; - VARG_IDENTIFIER s; - VARG_CONSTR c; - VARG_BINDERLIST binders; - VARG_BINDERLIST constructors ] -> - let la = join_binders binders in - let li = List.map - (function - | ([id],c) -> (id,c) - | _ -> bad_vernac_args "ONEINDUCTIVE") - constructors - in - (fun () -> build_mutual la [(s,c,li)] (isfinite f)) - | _ -> bad_vernac_args "ONEINDUCTIVE") - -let eq_la (id,ast) (id',ast') = id = id' & alpha_eq(ast,ast') - -let extract_la_lc = function - | [] -> anomaly "Vernacentries: empty list of inductive types" - | (la,lc)::rest -> - let rec check = function - | [] -> [] - | (la',lc)::rest -> - if (List.length la = List.length la') && - (List.for_all2 eq_la la la') - then - lc::(check rest) - else - error ("Parameters should be syntactically the same "^ - "for each inductive type") - in - (la,lc::(check rest)) - -let _ = - add "MUTUALINDUCTIVE" - (function - | [VARG_STRING f; VARG_VARGLIST indl] -> - let lac = - (List.map - (function - | (VARG_VARGLIST - [VARG_IDENTIFIER arid; - VARG_CONSTR arc; - VARG_BINDERLIST binders; - VARG_BINDERLIST lconstr]) - -> - let la = join_binders binders in - (la, (arid, arc, - List.map - (function - | ([id],c) -> (id,c) - | _ -> bad_vernac_args "") - lconstr)) - | _ -> bad_vernac_args "MUTUALINDUCTIVE") indl) - in - let (la,lnamearconstructs) = extract_la_lc lac in - fun () -> build_mutual la lnamearconstructs (isfinite f) - | _ -> bad_vernac_args "MUTUALINDUCTIVE") - -let _ = - add "OLDMUTUALINDUCTIVE" - (function - | [VARG_BINDERLIST binders; VARG_STRING f; VARG_VARGLIST indl] -> - let la = join_binders binders in - let lnamearconstructs = - (List.map - (function - | (VARG_VARGLIST - [VARG_IDENTIFIER arid; - VARG_CONSTR arc; - VARG_BINDERLIST lconstr]) - -> (arid, - arc, - List.map - (function - | ([id],c) -> (id,c) - | _ -> bad_vernac_args "OLDMUTUALINDUCTIVE") - lconstr) - | _ -> bad_vernac_args "") indl) - in - fun () -> build_mutual la lnamearconstructs (isfinite f) - | _ -> bad_vernac_args "MUTUALINDUCTIVE") - -let _ = - add "RECORD" - (function - | [VARG_STRING coe; - VARG_IDENTIFIER struc; - VARG_BINDERLIST binders; - VARG_CONSTR sort; - VARG_VARGLIST namec; - VARG_VARGLIST cfs] -> - let ps = join_binders binders in - let cfs = - List.map - (function - | (VARG_VARGLIST - [VARG_STRING str; VARG_STRING b; - VARG_IDENTIFIER id; VARG_CONSTR c]) -> - let assum = match b with - | "ASSUM" -> true - | "DEF" -> false - | _ -> bad_vernac_args "RECORD" in - (str = "COERCION", (id, assum, c)) - | _ -> bad_vernac_args "RECORD") - cfs in - let const = match namec with - | [] -> add_prefix "Build_" struc - | [VARG_IDENTIFIER id] -> id - | _ -> bad_vernac_args "RECORD" in - let iscoe = (coe = "COERCION") in - let s = interp_sort sort in - fun () -> Record.definition_structure (iscoe,struc,ps,cfs,const,s) - | _ -> bad_vernac_args "RECORD") - -let _ = - add "MUTUALRECURSIVE" - (function - | [VARG_VARGLIST lmi] -> - (let lnameargsardef = - List.map - (function - | (VARG_VARGLIST - [VARG_IDENTIFIER fid; - VARG_BINDERLIST farg; - VARG_CONSTR arf; - VARG_CONSTR ardef]) - -> (fid,join_binders farg,arf,ardef) - | _ -> bad_vernac_args "MUTUALRECURSIVE") lmi - in - fun () -> build_recursive lnameargsardef) - | _ -> bad_vernac_args "MUTUALRECURSIVE") - -let _ = - add "MUTUALCORECURSIVE" - (function - | [VARG_VARGLIST lmi] -> - let lnameardef = - List.map - (function - | (VARG_VARGLIST - [VARG_IDENTIFIER fid; - VARG_CONSTR arf; - VARG_CONSTR ardef]) - -> (fid,arf,ardef) - | _ -> bad_vernac_args "MUTUALCORECURSIVE") lmi - in - fun () -> build_corecursive lnameardef - | _ -> bad_vernac_args "MUTUALCORECURSIVE") - -let _ = - add "INDUCTIONSCHEME" - (function - | [VARG_VARGLIST lmi] -> - let lnamedepindsort = - List.map - (function - | (VARG_VARGLIST - [VARG_IDENTIFIER fid; - VARG_STRING depstr; - VARG_QUALID indid; - VARG_CONSTR sort]) -> - let dep = match depstr with - | "DEP" -> true - | "NODEP" -> false - | _ -> anomaly "Unexpected string" - in - (fid,dep,indid,sort) - | _ -> bad_vernac_args "INDUCTIONSCHEME") lmi - in - fun () -> build_scheme lnamedepindsort - | _ -> bad_vernac_args "INDUCTIONSCHEME") - -let _ = - add "SyntaxMacro" - (function - | [VARG_IDENTIFIER id;VARG_CONSTR com] -> - (fun () -> syntax_definition id com) - | [VARG_IDENTIFIER id;VARG_CONSTR com; VARG_NUMBER n] -> - (fun () -> - let rec aux t = function - | 0 -> t - | n -> aux (ope("APPLIST",[t;ope("ISEVAR",[])])) (n-1) - in - syntax_definition id (aux com n)) - | _ -> bad_vernac_args "SyntaxMacro") - -let _ = - add "Require" - (function - | VARG_STRING import :: VARG_STRING specif :: l -> - let l = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "Require") l in - let spec = - if specif="UNSPECIFIED" then - None - else - Some (specif="SPECIFICATION") in - let export = (import="EXPORT") in - fun () -> Library.require_module spec l export - | _ -> bad_vernac_args "Require") - -let _ = - add "RequireFrom" - (function - | [VARG_STRING import; VARG_STRING specif; - VARG_QUALID qid; VARG_STRING filename] -> - (fun () -> - Library.require_module_from_file - (if specif="UNSPECIFIED" then - None - else - Some (specif="SPECIFICATION")) - qid filename - (import="EXPORT")) - | _ -> bad_vernac_args "RequireFrom") - -let _ = - add "NOP" - (function - | [] -> (fun () -> ()) - | _ -> bad_vernac_args "NOP") - -let _ = - add "CLASS" + vinterp_add "CLASS" (function | [VARG_STRING kind; VARG_QUALID qid] -> let stre = if kind = "LOCAL" then make_strength_0() else - NeverDischarge + Nametab.NeverDischarge in fun () -> - let ref = Nametab.global dummy_loc qid in + let ref = Nametab.global (dummy_loc, qid) in Class.try_add_new_class ref stre; if_verbose message ((Nametab.string_of_qualid qid) ^ " is now a class") | _ -> bad_vernac_args "CLASS") -let cl_of_qualid qid = - match Nametab.repr_qualid qid with - | d, id when string_of_id id = "FUNCLASS" & repr_dirpath d = [] -> - Classops.CL_FUN - | d, id when string_of_id id = "SORTCLASS" & repr_dirpath d = [] -> - Classops.CL_SORT - | _ -> Class.class_of_ref (Nametab.global dummy_loc qid) - -let _ = - add "COERCION" - (function - | [VARG_STRING kind;VARG_STRING identity; - VARG_QUALID qid;VARG_QUALID qids;VARG_QUALID qidt] -> - let stre = - if (kind = "LOCAL") then - make_strength_0() - else - NeverDischarge - in - let isid = identity = "IDENTITY" in - let target = cl_of_qualid qidt in - let source = cl_of_qualid qids in - fun () -> - if isid then match Nametab.repr_qualid qid with - | d, id when repr_dirpath d = [] -> - Class.try_add_new_identity_coercion id stre source target - | _ -> bad_vernac_args "COERCION" - else - let ref = Nametab.global dummy_loc qid in - Class.try_add_new_coercion_with_target ref stre source target; - if_verbose - message - ((Nametab.string_of_qualid qid) ^ " is now a coercion") - | _ -> bad_vernac_args "COERCION") - -let _ = - add "PrintGRAPH" - (function - | [] -> (fun () -> ppnl (Prettyp.print_graph())) - | _ -> bad_vernac_args "PrintGRAPH") - -let _ = - add "PrintCLASSES" - (function - | [] -> (fun () -> ppnl (Prettyp.print_classes())) - | _ -> bad_vernac_args "PrintCLASSES") - -let _ = - add "PrintCOERCIONS" - (function - | [] -> (fun () -> ppnl (Prettyp.print_coercions())) - | _ -> bad_vernac_args "PrintCOERCIONS") - -let _ = - add "PrintPATH" - (function - | [VARG_QUALID qids;VARG_QUALID qidt] -> - (fun () -> - ppnl (Prettyp.print_path_between - (cl_of_qualid qids) (cl_of_qualid qidt))) - | _ -> bad_vernac_args "PrintPATH") - (* Meta-syntax commands *) - let _ = add "TOKEN" (function | [VARG_STRING s] -> (fun () -> Metasyntax.add_token_obj s) | _ -> bad_vernac_args "TOKEN") - -let _ = - add "GRAMMAR" - (function - | [VARG_STRING univ; VARG_ASTLIST al] -> - (fun () -> Metasyntax.add_grammar_obj univ al) - | _ -> bad_vernac_args "GRAMMAR") - -let _ = - add "SYNTAX" - (function - | [VARG_STRING whatfor; VARG_ASTLIST sel] -> - (fun () -> Metasyntax.add_syntax_obj whatfor sel) - | _ -> bad_vernac_args "SYNTAX") - -let _ = - add "TACDEF" - (let rec tacdef_fun lacc=function - (VARG_IDENTIFIER name)::(VARG_AST tacexp)::tl -> - tacdef_fun ((name,tacexp)::lacc) tl - |[] -> - fun () -> - List.iter (fun (name,ve) -> Tacinterp.add_tacdef name ve) lacc - |_ -> bad_vernac_args "TACDEF" - in - tacdef_fun []) - -let _ = - add "INFIX" - (function - | [VARG_AST assoc; VARG_NUMBER n; VARG_STRING inf; VARG_QUALID pref] -> - let ref = Astterm.globalize_qualid pref in - (fun () -> - Metasyntax.add_infix (Extend.gram_assoc assoc) n inf ref) - | _ -> bad_vernac_args "INFIX") - -let _ = - add "DISTFIX" - (function - | [VARG_AST assoc; VARG_NUMBER n; VARG_STRING s; VARG_QUALID pref] -> - let ref = Astterm.globalize_qualid pref in - (fun () -> - Metasyntax.add_distfix (Extend.gram_assoc assoc) n s ref) - | _ -> bad_vernac_args "DISTFIX") - -let _ = - add "PrintGrammar" - (function - | [VARG_IDENTIFIER uni; VARG_IDENTIFIER ent] -> - (fun () -> - Metasyntax.print_grammar (string_of_id uni) (string_of_id ent)) - | _ -> bad_vernac_args "PrintGrammar") +***) (* Search commands *) @@ -1460,407 +853,97 @@ let _ = | _ -> bad_vernac_args "Searchisos") ***) -open Goptions - -let _ = declare_int_option - { optsync=true; - optkey=SecondaryTable("Printing","Depth"); - optname="the printing depth"; - optread=Pp_control.get_depth_boxes; - optwrite=Pp_control.set_depth_boxes } - -let _ = - add "SetTableField" - (function - (* Hook for Printing Coercions *) - | (VARG_IDENTIFIER t) :: (VARG_IDENTIFIER f) :: l - when t = id_of_string "Printing" && f = id_of_string "Coercion" -> - let ql = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "PRINTING_COERCIONS_ON") l in - (fun () -> - List.iter - (fun qid -> - Classops.set_coercion_visibility true - (coercion_of_qualid dummy_loc qid)) - ql) - | [(VARG_IDENTIFIER t);(VARG_IDENTIFIER f);arg] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - match arg with - | VARG_NUMBER n -> set_int_option_value key n - | VARG_STRING s -> set_string_option_value key s - | _ -> error "No such type of option value") - | [(VARG_IDENTIFIER t);(VARG_IDENTIFIER f)] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - set_bool_option_value key true) - | [(VARG_IDENTIFIER t);arg] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - match arg with - | VARG_NUMBER n -> set_int_option_value key n - | VARG_STRING s -> set_string_option_value key s - | _ -> error "No such type of option value") - | [(VARG_IDENTIFIER t)] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - set_bool_option_value key true) - | _ -> bad_vernac_args "VernacOption") - -let _ = - add "UnsetTableField" - (function - | [(VARG_IDENTIFIER t);(VARG_IDENTIFIER f)] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - set_bool_option_value key false) - | [(VARG_IDENTIFIER t)] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - set_bool_option_value key false) - | _ -> bad_vernac_args "VernacOption") - -let _ = - add "AddTableField" - (function - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_QUALID v] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_ident_table key)#add (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_STRING s] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_string_table key)#add s - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_QUALID v] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_ident_table key)#add (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_STRING s] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_string_table key)#add s - with Not_found -> - error_undeclared_key key) - | _ -> bad_vernac_args "TableField") - -let _ = - add "RemoveTableField" - (function - (* Hook for Printing Coercions *) - | (VARG_IDENTIFIER t) :: (VARG_IDENTIFIER f) :: l - when t = id_of_string "Printing" && f = id_of_string "Coercion" -> - let ql = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "PRINTING_COERCIONS_OFF") l in - (fun () -> - List.iter - (fun qid -> - Classops.set_coercion_visibility false - (coercion_of_qualid dummy_loc qid)) - ql) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_QUALID v] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_ident_table key)#remove (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_STRING v] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_string_table key)#remove v - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_QUALID v] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_ident_table key)#remove (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_STRING v] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_string_table key)#remove v - with Not_found -> - error_undeclared_key key) - | _ -> bad_vernac_args "TableField") - -let _ = - add "MemTableField" - (function - (* Hook for Printing Coercions *) - | (VARG_IDENTIFIER t) :: (VARG_IDENTIFIER f) :: l - when t = id_of_string "Printing" && f = id_of_string "Coercion" -> - let ql = - List.map - (function - | VARG_QUALID qid -> qid - | _ -> bad_vernac_args "TEST_PRINTING_COERCIONS") l in - (fun () -> - List.iter - (fun qid -> - let coe = coercion_of_qualid dummy_loc qid in - if Classops.is_coercion_visible coe then - message - ("Printing of coercion "^(Nametab.string_of_qualid qid)^ - " is set") - else - message - ("Printing of coercion "^(Nametab.string_of_qualid qid)^ - " is unset")) - ql) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_QUALID v] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_ident_table key)#mem (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f; VARG_STRING v] -> - (fun () -> - let key = - SecondaryTable (string_of_id t,string_of_id f) in - try - (get_string_table key)#mem v - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_QUALID v] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_ident_table key)#mem (Nametab.global dummy_loc v) - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_IDENTIFIER v] -> - (fun () -> - let key = SecondaryTable (string_of_id t, string_of_id v) in - try - (get_ident_table key)#print - with not_found -> - try - (get_string_table key)#print - with not_found -> - try - print_option_value key - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t; VARG_STRING v] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_string_table key)#mem v - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_ident_table key)#print - with not_found -> - try - (get_string_table key)#print - with not_found -> - try - print_option_value key - with Not_found -> - error_undeclared_key key) - | _ -> bad_vernac_args "TableField") - -let _ = - add "PrintOption" - (function - | [VARG_IDENTIFIER t; VARG_IDENTIFIER f] -> - (fun () -> - let key = SecondaryTable (string_of_id t,string_of_id f) in - try - (get_ident_table key)#print - with not_found -> - try - (get_string_table key)#print - with not_found -> - try - print_option_value key - with Not_found -> - error_undeclared_key key) - | [VARG_IDENTIFIER t] -> - (fun () -> - let key = PrimaryTable (string_of_id t) in - try - (get_ident_table key)#print - with not_found -> - try - (get_string_table key)#print - with not_found -> - try - print_option_value key - with Not_found -> - if (string_of_id t) = "Tables" then - print_tables () - else - msg(print_name (Nametab.make_short_qualid t))) - | _ -> bad_vernac_args "TableField") - - -open Tactics -open Pfedit - -(* - The first one is a command that should be a tactic. It has been - added by Christine to patch an error in the design of the proof - machine, and enables to instantiate existential variables when - there are no more goals to solve. It cannot be a tactic since - all tactics fail if there are no further goals to prove. *) - -let _ = - vinterp_add "EXISTENTIAL" - (function - | [VARG_NUMBER n; VARG_CONSTR c] -> - (fun () -> instantiate_nth_evar_com n c) - | _ -> assert false) - -(* The second is a stupid macro that should be replaced by ``Exact - c. Save.'' all along the theories. *) - -let _ = - vinterp_add "PROOF" - (function - | [VARG_CONSTR c] -> - (fun () -> (* by (tactic_com exact c) *) - (* on experimente la synthese d'ise dans exact *) - by (dyn_exact_check [Command c]); - save_named true) - | _ -> assert false) - -let print_subgoals () = if_verbose show_open_subgoals_focused () - -let _ = - vinterp_add "SOLVE" - (function l -> - let (n,tcom) = match l with - | [VARG_NUMBER n;VARG_TACTIC tcom] -> (n,tcom) - | _ -> invalid_arg "SOLVE" - in - (fun () -> - if not (refining ()) then - error "Unknown command of the non proof-editing mode"; - solve_nth n (Tacinterp.hide_interp tcom); - print_subgoals(); - (* in case a strict subtree was completed, - go back to the top of the prooftree *) - if subtree_solved () then - (reset_top_of_tree (); print_subgoals()) - )) - -(* Coq syntax for ML commands *) - -let _ = vinterp_add "DeclareMLModule" - (fun l -> - let sl = - List.map - (function - | (VARG_STRING s) -> s - | _ -> anomaly "DeclareMLModule : not a string") l - in - fun () -> Mltop.declare_ml_modules sl) - -let _ = vinterp_add "AddMLPath" - (function - | [VARG_STRING s] -> - (fun () -> Mltop.add_ml_dir (glob s)) - | _ -> anomaly "AddMLPath : not a string") - -let _ = vinterp_add "RecAddMLPath" - (function - | [VARG_STRING s] -> - (fun () -> Mltop.add_rec_ml_dir (glob s)) - | _ -> anomaly "RecAddMLPath : not a string") - -let _ = vinterp_add "PrintMLPath" - (function - | [] -> (fun () -> Mltop.print_ml_path ()) - | _ -> anomaly "PrintMLPath : does not expect any argument") - -let _ = vinterp_add "PrintMLModules" - (function - | [] -> (fun () -> Mltop.print_ml_modules ()) - | _ -> anomaly "PrintMLModules : does not expect an argument") - -let _ = vinterp_add "DumpUniverses" - (function - | [] -> (fun () -> pp (Univ.pr_universes (Global.universes ()))) - | [VARG_STRING s] -> - (fun () -> let output = open_out s in - try - Univ.dump_universes output (Global.universes ()); - close_out output; - msg (str ("Universes written to file \""^s^"\".") ++ fnl ()) - with - e -> close_out output; raise e - ) - | _ -> - anomaly "DumpUniverses : expects a filename or nothing as argument") - -(* Simulate the Intro(s) tactic *) - -open Tactics - -let id_of_name = function - Anonymous -> id_of_string "H" - | Name id -> id;; - -let rec do_renum avoid gl = function - [] -> "" - | [n] -> (string_of_id (fresh_id avoid (id_of_name n) gl)) - | n :: l -> let id = fresh_id avoid (id_of_name n) gl in - (string_of_id id)^" "^(do_renum (id :: avoid) gl l) - -let _ = vinterp_add "SHOWINTRO" - (function - | [] -> - (fun () -> - let pf = get_pftreestate() in - let gl = nth_goal_of_pftreestate 1 pf in - let l,_= decompose_prod (strip_outer_cast (pf_concl gl)) in - try - let n = - List.hd (List.rev_map fst l) in - message (string_of_id (fresh_id [] (id_of_name n) gl)) - with Failure "hd" -> message "") - | _ -> bad_vernac_args "Show Intro") - -let _ = vinterp_add "SHOWINTROS" - (function - | [] -> - (fun () -> - let pf = get_pftreestate() in - let gl = nth_goal_of_pftreestate 1 pf in - let l,_= decompose_prod (strip_outer_cast (pf_concl gl)) in - let l = List.rev_map fst l in - message (do_renum [] gl l)) - | _ -> bad_vernac_args "Show Intros") +let interp c = match c with + (* Control (done in vernac) *) + | (VernacTime _ | VernacVar _ | VernacList _ | VernacLoad _) -> assert false + + (* Syntax *) + | VernacSyntax (whatfor,sel) -> vernac_syntax whatfor sel + | VernacTacticGrammar al -> Metasyntax.add_tactic_grammar al + | VernacGrammar (univ,al) -> vernac_grammar univ al + | VernacInfix (assoc,n,inf,qid) -> vernac_infix assoc n inf qid + | VernacDistfix (assoc,n,inf,qid) -> vernac_distfix assoc n inf qid + + (* Gallina *) + | VernacDefinition (k,id,red,t,otyp,f) -> vernac_definition k id red t otyp f + | VernacStartProof (k,id,t,top,f) -> vernac_start_proof k id t top f + | VernacEndProof (opaq,idopt) -> vernac_end_proof opaq idopt + | VernacExactProof c -> vernac_exact_proof c + | VernacAssumption (stre,l) -> vernac_assumption stre l + | VernacInductive (finite,l) -> vernac_inductive finite l + | VernacFixpoint l -> vernac_fixpoint l + | VernacCoFixpoint l -> vernac_cofixpoint l + | VernacScheme l -> vernac_scheme l + + (* Gallina extensions *) + | VernacRecord (coe,id,bl,s,idopt,fs) -> vernac_record coe id bl s idopt fs + | VernacBeginSection id -> vernac_begin_section id + | VernacEndSection id -> vernac_end_section id + | VernacRequire (export,spec,qidl) -> vernac_require export spec qidl + | VernacImport (export,qidl) -> vernac_import export qidl + | VernacCanonical qid -> vernac_canonical qid + | VernacCoercion (str,r,s,t) -> vernac_coercion str r s t + | VernacIdentityCoercion (str,id,s,t) -> vernac_identity_coercion str id s t + + (* Solving *) + | VernacSolve (n,tac) -> vernac_solve n tac + | VernacSolveExistential (n,c) -> vernac_solve_existential n c + + (* Auxiliary file and library management *) + | VernacRequireFrom (exp,spec,id,f) -> vernac_require_from exp spec id f + | VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias + | VernacRemoveLoadPath s -> vernac_remove_loadpath s + | VernacAddMLPath (isrec,s) -> vernac_add_ml_path isrec s + | VernacDeclareMLModule l -> vernac_declare_ml_module l + | VernacChdir s -> vernac_chdir s + + (* State management *) + | VernacWriteState s -> vernac_write_state s + | VernacRestoreState s -> vernac_restore_state s + + (* Resetting *) + | VernacResetName id -> vernac_reset_name id + | VernacResetInitial -> vernac_reset_initial () + | VernacBack n -> vernac_back n + + (* Commands *) + | VernacDeclareTacticDefinition (x,l) -> vernac_declare_tactic_definition x l + | VernacHints (dbnames,hints) -> vernac_hints dbnames hints + | VernacHintDestruct (id,l,p,n,tac) -> Dhyp.add_destructor_hint id l p n tac + | VernacSyntacticDefinition (id,c,n) -> vernac_syntactic_definition id c n + | VernacDeclareImplicits (qid,l) -> vernac_declare_implicits qid l + | VernacSetOpacity (opaq, qidl) -> List.iter (vernac_set_opacity opaq) qidl + | VernacSetOption (key,v) -> vernac_set_option key v + | VernacUnsetOption key -> vernac_unset_option key + | VernacRemoveOption (key,v) -> vernac_remove_option key v + | VernacAddOption (key,v) -> vernac_add_option key v + | VernacMemOption (key,v) -> vernac_mem_option key v + | VernacPrintOption key -> vernac_print_option key + | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval r g c + | VernacGlobalCheck c -> vernac_global_check c + | VernacPrint p -> vernac_print p + | VernacSearch (s,r) -> vernac_search s r + | VernacLocate l -> vernac_locate l + | VernacComments l -> if_verbose message ("Comments ok\n") + | VernacNop -> () + + (* Proof management *) +(* | VernacGoal c -> vernac_goal c*) + | VernacAbort id -> vernac_abort id + | VernacAbortAll -> vernac_abort_all () + | VernacRestart -> vernac_restart () + | VernacSuspend -> vernac_suspend () + | VernacResume id -> vernac_resume id + | VernacUndo n -> vernac_undo n + | VernacFocus n -> vernac_focus n + | VernacUnfocus -> vernac_unfocus () + | VernacGo g -> vernac_go g + | VernacShow s -> vernac_show s + | VernacCheckGuard -> vernac_check_guard () + | VernacDebug b -> vernac_debug b + + (* Toplevel control *) + | VernacToplevelControl e -> raise e + + (* Extensions *) + | VernacExtend (opn,args) -> Vernacinterp.call (opn,args) diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli index de9a2a092..85daf4375 100644 --- a/toplevel/vernacentries.mli +++ b/toplevel/vernacentries.mli @@ -12,29 +12,57 @@ open Names open Term open Vernacinterp +open Vernacexpr (*i*) (* Vernacular entries. This module registers almost all the vernacular entries, by side-effects using [Vernacinterp.vinterp_add]. *) +(* Moved to g_vernac.ml4 val join_binders : ('a list * 'b) list -> ('a * 'b) list -val add : string -> (vernac_arg list -> unit -> unit) -> unit +*) + +(* Synonym for Vernacinterp.vinterp_add +val add : string -> (Vernacexpr.vernac_arg list -> unit -> unit) -> unit +*) + val show_script : unit -> unit val show_prooftree : unit -> unit val show_open_subgoals : unit -> unit val show_nth_open_subgoal : int -> unit + +(* Merged with show_open_subgoals ! val show_open_subgoals_focused : unit -> unit +*) + val show_node : unit -> unit (* This function can be used by any command that want to observe terms in the context of the current goal, as for instance in pcoq *) -val get_current_context_of_args : vernac_arg list -> Evd.evar_map * Environ.env +val get_current_context_of_args : int option -> Evd.evar_map * Environ.env (* this function is used to analyse the extra arguments in search commands. - It is used in pcoq. *) -val inside_outside : vernac_arg list -> dir_path list * bool;; + It is used in pcoq. *) (*i anciennement: inside_outside i*) +(* +val interp_search_restriction : search_restriction -> dir_path list * bool +*) + +type pcoq_hook = { + start_proof : unit -> unit; + solve : int -> unit; + abort : string -> unit; + search : searchable -> dir_path list * bool -> unit; + print_name : Nametab.qualid Util.located -> unit; + print_check : Environ.unsafe_judgment -> unit; + print_eval : (constr -> constr) -> Environ.env -> Coqast.t -> Environ.unsafe_judgment -> unit; + show_goal : int option -> unit +} + +val set_pcoq_hook : pcoq_hook -> unit (* This function makes sure that the function given is argument is preceded by a command aborting all proofs if necessary. It is used in pcoq. *) val abort_refine : ('a -> unit) -> 'a -> unit;; + +val interp : Vernacexpr.vernac_expr -> unit diff --git a/toplevel/vernacinterp.ml b/toplevel/vernacinterp.ml index 039a42766..8e4189e2e 100644 --- a/toplevel/vernacinterp.ml +++ b/toplevel/vernacinterp.ml @@ -15,39 +15,18 @@ open Himsg open Proof_type open Tacinterp open Coqast +open Vernacexpr open Ast - -exception ProtectedLoop -exception Drop -exception Quit +open Extend let disable_drop e = if e <> Drop then e else UserError("Vernac.disable_drop",(str"Drop is forbidden.")) -type vernac_arg = - | VARG_VARGLIST of vernac_arg list - | VARG_STRING of string - | VARG_NUMBER of int - | VARG_NUMBERLIST of int list - | VARG_IDENTIFIER of identifier - | VARG_QUALID of Nametab.qualid - | VCALL of string * vernac_arg list - | VARG_CONSTR of Coqast.t - | VARG_CONSTRLIST of Coqast.t list - | VARG_TACTIC of Coqast.t - | VARG_TACTIC_ARG of tactic_arg - | VARG_BINDER of identifier list * Coqast.t - | VARG_BINDERLIST of (identifier list * Coqast.t) list - | VARG_AST of Coqast.t - | VARG_ASTLIST of Coqast.t list - | VARG_UNIT - | VARG_DYN of Dyn.t - - (* Table of vernac entries *) let vernac_tab = - (Hashtbl.create 51 : (string, vernac_arg list -> unit -> unit) Hashtbl.t) + (Hashtbl.create 51 : + (string, Tacexpr.raw_generic_argument list -> unit -> unit) Hashtbl.t) let vinterp_add s f = try @@ -73,46 +52,6 @@ let vinterp_map s = let vinterp_init () = Hashtbl.clear vernac_tab - -(* Conversion Coqast.t -> vernac_arg *) -let rec cvt_varg ast = - match ast with - | Node(_,"VERNACARGLIST",l) -> - VARG_VARGLIST (List.map cvt_varg l) - | Node(_,"VERNACCALL",(Str (_,na))::l) -> - VCALL (na,List.map cvt_varg l) - - | Nvar(_,id) -> VARG_IDENTIFIER id - | Node(loc,"QUALIDARG",p) -> VARG_QUALID (Astterm.interp_qualid p) - | Str(_,s) -> VARG_STRING s - | Id(_,s) -> VARG_STRING s - | Num(_,n) -> VARG_NUMBER n - | Node(_,"NONE",[]) -> VARG_UNIT - | Node(_,"CONSTR",[c]) -> VARG_CONSTR c - | Node(_,"CONSTRLIST",l) -> VARG_CONSTRLIST l - | Node(_,"TACTIC",[c]) -> VARG_TACTIC c - | Node(_,"BINDER",c::idl) -> - VARG_BINDER(List.map nvar_of_ast idl, c) - | Node(_,"BINDERLIST",l) -> - VARG_BINDERLIST - (List.map (compose (function - | (VARG_BINDER (x_0,x_1)) -> (x_0,x_1) - | _ -> assert false) cvt_varg) l) - | Node(_,"NUMBERLIST",ln) -> - VARG_NUMBERLIST (List.map num_of_ast ln) - - | Node(_,"AST",[a]) -> VARG_AST a - | Node(_,"ASTLIST",al) -> VARG_ASTLIST al - | Node(_,"TACTIC_ARG",[targ]) -> - let (evc,env)= Command.get_current_context () in - let ist = { evc=evc; env=env; lfun=[]; lmatch=[]; - goalopt=None; debug=get_debug ()} in - VARG_TACTIC_ARG (interp_tacarg ist targ) - | Node(_,"VERNACDYN",[Dynamic (_,d)]) -> VARG_DYN d - | _ -> anomaly_loc (Ast.loc ast, "Vernacinterp.cvt_varg", - (str "Unrecognizable ast node of vernac arg:" ++ - brk(1,0) ++ print_ast ast)) - (* Interpretation of a vernac command *) let call (opn,converted_args) = @@ -131,21 +70,6 @@ let call (opn,converted_args) = msgnl (str"Vernac Interpreter " ++ str !loc); raise e -let interp = function - | Node(_,opn,argl) as cmd -> - let converted_args = - try - List.map cvt_varg argl - with e -> - if !Options.debug then - msgnl (str"Vernac Interpreter " ++ str"Converting arguments"); - raise e - in - call (opn,converted_args) - | cmd -> - errorlabstrm "Vernac Interpreter" - (str"Malformed vernac AST : " ++ print_ast cmd) - let bad_vernac_args s = anomalylabstrm s (str"Vernac " ++ str s ++ str" called with bad arguments") diff --git a/toplevel/vernacinterp.mli b/toplevel/vernacinterp.mli index 2aa45f322..74f987ea4 100644 --- a/toplevel/vernacinterp.mli +++ b/toplevel/vernacinterp.mli @@ -9,47 +9,16 @@ (*i $Id$ i*) (*i*) -open Names -open Proof_type +open Tacexpr (*i*) -(* Interpretation of vernac phrases. *) +(* Interpretation of extended vernac phrases. *) -exception Drop -exception ProtectedLoop -exception Quit - val disable_drop : exn -> exn -type vernac_arg = - | VARG_VARGLIST of vernac_arg list - | VARG_STRING of string - | VARG_NUMBER of int - | VARG_NUMBERLIST of int list - | VARG_IDENTIFIER of identifier - | VARG_QUALID of Nametab.qualid - | VCALL of string * vernac_arg list - | VARG_CONSTR of Coqast.t - | VARG_CONSTRLIST of Coqast.t list - | VARG_TACTIC of Coqast.t - | VARG_TACTIC_ARG of tactic_arg - | VARG_BINDER of identifier list * Coqast.t - | VARG_BINDERLIST of (identifier list * Coqast.t) list - | VARG_AST of Coqast.t - | VARG_ASTLIST of Coqast.t list - | VARG_UNIT - | VARG_DYN of Dyn.t (* to put whatever you want *) - -val vinterp_add : string -> (vernac_arg list -> unit -> unit) -> unit +val vinterp_add : string -> (raw_generic_argument list -> unit -> unit) -> unit val overwriting_vinterp_add : - string -> (vernac_arg list -> unit -> unit) -> unit + string -> (raw_generic_argument list -> unit -> unit) -> unit -val vinterp_map : string -> vernac_arg list -> unit -> unit val vinterp_init : unit -> unit -val cvt_varg : Coqast.t -> vernac_arg -val call : string * vernac_arg list -> unit -val interp : Coqast.t -> unit - -(* raises an user error telling that the vernac command was called - with bad arguments *) -val bad_vernac_args : string -> 'a +val call : string * raw_generic_argument list -> unit -- cgit v1.2.3