diff options
author | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2005-12-26 20:07:21 +0000 |
---|---|---|
committer | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2005-12-26 20:07:21 +0000 |
commit | 52f4136ecf452162adb55c8de031b73c97dcdbac (patch) | |
tree | 8ac0a4c3584025a44067c6a96c6ce9d92ca93e78 /parsing | |
parent | 099fb1b4c5084bb899e4910e42c971cdfa81e1aa (diff) |
Renommage des Pp*new en Pp* (et déplacement dans parsing); renommage des G_*new en G_*.ml + autres petites modifications liées à suppression du traducteur
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@7740 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'parsing')
-rw-r--r-- | parsing/extend.mli | 10 | ||||
-rw-r--r-- | parsing/g_basevernac.ml4 | 536 | ||||
-rw-r--r-- | parsing/g_cases.ml4 | 79 | ||||
-rw-r--r-- | parsing/g_constr.ml4 | 545 | ||||
-rw-r--r-- | parsing/g_constrnew.ml4 | 330 | ||||
-rw-r--r-- | parsing/g_ltac.ml4 | 243 | ||||
-rw-r--r-- | parsing/g_ltacnew.ml4 | 197 | ||||
-rw-r--r-- | parsing/g_module.ml4 | 47 | ||||
-rw-r--r-- | parsing/g_prim.ml4 | 119 | ||||
-rw-r--r-- | parsing/g_primnew.ml4 | 87 | ||||
-rw-r--r-- | parsing/g_proofs.ml4 | 95 | ||||
-rw-r--r-- | parsing/g_proofsnew.ml4 | 126 | ||||
-rw-r--r-- | parsing/g_tactic.ml4 | 460 | ||||
-rw-r--r-- | parsing/g_tacticnew.ml4 | 402 | ||||
-rw-r--r-- | parsing/g_vernac.ml4 | 935 | ||||
-rw-r--r-- | parsing/g_vernacnew.ml4 | 733 | ||||
-rw-r--r-- | parsing/ppconstr.ml | 718 | ||||
-rw-r--r-- | parsing/ppconstr.mli | 79 | ||||
-rw-r--r-- | parsing/pptactic.ml | 729 | ||||
-rw-r--r-- | parsing/pptactic.mli | 13 | ||||
-rw-r--r-- | parsing/ppvernac.ml | 888 | ||||
-rw-r--r-- | parsing/ppvernac.mli | 31 | ||||
-rw-r--r-- | parsing/prettyp.ml | 32 | ||||
-rw-r--r-- | parsing/prettyp.mli | 1 | ||||
-rw-r--r-- | parsing/printer.ml | 15 | ||||
-rw-r--r-- | parsing/tactic_printer.ml | 4 |
26 files changed, 3718 insertions, 3736 deletions
diff --git a/parsing/extend.mli b/parsing/extend.mli index 6a51d738d..b092e766b 100644 --- a/parsing/extend.mli +++ b/parsing/extend.mli @@ -9,16 +9,6 @@ (*i $Id$ i*) open Util -(*i -open Pp -open Names -open Ast -open Coqast -open Ppextend -open Topconstr -open Genarg -open Mod_subst -i*) (**********************************************************************) (* constr entry keys *) diff --git a/parsing/g_basevernac.ml4 b/parsing/g_basevernac.ml4 deleted file mode 100644 index 49b29d1de..000000000 --- a/parsing/g_basevernac.ml4 +++ /dev/null @@ -1,536 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Coqast -open Extend -open Ppextend -open Vernacexpr -open Pcoq -open Vernac_ -open Goptions -open Constr -open Prim - -let vernac_kw = - [ "Quit"; "Load"; "Compile"; "Fixpoint"; "CoFixpoint"; - "Definition"; "Inductive"; "CoInductive"; - "Theorem"; "Variable"; "Axiom"; "Parameter"; "Hypothesis"; - "."; ">->" ] -let _ = - if !Options.v7 then - List.iter (fun s -> Lexer.add_token ("",s)) vernac_kw - -let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" -let lstring = Gram.Entry.create "lstring" - - -if !Options.v7 then -GEXTEND Gram - GLOBAL: class_rawexpr; - - class_rawexpr: - [ [ IDENT "FUNCLASS" -> FunClass - | IDENT "SORTCLASS" -> SortClass - | qid = global -> RefClass qid ] ] - ; -END; - -if !Options.v7 then -GEXTEND Gram - GLOBAL: command lstring; - - lstring: - [ [ s = STRING -> s ] ] - ; - comment: - [ [ c = constr -> CommentConstr c - | s = STRING -> CommentString s - | n = natural -> CommentInt n ] ] - ; - command: - [ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l - - (* System directory *) - | IDENT "Pwd" -> VernacChdir None - | IDENT "Cd" -> VernacChdir None - | IDENT "Cd"; dir = lstring -> 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"; fopt = OPT lstring -> - VernacPrint (PrintUniverses fopt) - - | IDENT "Locate"; l = locatable -> VernacLocate l - - (* Managing load paths *) - | IDENT "Add"; IDENT "LoadPath"; dir = lstring; alias = as_dirpath -> - VernacAddLoadPath (false, dir, alias) - | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = lstring; - alias = as_dirpath -> VernacAddLoadPath (true, dir, alias) - | IDENT "Remove"; IDENT "LoadPath"; dir = lstring -> - VernacRemoveLoadPath dir - - (* For compatibility *) - | IDENT "AddPath"; dir = lstring; alias = as_dirpath -> - VernacAddLoadPath (false, dir, alias) - | IDENT "AddRecPath"; dir = lstring; alias = as_dirpath -> - VernacAddLoadPath (true, dir, alias) - | IDENT "DelPath"; dir = lstring -> - VernacRemoveLoadPath dir - - (* Printing (careful factorization of entries) *) - | IDENT "Print"; p = printable -> VernacPrint p - | IDENT "Print"; qid = global -> VernacPrint (PrintName qid) - | IDENT "Print" -> VernacPrint PrintLocalContext - | IDENT "Print"; IDENT "Module"; "Type"; qid = global -> - VernacPrint (PrintModuleType qid) - | IDENT "Print"; IDENT "Module"; qid = global -> - VernacPrint (PrintModule qid) - | IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n) - | IDENT "About"; qid = global -> VernacPrint (PrintAbout qid) - - (* Searching the environment *) - | IDENT "Search"; qid = global; 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) - | IDENT "SearchAbout"; - sl = [ "["; l = LIST1 [ r = global -> SearchRef r - | s = lstring -> SearchString s ]; "]" -> l - | qid = global -> [SearchRef qid] ]; - l = in_or_out_modules -> - VernacSearch (SearchAbout sl, l) - - (* TODO: rapprocher Eval et Check *) - | IDENT "Eval"; r = Tactic.red_expr; "in"; - c = constr -> VernacCheckMayEval (Some r, None, c) - | IDENT "Check"; c = constr -> - VernacCheckMayEval (None, None, c) - | "Type"; c = constr -> VernacGlobalCheck c (* pas dans le RefMan *) - - | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = lstring -> - VernacAddMLPath (false, dir) - | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = lstring -> - VernacAddMLPath (true, dir) -(* - | IDENT "SearchIsos"; c = constr -> VernacSearch (SearchIsos c) -*) - - (* Pour intervenir sur les tables de paramtres *) - - | "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 global 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 "Term"; qid = global -> PrintOpaqueName qid - | IDENT "All" -> PrintFullContext - | IDENT "Section"; s = global -> PrintSectionContext s - | IDENT "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 "Ltac"; qid = global -> PrintLtac qid - | IDENT "Coercions" -> PrintCoercions - | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr - -> PrintCoercionPaths (s,t) - | IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions - | IDENT "Tables" -> PrintTables - | "Proof"; qid = global -> PrintOpaqueName qid - | IDENT "Hint" -> PrintHintGoal - | IDENT "Hint"; qid = global -> PrintHint qid - | IDENT "Hint"; "*" -> PrintHintDb - | IDENT "HintDb"; s = IDENT -> PrintHintDbName s - | IDENT "Rewrite"; IDENT "HintDb"; s = IDENT -> PrintRewriteHintDbName s - | IDENT "Scopes" -> PrintScopes - | IDENT "Scope"; s = IDENT -> PrintScope s - | IDENT "Visibility"; s = OPT IDENT -> PrintVisibility s - | IDENT "Implicit"; qid = global -> PrintImplicit qid ] ] - ; - locatable: - [ [ qid = global -> LocateTerm qid - | IDENT "File"; f = lstring -> LocateFile f - | IDENT "Library"; qid = global -> LocateLibrary qid - | IDENT "Module"; qid = global -> LocateModule qid - | s = lstring -> LocateNotation s ] ] - ; - option_value: - [ [ n = integer -> IntValue n - | s = lstring -> StringValue s ] ] - ; - option_ref_value: - [ [ id = global -> QualidRefValue id - | s = lstring -> StringRefValue s ] ] - ; - as_dirpath: - [ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ] - ; - in_or_out_modules: - [ [ IDENT "inside"; l = LIST1 global -> SearchInside l - | IDENT "outside"; l = LIST1 global -> SearchOutside l - | -> SearchOutside [] ] ] - ; -END - -(* Grammar extensions *) - -(* automatic translation of levels *) -let adapt_level n = - if n >= 10 then n*10 else - [| 0; 20; 30; 40; 50; 70; 80; 85; 90; 95; 100|].(n) - -let map_modl = function - | SetItemLevel(ids,NumLevel n) -> SetItemLevel(ids,NumLevel (adapt_level n)) - | SetLevel n -> SetLevel(adapt_level n) - | m -> m - -if !Options.v7 then -GEXTEND Gram - GLOBAL: syntax; - - univ: - [ [ univ = IDENT -> - set_default_action_parser (parser_type_from_name univ); univ ] ] - ; - grammar_tactic_level: - [ [ IDENT "simple_tactic" -> 0 - | IDENT "tactic1" -> 1 - | IDENT "tactic2" -> 2 - | IDENT "tactic3" -> 3 - | IDENT "tactic4" -> 4 - | IDENT "tactic5" -> 5 ] ] - ; - syntax: - [ [ IDENT "Token"; s = lstring -> - Pp.warning "Token declarations are now useless"; VernacNop - - | IDENT "Grammar"; IDENT "tactic"; lev = grammar_tactic_level; - OPT [ ":"; IDENT "tactic" ]; ":="; - OPT "|"; tl = LIST0 grammar_tactic_rule SEP "|" -> - VernacTacticGrammar (lev,tl) - - | IDENT "Grammar"; u = univ; - tl = LIST1 grammar_entry SEP "with" -> - VernacGrammar (rename_command_entry u,tl) - - | IDENT "Syntax"; u = univ; el = LIST1 syntax_entry SEP ";" -> - VernacSyntax (u,el) - - | IDENT "Uninterpreted"; IDENT "Notation"; local = locality; s = lstring; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - (s8,mv8) = - [IDENT "V8only"; - s8=OPT lstring; - mv8=OPT["(";mv8=LIST1 syntax_modifier SEP ","; ")" -> mv8] -> - (s8,mv8) - | -> (None,None)] -> - let s8 = match s8 with Some s -> s | _ -> s in - let mv8 = match mv8 with - Some mv8 -> mv8 - | _ -> List.map map_modl modl in - VernacSyntaxExtension (local,Some (s,modl),Some(s8,mv8)) - - | IDENT "Uninterpreted"; IDENT "V8Notation"; local = locality; s = lstring; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] -> - VernacSyntaxExtension (local,None,Some(s,modl)) - - | IDENT "Open"; local = locality; IDENT "Scope"; - sc = IDENT -> VernacOpenCloseScope (local,true, sc) - - | IDENT "Close"; local = locality; IDENT "Scope"; - sc = IDENT -> VernacOpenCloseScope (local,false,sc) - - | IDENT "Delimits"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT -> - VernacDelimiters (sc,key) - - | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with"; - refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl) - - | IDENT "Arguments"; IDENT "Scope"; qid = global; - "["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl) - - | IDENT "Infix"; local = locality; a = entry_prec; n = OPT natural; - op = lstring; - p = global; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc]; - mv8 = - [IDENT "V8only"; - a8=entry_prec; - n8=OPT natural; - op8=OPT lstring; - mv8=["("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> []] - -> - (match (a8,n8,mv8,op8) with - | None,None,[],None -> None - | _,_,mv8,_ -> - Some(op8,Metasyntax.merge_modifiers a8 n8 mv8)) - | -> (* Means: rules are based on V7 rules *) - Some (None,[]) ] -> - let mv = Metasyntax.merge_modifiers a n modl in - let v8 = Util.option_app (function (op8,mv8) -> - let op8 = match op8 with None -> op | Some op -> op in - let mv8 = - if mv8=[] then - let mv8 = List.map map_modl mv in - let mv8 = if List.for_all - (function SetLevel _ -> false | _ -> true) mv8 - then SetLevel 10 :: mv8 else mv8 in - let mv8 = if List.for_all - (function SetAssoc _ -> false | _ -> true) mv8 - then SetAssoc Gramext.LeftA :: mv8 else mv8 in - mv8 - else mv8 in - (op8,mv8)) mv8 in - VernacInfix (local,(op,mv),p,v8,sc) - | IDENT "Distfix"; local = locality; a = entry_prec; n = natural; - s = lstring; p = global; sc = OPT [ ":"; sc = IDENT -> sc ] -> - let (a,s,c) = Metasyntax.translate_distfix a s p in - let mv = Some(s,[SetLevel n;SetAssoc a]) in - VernacNotation (local,c,mv,mv,sc) -(* - VernacDistfix (local,a,n,s,p,sc) -*) - | IDENT "Notation"; local = locality; id = ident; ":="; c = constr; - b = [ "("; IDENT "only"; IDENT "parsing"; ")" -> true | -> false ] -> - VernacSyntacticDefinition (id,c,local,b) - | IDENT "Notation"; local = locality; s = lstring; ":="; c = constr; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ]; - (s8,mv8) = - [IDENT "V8only"; - s8=OPT lstring; - mv8=OPT["(";mv8=LIST1 syntax_modifier SEP ","; ")" -> mv8] -> - (s8,mv8) - | -> (* Means: rules are based on V7 rules *) - None, Some [] ] -> - let smv8 = match s8,mv8 with - | None, None -> None (* = only interpretation *) - | Some s8, None -> Some (s8,[]) (* = only interp, new s *) - | None, Some [] -> Some (s,List.map map_modl modl) (*like v7*) - | None, Some mv8 -> Some (s,mv8) (* s like v7 *) - | Some s8, Some mv8 -> Some (s8,mv8) in - VernacNotation (local,c,Some(s,modl),smv8,sc) - | IDENT "V8Notation"; local = locality; s = lstring; ":="; c = constr; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ] -> - VernacNotation (local,c,None,Some(s,modl),sc) - - | IDENT "V8Infix"; local = locality; op8 = lstring; p = global; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc] -> - let mv8 = Metasyntax.merge_modifiers None None modl in - VernacInfix (local,("",[]),p,Some (op8,mv8),sc) - - (* "Print" "Grammar" should be here but is in "command" entry in order - to factorize with other "Print"-based vernac entries *) - ] ] - ; - locality: - [ [ IDENT "Local" -> true | -> false ] ] - ; - level: - [ [ IDENT "level"; n = natural -> NumLevel n - | IDENT "next"; IDENT "level" -> NextLevel ] ] - ; - syntax_modifier: - [ [ x = IDENT; IDENT "at"; lev = level -> SetItemLevel ([x],lev) - | x = IDENT; ","; l = LIST1 IDENT SEP ","; IDENT "at"; lev = level -> - SetItemLevel (x::l,lev) - | IDENT "at"; IDENT "level"; n = natural -> SetLevel n - | IDENT "left"; IDENT "associativity" -> SetAssoc Gramext.LeftA - | IDENT "right"; IDENT "associativity" -> SetAssoc Gramext.RightA - | IDENT "no"; IDENT "associativity" -> SetAssoc Gramext.NonA - | x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ) - | IDENT "only"; IDENT "parsing" -> SetOnlyParsing - | IDENT "format"; s = [s = lstring -> (loc,s)] -> SetFormat s ] ] - ; - syntax_extension_type: - [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference - | IDENT "bigint" -> ETBigint - | i=IDENT -> ETOther ("constr",i) - ] ] - ; - opt_scope: - [ [ IDENT "_" -> None | sc = IDENT -> Some sc ] ] - ; - (* Syntax entries for Grammar. Only grammar_entry is exported *) - grammar_entry: - [[ nont = IDENT; set_entry_type; ":="; - ep = entry_prec; OPT "|"; rl = LIST0 grammar_rule SEP "|" -> - (rename_command_entry nont,ep,rl) ]] - ; - entry_prec: - [[ IDENT "LEFTA" -> Some Gramext.LeftA - | IDENT "RIGHTA" -> Some Gramext.RightA - | IDENT "NONA" -> Some Gramext.NonA - | -> None ]] - ; - grammar_tactic_rule: - [[ name = rule_name; "["; pil = LIST1 production_item; "]"; - "->"; "["; t = Tactic.tactic; "]" -> (name, pil, t) ]] - ; - grammar_rule: - [[ name = rule_name; "["; pil = LIST0 production_item; "]"; "->"; - a = action -> (name, pil, a) ]] - ; - rule_name: - [[ name = IDENT -> name ]] - ; - production_item: - [[ s = lstring -> VTerm s - | nt = non_terminal; po = OPT [ "("; p = METAIDENT; ")" -> p ] -> - match po with - | Some p -> VNonTerm (loc,nt,Some (Names.id_of_string p)) - | _ -> VNonTerm (loc,nt,None) ]] - ; - non_terminal: - [[ u = IDENT; ":"; nt = IDENT -> - NtQual(rename_command_entry u, rename_command_entry nt) - | nt = IDENT -> NtShort (rename_command_entry nt) ]] - ; - - - (* Syntax entries for Syntax. Only syntax_entry is exported *) - syntax_entry: - [ [ IDENT "level"; p = precedence; ":"; - OPT "|"; rl = LIST1 syntax_rule SEP "|" -> (p,rl) ] ] - ; - syntax_rule: - [ [ nm = IDENT; "["; s = astpat; "]"; "->"; u = unparsing -> (nm,s,u) ] ] - ; - precedence: - [ [ a = natural -> a -(* | "["; a1 = natural; a2 = natural; a3 = natural; "]" -> (a1,a2,a3)*) - ] ] - ; - unparsing: - [ [ "["; ll = LIST0 next_hunks; "]" -> ll ] ] - ; - next_hunks: - [ [ IDENT "FNL" -> UNP_FNL - | IDENT "TAB" -> UNP_TAB - | c = lstring -> RO c - | "["; - x = - [ 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 (ext,pr) -> PH (e,ext,pr) - | None -> PH (e,None,Any) - ]] - ; - box: - [ [ "<"; bk = box_kind; ">" -> bk ] ] - ; - box_kind: - [ [ 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" -> None, L - | IDENT "E" -> None, E - | pprim = lstring; precrec = OPT [ ":"; p = precedence -> p ] -> - match precrec with - | Some p -> Some pprim, Prec p - | None -> Some pprim, Any ] ] - ; - (* meta-syntax entries *) - astpat: - [ [ "<<" ; a = Prim.ast; ">>" -> a - | a = Constr.constr -> - Termast.ast_of_rawconstr - (Constrintern.intern_constr Evd.empty (Global.env()) a) - ] ] - ; - 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: - [[ IDENT "ast"; IDENT "list" -> Util.error "type ast list no longer supported" - | IDENT "ast" -> Util.error "type ast no longer supported" - | IDENT "constr" -> ConstrParser - | IDENT "pattern" -> CasesPatternParser - | IDENT "tactic" -> assert false - | IDENT "vernac" -> Util.error "vernac extensions no longer supported" ] ] - ; -END diff --git a/parsing/g_cases.ml4 b/parsing/g_cases.ml4 deleted file mode 100644 index 7ad53edb9..000000000 --- a/parsing/g_cases.ml4 +++ /dev/null @@ -1,79 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pcoq -open Constr -open Topconstr -open Term -open Libnames - -open Prim - -let pair loc = - Qualid (loc, Libnames.qualid_of_string "Coq.Init.Datatypes.pair") - -if !Options.v7 then -GEXTEND Gram - GLOBAL: operconstr pattern; - - pattern: - [ [ r = Prim.reference -> CPatAtom (loc,Some r) - | IDENT "_" -> CPatAtom (loc,None) - (* Hack to parse syntax "(n)" as a natural number *) - | "("; G_constr.test_int_rparen; n = INT; ")" -> - (* Delimiter "N" moved to "nat" in V7 *) - CPatDelimiters (loc,"nat",CPatNumeral (loc,Bigint.of_string n)) - | "("; p = compound_pattern; ")" -> p - | n = INT -> CPatNumeral (loc,Bigint.of_string n) - | "-"; n = INT -> - let n = Bigint.of_string n in - if n = Bigint.zero then - CPatNotation(loc,"( _ )",[CPatNumeral (loc,n)]) - else - CPatNumeral (loc,Bigint.neg n) - | "'"; G_constr.test_ident_colon; key = IDENT; ":"; c = pattern; "'" -> - CPatDelimiters (loc,key,c) - ] ] - ; - compound_pattern: - [ [ p = pattern ; lp = LIST1 pattern -> - (match p with - | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp) - | _ -> Util.user_err_loc - (loc, "compound_pattern", Pp.str "Constructor expected")) - | p = pattern; "as"; id = base_ident -> - CPatAlias (loc, p, id) - | p1 = pattern; ","; p2 = pattern -> - CPatCstr (loc, pair loc, [p1; p2]) - | p = pattern -> p ] ] - ; - equation: - [ [ lhs = LIST1 pattern; "=>"; rhs = operconstr LEVEL "9" -> (loc,lhs,rhs) ] ] - ; - ne_eqn_list: - [ [ leqn = LIST1 equation SEP "|" -> leqn ] ] - ; - operconstr: LEVEL "1" - [ [ "<"; p = annot; ">"; "Cases"; lc = LIST1 constr; "of"; - OPT "|"; eqs = ne_eqn_list; "end" -> - let lc = List.map (fun c -> c,(None,None)) lc in - CCases (loc, (Some p,None), lc, eqs) - | "Cases"; lc = LIST1 constr; "of"; - OPT "|"; eqs = ne_eqn_list; "end" -> - let lc = List.map (fun c -> c,(None,None)) lc in - CCases (loc, (None,None), lc, eqs) - | "<"; p = annot; ">"; "Cases"; lc = LIST1 constr; "of"; "end" -> - let lc = List.map (fun c -> c,(None,None)) lc in - CCases (loc, (Some p,None), lc, []) - | "Cases"; lc = LIST1 constr; "of"; "end" -> - let lc = List.map (fun c -> c,(None,None)) lc in - CCases (loc, (None,None), lc, []) ] ] - ; -END; diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 722703c1f..1f8422111 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -10,125 +10,125 @@ open Pcoq open Constr +open Prim open Rawterm open Term open Names open Libnames -open Prim open Topconstr -(* Initialize the lexer *) +open Util + let constr_kw = - [ "Cases"; "of"; "with"; "end"; "as"; "in"; "Prop"; "Set"; "Type"; - ":"; "("; ")"; "["; "]"; "{"; "}"; ","; ";"; "->"; "="; ":="; "!"; - "::"; "<:"; ":<"; "=>"; "<"; ">"; "|"; "?"; "/"; - "<->"; "\\/"; "/\\"; "`"; "``"; "&"; "*"; "+"; "@"; "^"; "#"; "-"; - "~"; "'"; "<<"; ">>"; "<>"; ".." - ] -let _ = - if !Options.v7 then - List.iter (fun s -> Lexer.add_token ("",s)) constr_kw -(* "let" is not a keyword because #Core#let.cci would not parse. - Is it still accurate ? *) + [ "forall"; "fun"; "match"; "fix"; "cofix"; "with"; "in"; "for"; + "end"; "as"; "let"; "if"; "then"; "else"; "return"; + "Prop"; "Set"; "Type"; ".("; "_"; ".." ] +let _ = List.iter (fun s -> Lexer.add_token("",s)) constr_kw -let coerce_to_var = function - | CRef (Ident (_,id)) -> id - | ast -> Util.user_err_loc - (constr_loc ast,"Ast.coerce_to_var", - (Pp.str"This expression should be a simple identifier")) +let mk_cast = function + (c,(_,None)) -> c + | (c,(_,Some ty)) -> CCast(join_loc (constr_loc c) (constr_loc ty), c, DEFAULTcast,ty) -let coerce_to_name = function - | CRef (Ident (loc,id)) -> (loc, Name id) - | ast -> Util.user_err_loc - (constr_loc ast,"Ast.coerce_to_var", - (Pp.str"This expression should be a simple identifier")) +let mk_lam = function + ([],c) -> c + | (bl,c) -> CLambdaN(constr_loc c, bl,c) -let set_loc loc = function - | CRef(Ident(_,i)) -> CRef(Ident(loc,i)) - | CRef(Qualid(_,q)) -> CRef(Qualid(loc,q)) - | CFix(_,x,a) -> CFix(loc,x,a) - | CCoFix(_,x,a) -> CCoFix(loc,x,a) - | CArrow(_,a,b) -> CArrow(loc,a,b) - | CProdN(_,bl,a) -> CProdN(loc,bl,a) - | CLambdaN(_,bl,a) -> CLambdaN(loc,bl,a) - | CLetIn(_,x,a,b) -> CLetIn(loc,x,a,b) - | CAppExpl(_,f,a) -> CAppExpl(loc,f,a) - | CApp(_,f,a) -> CApp(loc,f,a) - | CCases(_,p,a,br) -> CCases(loc,p,a,br) - | COrderedCase(_,s,p,a,br) -> COrderedCase(loc,s,p,a,br) - | CLetTuple(_,ids,p,a,b) -> CLetTuple(loc,ids,p,a,b) - | CIf(_,e,p,a,b) -> CIf(loc,e,p,a,b) - | CHole _ -> CHole loc - | CPatVar(_,v) -> CPatVar(loc,v) - | CEvar(_,ev) -> CEvar(loc,ev) - | CSort(_,s) -> CSort(loc,s) - | CCast(_,a,k,b) -> CCast(loc,a,k,b) - | CNotation(_,n,l) -> CNotation(loc,n,l) - | CNumeral(_,i) -> CNumeral(loc,i) - | CDelimiters(_,s,e) -> CDelimiters(loc,s,e) - | CDynamic(_,d) -> CDynamic(loc,d) +let loc_of_binder_let = function + | LocalRawAssum ((loc,_)::_,_)::_ -> loc + | LocalRawDef ((loc,_),_)::_ -> loc + | _ -> dummy_loc -open Util +let rec mkCProdN loc bll c = + match bll with + | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> + CProdN (loc,[idl,t],mkCProdN (join_loc loc1 loc) bll c) + | LocalRawDef ((loc1,_) as id,b) :: bll -> + CLetIn (loc,id,b,mkCProdN (join_loc loc1 loc) bll c) + | [] -> c + | LocalRawAssum ([],_) :: bll -> mkCProdN loc bll c -let rec abstract_constr loc c = function +let rec mkCLambdaN loc bll c = + match bll with + | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> + CLambdaN (loc,[idl,t],mkCLambdaN (join_loc loc1 loc) bll c) + | LocalRawDef ((loc1,_) as id,b) :: bll -> + CLetIn (loc,id,b,mkCLambdaN (join_loc loc1 loc) bll c) | [] -> c - | LocalRawDef ((loc',x),b)::bl -> - CLetIn (join_loc loc' loc, (loc', x), b, abstract_constr loc c bl) - | LocalRawAssum (nal,t)::bl -> - let loc' = join_loc (fst (List.hd nal)) loc in - CLambdaN(loc', [nal, t], abstract_constr loc c bl) + | LocalRawAssum ([],_) :: bll -> mkCLambdaN loc bll c -(* Hack to parse "(n)" as nat without conflicts with the (useless) *) -(* admissible notation "(n)" *) -let test_int_rparen = - Gram.Entry.of_parser "test_int_rparen" - (fun strm -> - match Stream.npeek 1 strm with - | [("INT", _)] -> - begin match Stream.npeek 2 strm with - | [_; ("", ")")] -> () - | _ -> raise Stream.Failure - end - | _ -> raise Stream.Failure) +let rec index_of_annot loc bl ann = + match names_of_local_assums bl,ann with + | [_], None -> 0 + | lids, Some x -> + let ids = List.map snd lids in + (try list_index (snd x) ids - 1 + with Not_found -> + user_err_loc(fst x,"index_of_annot", Pp.str"no such fix variable")) + | _ -> user_err_loc(loc,"index_of_annot", + Pp.str "cannot guess decreasing argument of fix") -(* Hack to parse "n" at level 0 without conflicting with "n!" at level 91 *) -(* admissible notation "(n)" *) -let test_int_bang = - Gram.Entry.of_parser "test_int_bang" - (fun strm -> - match Stream.npeek 1 strm with - | [("INT", n)] -> - begin match Stream.npeek 2 strm with - | [_; ("", "!")] -> () - | _ -> raise Stream.Failure - end - | _ -> raise Stream.Failure) +let mk_fixb (id,bl,ann,body,(loc,tyc)) = + let n = index_of_annot (fst id) bl ann in + let ty = match tyc with + Some ty -> ty + | None -> CHole loc in + (snd id,n,bl,ty,body) + +let mk_cofixb (id,bl,ann,body,(loc,tyc)) = + let _ = option_app (fun (aloc,_) -> + Util.user_err_loc + (aloc,"Constr:mk_cofixb", + Pp.str"Annotation forbidden in cofix expression")) ann in + let ty = match tyc with + Some ty -> ty + | None -> CHole loc in + (snd id,bl,ty,body) + +let mk_fix(loc,kw,id,dcls) = + if kw then + let fb = List.map mk_fixb dcls in + CFix(loc,id,fb) + else + let fb = List.map mk_cofixb dcls in + CCoFix(loc,id,fb) + +let mk_single_fix (loc,kw,dcl) = + let (id,_,_,_,_) = dcl in mk_fix(loc,kw,id,[dcl]) -(* Hack to parse "`id:...`" at level 0 without conflicting with - "`...`" from ZArith *) -let test_ident_colon = - Gram.Entry.of_parser "test_ident_colon" +let binder_constr = + create_constr_entry (get_univ "constr") "binder_constr" + +(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *) +(* admissible notation "(x t)" *) +let lpar_id_coloneq = + Gram.Entry.of_parser "test_lpar_id_coloneq" (fun strm -> match Stream.npeek 1 strm with - | [("IDENT", _)] -> - begin match Stream.npeek 2 strm with - | [_; ("", ":")] -> () - | _ -> raise Stream.Failure - end + | [("","(")] -> + (match Stream.npeek 2 strm with + | [_; ("IDENT",s)] -> + (match Stream.npeek 3 strm with + | [_; _; ("", ":=")] -> + Stream.junk strm; Stream.junk strm; Stream.junk strm; + Names.id_of_string s + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) | _ -> raise Stream.Failure) -if !Options.v7 then GEXTEND Gram - GLOBAL: operconstr lconstr constr sort global constr_pattern Constr.ident annot - (*ne_name_comma_list*); + GLOBAL: binder_constr lconstr constr operconstr sort global + constr_pattern lconstr_pattern Constr.ident binder binder_let pattern; Constr.ident: [ [ id = Prim.ident -> id (* This is used in quotations and Syntax *) | id = METAIDENT -> id_of_string id ] ] ; + Prim.name: + [ [ "_" -> (loc, Anonymous) ] ] + ; global: [ [ r = Prim.reference -> r @@ -138,238 +138,193 @@ GEXTEND Gram constr_pattern: [ [ c = constr -> c ] ] ; - ne_constr_list: - [ [ cl = LIST1 constr -> cl ] ] + lconstr_pattern: + [ [ c = lconstr -> c ] ] ; sort: [ [ "Set" -> RProp Pos | "Prop" -> RProp Null | "Type" -> RType None ] ] ; - constr: - [ [ c = operconstr LEVEL "8" -> c ] ] - ; lconstr: - [ [ c = operconstr LEVEL "10" -> c ] ] + [ [ c = operconstr LEVEL "200" -> c ] ] + ; + constr: + [ [ c = operconstr LEVEL "9" -> c ] ] ; operconstr: - [ "10" RIGHTA - [ "!"; f = global; args = LIST0 (operconstr LEVEL "9") -> - CAppExpl (loc, (None,f), args) -(* - | "!"; f = global; "with"; b = binding_list -> - <:ast< (APPLISTWITH $f $b) >> -*) - | f = operconstr; args = LIST1 constr91 -> CApp (loc, (None,f), args) ] - | "9" RIGHTA - [ c1 = operconstr; "::"; c2 = operconstr LEVEL "9" -> - CCast (loc, c1, Term.DEFAULTcast,c2) ] - | "8" RIGHTA - [ c1 = operconstr; "->"; c2 = operconstr LEVEL "8"-> CArrow (loc, c1, c2) ] - | "1" RIGHTA - [ "<"; p = annot; ">"; IDENT "Match"; c = constr; "with"; - cl = LIST0 constr; "end" -> - COrderedCase (loc, MatchStyle, Some p, c, cl) - | "<"; p = annot; ">"; IDENT "Case"; c = constr; "of"; - cl = LIST0 constr; "end" -> - COrderedCase (loc, RegularStyle, Some p, c, cl) - | IDENT "Case"; c = constr; "of"; cl = LIST0 constr; "end" -> - COrderedCase (loc, RegularStyle, None, c, cl) - | IDENT "Match"; c = constr; "with"; cl = LIST1 constr; "end" -> - COrderedCase (loc, MatchStyle, None, c, cl) - | IDENT "let"; "("; b = ne_name_comma_list; ")"; "="; - c = constr; "in"; c1 = constr -> - (* TODO: right loc *) - COrderedCase - (loc, LetStyle, None, c, [CLambdaN (loc, [b, CHole loc], c1)]) - | IDENT "let"; na = name; "="; c = opt_casted_constr; - "in"; c1 = constr -> - CLetIn (loc, na, c, c1) - | IDENT "if"; c1 = constr; - IDENT "then"; c2 = constr; - IDENT "else"; c3 = constr -> - COrderedCase (loc, IfStyle, None, c1, [c2; c3]) - | "<"; p = annot; ">"; - IDENT "let"; "("; b = ne_name_comma_list; ")"; "="; c = constr; - "in"; c1 = constr -> - (* TODO: right loc *) - COrderedCase (loc, LetStyle, Some p, c, - [CLambdaN (loc, [b, CHole loc], c1)]) - | "<"; p = annot; ">"; - IDENT "if"; c1 = constr; - IDENT "then"; c2 = constr; - IDENT "else"; c3 = constr -> - COrderedCase (loc, IfStyle, Some p, c1, [c2; c3]) - | ".."; c = operconstr LEVEL "0"; ".." -> + [ "200" RIGHTA + [ c = binder_constr -> c ] + | "100" RIGHTA + [ c1 = operconstr; ":"; c2 = binder_constr -> CCast(loc,c1,DEFAULTcast,c2) + | c1 = operconstr; ":"; c2 = SELF -> CCast(loc,c1,DEFAULTcast,c2) ] + | "99" RIGHTA [ ] + | "90" RIGHTA + [ c1 = operconstr; "->"; c2 = binder_constr -> CArrow(loc,c1,c2) + | c1 = operconstr; "->"; c2 = SELF -> CArrow(loc,c1,c2)] + | "10" LEFTA + [ f=operconstr; args=LIST1 appl_arg -> CApp(loc,(None,f),args) + | "@"; f=global; args=LIST0 NEXT -> CAppExpl(loc,(None,f),args) ] + | "9" + [ ".."; c = operconstr LEVEL "0"; ".." -> CAppExpl (loc,(None,Ident (loc,Topconstr.ldots_var)),[c]) ] - | "0" RIGHTA - [ "?" -> CHole loc - | "?"; n = Prim.natural -> CPatVar (loc, (false,Pattern.patvar_of_int n)) - | bll = binders; c = constr -> abstract_constr loc c bll - (* Hack to parse syntax "(n)" as a natural number *) - | "("; test_int_rparen; n = INT; ")" -> - (* Delimiter "N" moved to "nat" in V7 *) - CDelimiters (loc,"nat",CNumeral (loc,Bigint.of_string n)) - | "("; lc1 = lconstr; ":"; c = constr; (bl,body) = product_tail -> - let id = coerce_to_name lc1 in - CProdN (loc, ([id], c)::bl, body) -(* TODO: Syntaxe (_:t...)t et (_,x...)t *) - | "("; lc1 = lconstr; ","; lc2 = lconstr; ":"; c = constr; - (bl,body) = product_tail -> - let id1 = coerce_to_name lc1 in - let id2 = coerce_to_name lc2 in - CProdN (loc, ([id1; id2], c)::bl, body) - | "("; lc1 = lconstr; ","; lc2 = lconstr; ","; - idl = ne_name_comma_list; ":"; c = constr; (bl,body) = product_tail -> - let id1 = coerce_to_name lc1 in - let id2 = coerce_to_name lc2 in - CProdN (loc, (id1::id2::idl, c)::bl, body) - | "("; lc1 = lconstr; ")" -> - if Options.do_translate() then set_loc loc lc1 else lc1 - | "("; lc1 = lconstr; ")"; "@"; "["; cl = ne_constr_list; "]" -> - (match lc1 with - | CPatVar (loc2,(false,n)) -> - CApp (loc,(None,CPatVar (loc2, (true,n))), List.map (fun c -> c, None) cl) - | _ -> - Util.error "Second-order pattern-matching expects a head metavariable") - | IDENT "Fix"; id = identref; "{"; fbinders = fixbinders; "}" -> - CFix (loc, id, fbinders) - | IDENT "CoFix"; id = identref; "{"; fbinders = cofixbinders; "}" -> - CCoFix (loc, id, fbinders) - | IDENT "Prefix" ; "(" ; s = STRING ; cl = LIST0 constr ; ")" -> - CNotation(loc, s, cl) - | s = sort -> CSort (loc, s) - | v = global -> CRef v - | (b,n) = bigint -> - if n = Bigint.zero & b then CNotation(loc,"( _ )",[CNumeral (loc,n)]) - else CNumeral (loc,n) - | "!"; f = global -> CAppExpl (loc,(None,f),[]) - | "'"; test_ident_colon; key = IDENT; ":"; c = constr; "'" -> - (* Delimiter "N" implicitly moved to "nat" in V7 *) - let key = if key = "N" then "nat" else key in - let key = if key = "P" then "positive" else key in - let key = if key = "T" then "type" else key in - CDelimiters (loc,key,c) ] ] + | "1" LEFTA + [ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" -> + CApp(loc,(Some (List.length args+1),CRef f),args@[c,None]) + | c=operconstr; ".("; "@"; f=global; + args=LIST0 (operconstr LEVEL "9"); ")" -> + CAppExpl(loc,(Some (List.length args+1),f),args@[c]) + | c=operconstr; "%"; key=IDENT -> CDelimiters (loc,key,c) ] + | "0" + [ c=atomic_constr -> c + | c=match_constr -> c + | "("; c = operconstr LEVEL "200"; ")" -> + (match c with + CNumeral(_,z) when Bigint.is_pos_or_zero z -> + CNotation(loc,"( _ )",[c]) + | _ -> c) ] ] + ; + binder_constr: + [ [ "forall"; bl = binder_list; ","; c = operconstr LEVEL "200" -> + mkCProdN loc bl c + | "fun"; bl = binder_list; "=>"; c = operconstr LEVEL "200" -> + mkCLambdaN loc bl c + | "let"; id=name; bl = LIST0 binder_let; ty = type_cstr; ":="; + c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" -> + let loc1 = loc_of_binder_let bl in + CLetIn(loc,id,mkCLambdaN loc1 bl (mk_cast(c1,ty)),c2) + | "let"; fx = single_fix; "in"; c = operconstr LEVEL "200" -> + let fixp = mk_single_fix fx in + let (li,id) = match fixp with + CFix(_,id,_) -> id + | CCoFix(_,id,_) -> id + | _ -> assert false in + CLetIn(loc,(li,Name id),fixp,c) + | "let"; lb = ["("; l=LIST0 name SEP ","; ")" -> l | "()" -> []]; + po = return_type; + ":="; c1 = operconstr LEVEL "200"; "in"; + c2 = operconstr LEVEL "200" -> + CLetTuple (loc,List.map snd lb,po,c1,c2) + | "if"; c=operconstr LEVEL "200"; po = return_type; + "then"; b1=operconstr LEVEL "200"; + "else"; b2=operconstr LEVEL "200" -> + CIf (loc, c, po, b1, b2) + | c=fix_constr -> c ] ] ; - constr91: - [ [ test_int_bang; n = INT; "!"; c = operconstr LEVEL "9" -> - (c, Some (loc,ExplByPos (int_of_string n))) - | c = operconstr LEVEL "9" -> (c, None) ] ] + appl_arg: + [ [ id = lpar_id_coloneq; c=lconstr; ")" -> + (c,Some (loc,ExplByName id)) + | c=constr -> (c,None) ] ] ; - (* annot and product_annot_tail are hacks to forbid concrete syntax *) - (* ">" (e.g. for gt, Zgt, ...) in annotations *) - annot: - [ RIGHTA - [ bll = binders; c = annot -> abstract_constr loc c bll - | "("; lc1 = lconstr; ":"; c = constr; (bl,body) = product_annot_tail -> - let id = coerce_to_name lc1 in - CProdN (loc, ([id], c)::bl, body) - | "("; lc1 = lconstr; ","; lc2 = lconstr; ":"; c = constr; - (bl,body) = product_annot_tail -> - let id1 = coerce_to_name lc1 in - let id2 = coerce_to_name lc2 in - CProdN (loc, ([id1; id2], c)::bl, body) - | "("; lc1 = lconstr; ","; lc2 = lconstr; ","; - idl = ne_name_comma_list; ":"; c = constr; - (bl,body) = product_annot_tail -> - let id1 = coerce_to_name lc1 in - let id2 = coerce_to_name lc2 in - CProdN (loc, (id1::id2::idl, c)::bl, body) - | "("; lc1 = lconstr; ")" -> lc1 - | c1 = annot; "->"; c2 = annot -> CArrow (loc, c1, c2) ] - | RIGHTA - [ c1 = annot; "\\/"; c2 = annot -> CNotation (loc, "_ \\/ _", [c1;c2]) ] - | RIGHTA - [ c1 = annot; "/\\"; c2 = annot -> CNotation (loc, "_ /\\ _", [c1;c2]) ] - | RIGHTA - [ "~"; c = SELF -> CNotation (loc, "~ _", [c]) ] - | RIGHTA - [ c1 = SELF; "=="; c2 = NEXT -> CNotation (loc, "_ == _", [c1;c2]) ] - | RIGHTA - [ c1 = SELF; "="; c2 = NEXT -> CNotation (loc, "_ = _", [c1;c2]) ] - | [ c = operconstr LEVEL "4L" -> c ] ] + atomic_constr: + [ [ g=global -> CRef g + | s=sort -> CSort(loc,s) + | n=INT -> CNumeral (loc, Bigint.of_string n) + | "_" -> CHole loc + | "?"; id=ident -> CPatVar(loc,(false,id)) ] ] ; - product_annot_tail: - [ [ ";"; idl = ne_name_comma_list; ":"; c = constr; - (bl,c2) = product_annot_tail -> ((idl, c)::bl, c2) - | ";"; idl = ne_name_comma_list; (bl,c2) = product_annot_tail -> - ((idl, CHole (fst (List.hd idl)))::bl, c2) - | ")"; c = annot -> ([], c) ] ] + fix_constr: + [ [ fx1=single_fix -> mk_single_fix fx1 + | (_,kw,dcl1)=single_fix; "with"; dcls=LIST1 fix_decl SEP "with"; + "for"; id=identref -> + mk_fix(loc,kw,id,dcl1::dcls) + ] ] ; - ne_name_comma_list: - [ [ nal = LIST1 name SEP "," -> nal ] ] + single_fix: + [ [ kw=fix_kw; dcl=fix_decl -> (loc,kw,dcl) ] ] ; - name_comma_list_tail: - [ [ ","; idl = ne_name_comma_list -> idl - | -> [] ] ] + fix_kw: + [ [ "fix" -> true + | "cofix" -> false ] ] ; - opt_casted_constr: - [ [ c = constr; ":"; t = constr -> CCast (loc, c, DEFAULTcast, t) - | c = constr -> c ] ] + fix_decl: + [ [ id=identref; bl=LIST0 binder_let; ann=fixannot; ty=type_cstr; ":="; + c=operconstr LEVEL "200" -> (id,bl,ann,c,ty) ] ] ; - vardecls: - [ [ na = name; nal = name_comma_list_tail; c = type_option -> - LocalRawAssum (na::nal,c) - | na = name; "="; c = opt_casted_constr -> - LocalRawDef (na, c) - | na = name; ":="; c = opt_casted_constr -> - LocalRawDef (na, c) - - (* This is used in quotations *) - | id = METAIDENT; c = type_option -> LocalRawAssum ([loc, Name (id_of_string id)], c) - ] ] + fixannot: + [ [ "{"; IDENT "struct"; id=name; "}" -> Some id + | -> None ] ] ; - ne_vardecls_list: - [ [ id = vardecls; ";"; idl = ne_vardecls_list -> id :: idl - | id = vardecls -> [id] ] ] + match_constr: + [ [ "match"; ci=LIST1 case_item SEP ","; ty=OPT case_type; "with"; + br=branches; "end" -> CCases(loc,ty,ci,br) ] ] ; - binders: - [ [ "["; bl = ne_vardecls_list; "]" -> bl ] ] + case_item: + [ [ c=operconstr LEVEL "100"; p=pred_pattern -> (c,p) ] ] ; - simple_params: - [ [ idl = LIST1 name SEP ","; ":"; c = constr -> (idl, c) - | idl = LIST1 name SEP "," -> (idl, CHole loc) - ] ] + pred_pattern: + [ [ ona = OPT ["as"; id=name -> snd id]; + ty = OPT ["in"; t=lconstr -> t] -> (ona,ty) ] ] ; - simple_binders: - [ [ "["; bll = LIST1 simple_params SEP ";"; "]" -> bll ] ] + case_type: + [ [ "return"; ty = operconstr LEVEL "100" -> ty ] ] ; - ne_simple_binders_list: - [ [ bll = LIST1 simple_binders -> List.flatten bll ] ] + return_type: + [ [ a = OPT [ na = OPT["as"; id=name -> snd id]; + ty = case_type -> (na,ty) ] -> + match a with + | None -> None, None + | Some (na,t) -> (na, Some t) + ] ] ; - type_option: - [ [ ":"; c = constr -> c - | -> CHole loc ] ] + branches: + [ [ OPT"|"; br=LIST0 eqn SEP "|" -> br ] ] ; - fixbinder: - [ [ id = base_ident; "/"; recarg = natural; ":"; type_ = constr; - ":="; def = constr -> - (id, recarg-1, [], type_, def) - | id = base_ident; bl = ne_simple_binders_list; ":"; type_ = constr; - ":="; def = constr -> - let ni = List.length (List.flatten (List.map fst bl)) -1 in - let bl = List.map (fun(nal,ty)->LocalRawAssum(nal,ty)) bl in - (id, ni, bl, type_, def) ] ] + eqn: + [ [ pl = LIST1 pattern SEP ","; "=>"; rhs = lconstr -> (loc,pl,rhs) ] ] ; - fixbinders: - [ [ fbs = LIST1 fixbinder SEP "with" -> fbs ] ] + pattern: + [ "200" RIGHTA [ ] + | "100" LEFTA + [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> CPatOr (loc,p::pl) ] + | "99" RIGHTA [ ] + | "10" LEFTA + [ p = pattern; lp = LIST1 (pattern LEVEL "0") -> + (match p with + | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp) + | _ -> Util.user_err_loc + (cases_pattern_loc p, "compound_pattern", + Pp.str "Constructor expected")) + | p = pattern; "as"; id = ident -> + CPatAlias (loc, p, id) + | c = pattern; "%"; key=IDENT -> + CPatDelimiters (loc,key,c) ] + | "0" + [ r = Prim.reference -> CPatAtom (loc,Some r) + | "_" -> CPatAtom (loc,None) + | "("; p = pattern LEVEL "200"; ")" -> + (match p with + CPatNumeral(_,z) when Bigint.is_pos_or_zero z -> + CPatNotation(loc,"( _ )",[p]) + | _ -> p) + | n = INT -> CPatNumeral (loc, Bigint.of_string n) ] ] ; - cofixbinder: - [ [ id = base_ident; ":"; type_ = constr; ":="; def = constr -> - (id, [],type_, def) ] ] + binder_list: + [ [ idl=LIST1 name; bl=LIST0 binder_let -> + LocalRawAssum (idl,CHole loc)::bl + | idl=LIST1 name; ":"; c=lconstr -> + [LocalRawAssum (idl,c)] + | "("; idl=LIST1 name; ":"; c=lconstr; ")"; bl=LIST0 binder_let -> + LocalRawAssum (idl,c)::bl ] ] ; - cofixbinders: - [ [ fbs = LIST1 cofixbinder SEP "with" -> fbs ] ] + binder_let: + [ [ id=name -> + LocalRawAssum ([id],CHole loc) + | "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" -> + LocalRawAssum (id::idl,c) + | "("; id=name; ":"; c=lconstr; ")" -> + LocalRawAssum ([id],c) + | "("; id=name; ":="; c=lconstr; ")" -> + LocalRawDef (id,c) + | "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" -> + LocalRawDef (id,CCast (join_loc (constr_loc t) loc,c,DEFAULTcast,t)) + ] ] ; - product_tail: - [ [ ";"; idl = ne_name_comma_list; ":"; c = constr; - (bl,c2) = product_tail -> ((idl, c)::bl, c2) - | ";"; idl = ne_name_comma_list; (bl,c2) = product_tail -> - ((idl, CHole (fst (List.hd idl)))::bl, c2) - | ")"; c = constr -> ([], c) ] ] + binder: + [ [ id=name -> ([id],CHole loc) + | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,c) ] ] ; - bigint: - [ [ i = INT -> true, Bigint.of_string i - | "-"; i = INT -> false, Bigint.neg (Bigint.of_string i) ] ] + type_cstr: + [ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ] ; -END;; + END;; diff --git a/parsing/g_constrnew.ml4 b/parsing/g_constrnew.ml4 deleted file mode 100644 index 1f8422111..000000000 --- a/parsing/g_constrnew.ml4 +++ /dev/null @@ -1,330 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pcoq -open Constr -open Prim -open Rawterm -open Term -open Names -open Libnames -open Topconstr - -open Util - -let constr_kw = - [ "forall"; "fun"; "match"; "fix"; "cofix"; "with"; "in"; "for"; - "end"; "as"; "let"; "if"; "then"; "else"; "return"; - "Prop"; "Set"; "Type"; ".("; "_"; ".." ] - -let _ = List.iter (fun s -> Lexer.add_token("",s)) constr_kw - -let mk_cast = function - (c,(_,None)) -> c - | (c,(_,Some ty)) -> CCast(join_loc (constr_loc c) (constr_loc ty), c, DEFAULTcast,ty) - -let mk_lam = function - ([],c) -> c - | (bl,c) -> CLambdaN(constr_loc c, bl,c) - -let loc_of_binder_let = function - | LocalRawAssum ((loc,_)::_,_)::_ -> loc - | LocalRawDef ((loc,_),_)::_ -> loc - | _ -> dummy_loc - -let rec mkCProdN loc bll c = - match bll with - | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> - CProdN (loc,[idl,t],mkCProdN (join_loc loc1 loc) bll c) - | LocalRawDef ((loc1,_) as id,b) :: bll -> - CLetIn (loc,id,b,mkCProdN (join_loc loc1 loc) bll c) - | [] -> c - | LocalRawAssum ([],_) :: bll -> mkCProdN loc bll c - -let rec mkCLambdaN loc bll c = - match bll with - | LocalRawAssum ((loc1,_)::_ as idl,t) :: bll -> - CLambdaN (loc,[idl,t],mkCLambdaN (join_loc loc1 loc) bll c) - | LocalRawDef ((loc1,_) as id,b) :: bll -> - CLetIn (loc,id,b,mkCLambdaN (join_loc loc1 loc) bll c) - | [] -> c - | LocalRawAssum ([],_) :: bll -> mkCLambdaN loc bll c - -let rec index_of_annot loc bl ann = - match names_of_local_assums bl,ann with - | [_], None -> 0 - | lids, Some x -> - let ids = List.map snd lids in - (try list_index (snd x) ids - 1 - with Not_found -> - user_err_loc(fst x,"index_of_annot", Pp.str"no such fix variable")) - | _ -> user_err_loc(loc,"index_of_annot", - Pp.str "cannot guess decreasing argument of fix") - -let mk_fixb (id,bl,ann,body,(loc,tyc)) = - let n = index_of_annot (fst id) bl ann in - let ty = match tyc with - Some ty -> ty - | None -> CHole loc in - (snd id,n,bl,ty,body) - -let mk_cofixb (id,bl,ann,body,(loc,tyc)) = - let _ = option_app (fun (aloc,_) -> - Util.user_err_loc - (aloc,"Constr:mk_cofixb", - Pp.str"Annotation forbidden in cofix expression")) ann in - let ty = match tyc with - Some ty -> ty - | None -> CHole loc in - (snd id,bl,ty,body) - -let mk_fix(loc,kw,id,dcls) = - if kw then - let fb = List.map mk_fixb dcls in - CFix(loc,id,fb) - else - let fb = List.map mk_cofixb dcls in - CCoFix(loc,id,fb) - -let mk_single_fix (loc,kw,dcl) = - let (id,_,_,_,_) = dcl in mk_fix(loc,kw,id,[dcl]) - -let binder_constr = - create_constr_entry (get_univ "constr") "binder_constr" - -(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *) -(* admissible notation "(x t)" *) -let lpar_id_coloneq = - Gram.Entry.of_parser "test_lpar_id_coloneq" - (fun strm -> - match Stream.npeek 1 strm with - | [("","(")] -> - (match Stream.npeek 2 strm with - | [_; ("IDENT",s)] -> - (match Stream.npeek 3 strm with - | [_; _; ("", ":=")] -> - Stream.junk strm; Stream.junk strm; Stream.junk strm; - Names.id_of_string s - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - - -GEXTEND Gram - GLOBAL: binder_constr lconstr constr operconstr sort global - constr_pattern lconstr_pattern Constr.ident binder binder_let pattern; - Constr.ident: - [ [ id = Prim.ident -> id - - (* This is used in quotations and Syntax *) - | id = METAIDENT -> id_of_string id ] ] - ; - Prim.name: - [ [ "_" -> (loc, Anonymous) ] ] - ; - global: - [ [ r = Prim.reference -> r - - (* This is used in quotations *) - | id = METAIDENT -> Ident (loc,id_of_string id) ] ] - ; - constr_pattern: - [ [ c = constr -> c ] ] - ; - lconstr_pattern: - [ [ c = lconstr -> c ] ] - ; - sort: - [ [ "Set" -> RProp Pos - | "Prop" -> RProp Null - | "Type" -> RType None ] ] - ; - lconstr: - [ [ c = operconstr LEVEL "200" -> c ] ] - ; - constr: - [ [ c = operconstr LEVEL "9" -> c ] ] - ; - operconstr: - [ "200" RIGHTA - [ c = binder_constr -> c ] - | "100" RIGHTA - [ c1 = operconstr; ":"; c2 = binder_constr -> CCast(loc,c1,DEFAULTcast,c2) - | c1 = operconstr; ":"; c2 = SELF -> CCast(loc,c1,DEFAULTcast,c2) ] - | "99" RIGHTA [ ] - | "90" RIGHTA - [ c1 = operconstr; "->"; c2 = binder_constr -> CArrow(loc,c1,c2) - | c1 = operconstr; "->"; c2 = SELF -> CArrow(loc,c1,c2)] - | "10" LEFTA - [ f=operconstr; args=LIST1 appl_arg -> CApp(loc,(None,f),args) - | "@"; f=global; args=LIST0 NEXT -> CAppExpl(loc,(None,f),args) ] - | "9" - [ ".."; c = operconstr LEVEL "0"; ".." -> - CAppExpl (loc,(None,Ident (loc,Topconstr.ldots_var)),[c]) ] - | "1" LEFTA - [ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" -> - CApp(loc,(Some (List.length args+1),CRef f),args@[c,None]) - | c=operconstr; ".("; "@"; f=global; - args=LIST0 (operconstr LEVEL "9"); ")" -> - CAppExpl(loc,(Some (List.length args+1),f),args@[c]) - | c=operconstr; "%"; key=IDENT -> CDelimiters (loc,key,c) ] - | "0" - [ c=atomic_constr -> c - | c=match_constr -> c - | "("; c = operconstr LEVEL "200"; ")" -> - (match c with - CNumeral(_,z) when Bigint.is_pos_or_zero z -> - CNotation(loc,"( _ )",[c]) - | _ -> c) ] ] - ; - binder_constr: - [ [ "forall"; bl = binder_list; ","; c = operconstr LEVEL "200" -> - mkCProdN loc bl c - | "fun"; bl = binder_list; "=>"; c = operconstr LEVEL "200" -> - mkCLambdaN loc bl c - | "let"; id=name; bl = LIST0 binder_let; ty = type_cstr; ":="; - c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" -> - let loc1 = loc_of_binder_let bl in - CLetIn(loc,id,mkCLambdaN loc1 bl (mk_cast(c1,ty)),c2) - | "let"; fx = single_fix; "in"; c = operconstr LEVEL "200" -> - let fixp = mk_single_fix fx in - let (li,id) = match fixp with - CFix(_,id,_) -> id - | CCoFix(_,id,_) -> id - | _ -> assert false in - CLetIn(loc,(li,Name id),fixp,c) - | "let"; lb = ["("; l=LIST0 name SEP ","; ")" -> l | "()" -> []]; - po = return_type; - ":="; c1 = operconstr LEVEL "200"; "in"; - c2 = operconstr LEVEL "200" -> - CLetTuple (loc,List.map snd lb,po,c1,c2) - | "if"; c=operconstr LEVEL "200"; po = return_type; - "then"; b1=operconstr LEVEL "200"; - "else"; b2=operconstr LEVEL "200" -> - CIf (loc, c, po, b1, b2) - | c=fix_constr -> c ] ] - ; - appl_arg: - [ [ id = lpar_id_coloneq; c=lconstr; ")" -> - (c,Some (loc,ExplByName id)) - | c=constr -> (c,None) ] ] - ; - atomic_constr: - [ [ g=global -> CRef g - | s=sort -> CSort(loc,s) - | n=INT -> CNumeral (loc, Bigint.of_string n) - | "_" -> CHole loc - | "?"; id=ident -> CPatVar(loc,(false,id)) ] ] - ; - fix_constr: - [ [ fx1=single_fix -> mk_single_fix fx1 - | (_,kw,dcl1)=single_fix; "with"; dcls=LIST1 fix_decl SEP "with"; - "for"; id=identref -> - mk_fix(loc,kw,id,dcl1::dcls) - ] ] - ; - single_fix: - [ [ kw=fix_kw; dcl=fix_decl -> (loc,kw,dcl) ] ] - ; - fix_kw: - [ [ "fix" -> true - | "cofix" -> false ] ] - ; - fix_decl: - [ [ id=identref; bl=LIST0 binder_let; ann=fixannot; ty=type_cstr; ":="; - c=operconstr LEVEL "200" -> (id,bl,ann,c,ty) ] ] - ; - fixannot: - [ [ "{"; IDENT "struct"; id=name; "}" -> Some id - | -> None ] ] - ; - match_constr: - [ [ "match"; ci=LIST1 case_item SEP ","; ty=OPT case_type; "with"; - br=branches; "end" -> CCases(loc,ty,ci,br) ] ] - ; - case_item: - [ [ c=operconstr LEVEL "100"; p=pred_pattern -> (c,p) ] ] - ; - pred_pattern: - [ [ ona = OPT ["as"; id=name -> snd id]; - ty = OPT ["in"; t=lconstr -> t] -> (ona,ty) ] ] - ; - case_type: - [ [ "return"; ty = operconstr LEVEL "100" -> ty ] ] - ; - return_type: - [ [ a = OPT [ na = OPT["as"; id=name -> snd id]; - ty = case_type -> (na,ty) ] -> - match a with - | None -> None, None - | Some (na,t) -> (na, Some t) - ] ] - ; - branches: - [ [ OPT"|"; br=LIST0 eqn SEP "|" -> br ] ] - ; - eqn: - [ [ pl = LIST1 pattern SEP ","; "=>"; rhs = lconstr -> (loc,pl,rhs) ] ] - ; - pattern: - [ "200" RIGHTA [ ] - | "100" LEFTA - [ p = pattern; "|"; pl = LIST1 pattern SEP "|" -> CPatOr (loc,p::pl) ] - | "99" RIGHTA [ ] - | "10" LEFTA - [ p = pattern; lp = LIST1 (pattern LEVEL "0") -> - (match p with - | CPatAtom (_, Some r) -> CPatCstr (loc, r, lp) - | _ -> Util.user_err_loc - (cases_pattern_loc p, "compound_pattern", - Pp.str "Constructor expected")) - | p = pattern; "as"; id = ident -> - CPatAlias (loc, p, id) - | c = pattern; "%"; key=IDENT -> - CPatDelimiters (loc,key,c) ] - | "0" - [ r = Prim.reference -> CPatAtom (loc,Some r) - | "_" -> CPatAtom (loc,None) - | "("; p = pattern LEVEL "200"; ")" -> - (match p with - CPatNumeral(_,z) when Bigint.is_pos_or_zero z -> - CPatNotation(loc,"( _ )",[p]) - | _ -> p) - | n = INT -> CPatNumeral (loc, Bigint.of_string n) ] ] - ; - binder_list: - [ [ idl=LIST1 name; bl=LIST0 binder_let -> - LocalRawAssum (idl,CHole loc)::bl - | idl=LIST1 name; ":"; c=lconstr -> - [LocalRawAssum (idl,c)] - | "("; idl=LIST1 name; ":"; c=lconstr; ")"; bl=LIST0 binder_let -> - LocalRawAssum (idl,c)::bl ] ] - ; - binder_let: - [ [ id=name -> - LocalRawAssum ([id],CHole loc) - | "("; id=name; idl=LIST1 name; ":"; c=lconstr; ")" -> - LocalRawAssum (id::idl,c) - | "("; id=name; ":"; c=lconstr; ")" -> - LocalRawAssum ([id],c) - | "("; id=name; ":="; c=lconstr; ")" -> - LocalRawDef (id,c) - | "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" -> - LocalRawDef (id,CCast (join_loc (constr_loc t) loc,c,DEFAULTcast,t)) - ] ] - ; - binder: - [ [ id=name -> ([id],CHole loc) - | "("; idl=LIST1 name; ":"; c=lconstr; ")" -> (idl,c) ] ] - ; - type_cstr: - [ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ] - ; - END;; diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4 index e73682d22..313886e9a 100644 --- a/parsing/g_ltac.ml4 +++ b/parsing/g_ltac.ml4 @@ -10,12 +10,10 @@ open Pp open Util -open Ast open Topconstr open Rawterm open Tacexpr open Vernacexpr -open Ast open Pcoq open Prim open Tactic @@ -39,162 +37,149 @@ let arg_of_expr = function (* Tactics grammar rules *) -if !Options.v7 then GEXTEND Gram - GLOBAL: tactic Vernac_.command tactic_arg - tactic_expr5 tactic_expr4 tactic_expr3 tactic_expr2; + GLOBAL: tactic Vernac_.command tactic_expr tactic_arg constr_may_eval; -(* - GLOBAL: tactic_atom tactic_atom0 input_fun; -*) + tactic_expr: + [ "5" LEFTA + [ ta0 = tactic_expr; ";"; ta1 = tactic_expr -> TacThen (ta0, ta1) + | ta = tactic_expr; ";"; + "["; lta = LIST0 OPT tactic_expr SEP "|"; "]" -> + let lta = List.map (function None -> TacId "" | Some t -> t) lta in + TacThens (ta, lta) ] + | "4" + [ ] + | "3" RIGHTA + [ IDENT "try"; ta = tactic_expr -> TacTry ta + | IDENT "do"; n = int_or_var; ta = tactic_expr -> TacDo (n,ta) + | IDENT "repeat"; ta = tactic_expr -> TacRepeat ta + | IDENT "progress"; ta = tactic_expr -> TacProgress ta + | IDENT "info"; tc = tactic_expr -> TacInfo tc +(*To do: put Abstract in Refiner*) + | IDENT "abstract"; tc = NEXT -> TacAbstract (tc,None) + | IDENT "abstract"; tc = NEXT; "using"; s = ident -> + TacAbstract (tc,Some s) ] +(*End of To do*) + | "2" RIGHTA + [ ta0 = tactic_expr; "||"; ta1 = tactic_expr -> TacOrelse (ta0,ta1) ] + | "1" RIGHTA + [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr -> + TacFun (it,body) + | "let"; IDENT "rec"; rcl = LIST1 rec_clause SEP "with"; "in"; + body = tactic_expr -> TacLetRecIn (rcl,body) + | "let"; llc = LIST1 let_clause SEP "with"; "in"; + u = tactic_expr -> TacLetIn (make_letin_clause loc llc,u) + | b = match_key; IDENT "goal"; "with"; mrl = match_context_list; "end" -> + TacMatchContext (b,false,mrl) + | b = match_key; IDENT "reverse"; IDENT "goal"; "with"; + mrl = match_context_list; "end" -> + TacMatchContext (b,true,mrl) + | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> + TacMatch (b,c,mrl) + | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> + TacFirst l + | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> + TacSolve l + | IDENT "idtac"; s = [ s = STRING -> s | -> ""] -> TacId s + | IDENT "fail"; n = [ n = int_or_var -> n | -> fail_default_value ]; + s = [ s = STRING -> s | -> ""] -> TacFail (n,s) + | IDENT "external"; com = STRING; req = STRING; la = LIST1 tactic_arg -> + TacArg (TacExternal (loc,com,req,la)) + | st = simple_tactic -> TacAtom (loc,st) + | a = may_eval_arg -> TacArg(a) + | IDENT "constr"; ":"; c = Constr.constr -> + TacArg(ConstrMayEval(ConstrTerm c)) + | IDENT "ipattern"; ":"; ipat = simple_intropattern -> + TacArg(IntroPattern ipat) + | r = reference; la = LIST1 tactic_arg -> + TacArg(TacCall (loc,r,la)) + | r = reference -> TacArg (Reference r) ] + | "0" + [ "("; a = tactic_expr; ")" -> a + | a = tactic_atom -> TacArg a ] ] + ; + (* Tactic arguments *) + tactic_arg: + [ [ IDENT "ltac"; ":"; a = tactic_expr LEVEL "0" -> arg_of_expr a + | IDENT "ipattern"; ":"; ipat = simple_intropattern -> IntroPattern ipat + | a = may_eval_arg -> a + | a = tactic_atom -> a + | c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ] + ; + may_eval_arg: + [ [ c = constr_eval -> ConstrMayEval c + | IDENT "fresh"; s = OPT STRING -> TacFreshId s ] ] + ; + constr_eval: + [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> + ConstrEval (rtc,c) + | IDENT "context"; id = identref; "["; c = Constr.lconstr; "]" -> + ConstrContext (id,c) + | IDENT "type"; IDENT "of"; c = Constr.constr -> + ConstrTypeOf c ] ] + ; + constr_may_eval: (* For extensions *) + [ [ c = constr_eval -> c + | c = Constr.constr -> ConstrTerm c ] ] + ; + tactic_atom: + [ [ id = METAIDENT -> MetaIdArg (loc,id) + | r = reference -> Reference r + | "()" -> TacVoid ] ] + ; + match_key: + [ [ "match" -> false ] ] + ; input_fun: - [ [ l = base_ident -> Some l - | "()" -> None ] ] + [ [ "_" -> None + | l = ident -> Some l ] ] ; let_clause: - [ [ id = identref; "="; te = tactic_letarg -> LETCLAUSE (id, None, te) - | id = base_ident; ":"; c = Constr.constr; ":="; "Proof" -> - LETTOPCLAUSE (id, c) - | id = identref; ":"; c = constrarg; ":="; te = tactic_letarg -> - LETCLAUSE (id, Some (TacArg(ConstrMayEval c)), te) - | id = base_ident; ":"; c = Constr.constr -> - LETTOPCLAUSE (id, c) ] ] + [ [ id = identref; ":="; te = tactic_expr -> + LETCLAUSE (id, None, arg_of_expr te) + | id = identref; args = LIST1 input_fun; ":="; te = tactic_expr -> + LETCLAUSE (id, None, arg_of_expr (TacFun(args,te))) ] ] ; rec_clause: - [ [ name = identref; it = LIST1 input_fun; "->"; body = tactic_expr -> + [ [ name = identref; it = LIST1 input_fun; ":="; body = tactic_expr -> (name,(it,body)) ] ] ; match_pattern: - [ [ 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 ] ] + [ [ IDENT "context"; oid = OPT Constr.ident; + "["; pc = Constr.lconstr_pattern; "]" -> + Subterm (oid, pc) + | pc = Constr.lconstr_pattern -> Term pc ] ] ; match_hyps: [ [ na = name; ":"; mp = match_pattern -> Hyp (na, mp) ] ] ; match_context_rule: - [ [ "["; largs = LIST0 match_hyps SEP ";"; "|-"; mp = match_pattern; "]"; - "->"; te = tactic_expr -> Pat (largs, mp, te) - | IDENT "_"; "->"; te = tactic_expr -> All te ] ] + [ [ largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; + "=>"; te = tactic_expr -> Pat (largs, mp, te) + | "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; + "]"; "=>"; te = tactic_expr -> Pat (largs, mp, te) + | "_"; "=>"; te = tactic_expr -> All te ] ] ; match_context_list: [ [ mrl = LIST1 match_context_rule SEP "|" -> mrl | "|"; mrl = LIST1 match_context_rule SEP "|" -> mrl ] ] ; match_rule: - [ [ "["; mp = match_pattern; "]"; "->"; te = tactic_expr -> Pat ([],mp,te) - | IDENT "_"; "->"; te = tactic_expr -> All te ] ] + [ [ mp = match_pattern; "=>"; te = tactic_expr -> Pat ([],mp,te) + | "_"; "=>"; te = tactic_expr -> All te ] ] ; match_list: [ [ mrl = LIST1 match_rule SEP "|" -> mrl | "|"; mrl = LIST1 match_rule SEP "|" -> mrl ] ] ; - tactic_expr: - [ [ 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 = int_or_var; 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 -> - TacFun (it,body) - | IDENT "Rec"; rc = rec_clause -> - warning "'Rec f ...' is obsolete; use 'Rec f ... In f' instead"; - TacLetRecIn ([rc],TacArg (Reference (Libnames.Ident (fst rc)))) - | IDENT "Rec"; rc = rec_clause; rcl = LIST0 rec_clause SEP "And"; - [IDENT "In" | "in"]; body = tactic_expr -> TacLetRecIn (rc::rcl,body) - | IDENT "Let"; llc = LIST1 let_clause SEP "And"; IDENT "In"; - u = tactic_expr -> TacLetIn (make_letin_clause loc llc,u) - - | IDENT "Match"; IDENT "Context"; IDENT "With"; mrl = match_context_list - -> TacMatchContext (false,false,mrl) - | IDENT "Match"; IDENT "Reverse"; IDENT "Context"; IDENT "With"; mrl = match_context_list - -> TacMatchContext (false,true,mrl) - | IDENT "Match"; c = constrarg; IDENT "With"; mrl = match_list -> - TacMatch (false,TacArg(ConstrMayEval c),mrl) -(*To do: put Abstract in Refiner*) - | IDENT "Abstract"; tc = tactic_expr -> TacAbstract (tc,None) - | IDENT "Abstract"; tc = tactic_expr; "using"; s = base_ident -> - TacAbstract (tc,Some s) -(*End of To do*) - | IDENT "First" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacFirst l - | IDENT "Solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacSolve l - | IDENT "Idtac" ; s = [ s = STRING -> s | -> ""] -> TacId s - | IDENT "Fail"; n = [ n = int_or_var -> n | -> fail_default_value ]; - s = [ s = STRING -> s | -> ""] -> TacFail (n,s) - | 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 = red_expr; "in"; c = Constr.constr -> - ConstrMayEval (ConstrEval (rtc,c)) - | IDENT "Inst"; id = identref; "["; c = Constr.constr; "]" -> - ConstrMayEval (ConstrContext (id,c)) - | IDENT "Check"; c = Constr.constr -> - ConstrMayEval (ConstrTypeOf c) - | IDENT "FreshId"; s = OPT STRING -> TacFreshId s - | IDENT "ipattern"; ":"; ipat = simple_intropattern -> IntroPattern ipat - | r = reference -> Reference r - | ta = tactic_arg0 -> ta ] ] - ; - tactic_arg1: - [ [ IDENT "Eval"; rtc = red_expr; "in"; c = Constr.constr -> - ConstrMayEval (ConstrEval (rtc,c)) - | IDENT "Inst"; id = identref; "["; c = Constr.constr; "]" -> - ConstrMayEval (ConstrContext (id,c)) - | IDENT "Check"; c = Constr.constr -> - ConstrMayEval (ConstrTypeOf c) - | IDENT "FreshId"; s = OPT STRING -> TacFreshId s - | IDENT "ipattern"; ":"; ipat = simple_intropattern -> IntroPattern ipat - | r = reference; la = LIST1 tactic_arg0 -> TacCall (loc,r,la) - | r = reference -> Reference r - | ta = tactic_arg0 -> ta ] ] - ; - tactic_arg0: - [ [ "("; a = tactic_expr; ")" -> arg_of_expr a - | "()" -> TacVoid - | r = reference -> Reference r - | n = integer -> Integer n - | id = METAIDENT -> MetaIdArg (loc,id) - | "?" -> ConstrMayEval (ConstrTerm (CHole loc)) - | "?"; n = natural -> ConstrMayEval (ConstrTerm (CPatVar (loc,(false,Pattern.patvar_of_int n)))) - | "'"; c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ] - ; (* Definitions for tactics *) +(* deftok: [ [ IDENT "Meta" | IDENT "Tactic" ] ] ; +*) tacdef_body: [ [ name = identref; it=LIST1 input_fun; ":="; body = tactic_expr -> (name, TacFun (it, body)) @@ -205,10 +190,8 @@ GEXTEND Gram [ [ tac = tactic_expr -> tac ] ] ; Vernac_.command: - [ [ deftok; "Definition"; b = tacdef_body -> - VernacDeclareTacticDefinition (false, [b]) - | IDENT "Recursive"; deftok; "Definition"; - l = LIST1 tacdef_body SEP "And" -> + [ [ IDENT "Ltac"; + l = LIST1 tacdef_body SEP "with" -> VernacDeclareTacticDefinition (true, l) ] ] ; END diff --git a/parsing/g_ltacnew.ml4 b/parsing/g_ltacnew.ml4 deleted file mode 100644 index 313886e9a..000000000 --- a/parsing/g_ltacnew.ml4 +++ /dev/null @@ -1,197 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pp -open Util -open Topconstr -open Rawterm -open Tacexpr -open Vernacexpr -open Pcoq -open Prim -open Tactic - -type let_clause_kind = - | LETTOPCLAUSE of Names.identifier * constr_expr - | LETCLAUSE of - (Names.identifier Util.located * raw_tactic_expr option * raw_tactic_arg) - -let fail_default_value = Genarg.ArgArg 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 loc = List.map (out_letin_clause loc) - -let arg_of_expr = function - TacArg a -> a - | e -> Tacexp (e:raw_tactic_expr) - -(* Tactics grammar rules *) - -GEXTEND Gram - GLOBAL: tactic Vernac_.command tactic_expr tactic_arg constr_may_eval; - - tactic_expr: - [ "5" LEFTA - [ ta0 = tactic_expr; ";"; ta1 = tactic_expr -> TacThen (ta0, ta1) - | ta = tactic_expr; ";"; - "["; lta = LIST0 OPT tactic_expr SEP "|"; "]" -> - let lta = List.map (function None -> TacId "" | Some t -> t) lta in - TacThens (ta, lta) ] - | "4" - [ ] - | "3" RIGHTA - [ IDENT "try"; ta = tactic_expr -> TacTry ta - | IDENT "do"; n = int_or_var; ta = tactic_expr -> TacDo (n,ta) - | IDENT "repeat"; ta = tactic_expr -> TacRepeat ta - | IDENT "progress"; ta = tactic_expr -> TacProgress ta - | IDENT "info"; tc = tactic_expr -> TacInfo tc -(*To do: put Abstract in Refiner*) - | IDENT "abstract"; tc = NEXT -> TacAbstract (tc,None) - | IDENT "abstract"; tc = NEXT; "using"; s = ident -> - TacAbstract (tc,Some s) ] -(*End of To do*) - | "2" RIGHTA - [ ta0 = tactic_expr; "||"; ta1 = tactic_expr -> TacOrelse (ta0,ta1) ] - | "1" RIGHTA - [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr -> - TacFun (it,body) - | "let"; IDENT "rec"; rcl = LIST1 rec_clause SEP "with"; "in"; - body = tactic_expr -> TacLetRecIn (rcl,body) - | "let"; llc = LIST1 let_clause SEP "with"; "in"; - u = tactic_expr -> TacLetIn (make_letin_clause loc llc,u) - | b = match_key; IDENT "goal"; "with"; mrl = match_context_list; "end" -> - TacMatchContext (b,false,mrl) - | b = match_key; IDENT "reverse"; IDENT "goal"; "with"; - mrl = match_context_list; "end" -> - TacMatchContext (b,true,mrl) - | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> - TacMatch (b,c,mrl) - | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacFirst l - | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - TacSolve l - | IDENT "idtac"; s = [ s = STRING -> s | -> ""] -> TacId s - | IDENT "fail"; n = [ n = int_or_var -> n | -> fail_default_value ]; - s = [ s = STRING -> s | -> ""] -> TacFail (n,s) - | IDENT "external"; com = STRING; req = STRING; la = LIST1 tactic_arg -> - TacArg (TacExternal (loc,com,req,la)) - | st = simple_tactic -> TacAtom (loc,st) - | a = may_eval_arg -> TacArg(a) - | IDENT "constr"; ":"; c = Constr.constr -> - TacArg(ConstrMayEval(ConstrTerm c)) - | IDENT "ipattern"; ":"; ipat = simple_intropattern -> - TacArg(IntroPattern ipat) - | r = reference; la = LIST1 tactic_arg -> - TacArg(TacCall (loc,r,la)) - | r = reference -> TacArg (Reference r) ] - | "0" - [ "("; a = tactic_expr; ")" -> a - | a = tactic_atom -> TacArg a ] ] - ; - (* Tactic arguments *) - tactic_arg: - [ [ IDENT "ltac"; ":"; a = tactic_expr LEVEL "0" -> arg_of_expr a - | IDENT "ipattern"; ":"; ipat = simple_intropattern -> IntroPattern ipat - | a = may_eval_arg -> a - | a = tactic_atom -> a - | c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ] - ; - may_eval_arg: - [ [ c = constr_eval -> ConstrMayEval c - | IDENT "fresh"; s = OPT STRING -> TacFreshId s ] ] - ; - constr_eval: - [ [ IDENT "eval"; rtc = red_expr; "in"; c = Constr.constr -> - ConstrEval (rtc,c) - | IDENT "context"; id = identref; "["; c = Constr.lconstr; "]" -> - ConstrContext (id,c) - | IDENT "type"; IDENT "of"; c = Constr.constr -> - ConstrTypeOf c ] ] - ; - constr_may_eval: (* For extensions *) - [ [ c = constr_eval -> c - | c = Constr.constr -> ConstrTerm c ] ] - ; - tactic_atom: - [ [ id = METAIDENT -> MetaIdArg (loc,id) - | r = reference -> Reference r - | "()" -> TacVoid ] ] - ; - match_key: - [ [ "match" -> false ] ] - ; - input_fun: - [ [ "_" -> None - | l = ident -> Some l ] ] - ; - let_clause: - [ [ id = identref; ":="; te = tactic_expr -> - LETCLAUSE (id, None, arg_of_expr te) - | id = identref; args = LIST1 input_fun; ":="; te = tactic_expr -> - LETCLAUSE (id, None, arg_of_expr (TacFun(args,te))) ] ] - ; - rec_clause: - [ [ name = identref; it = LIST1 input_fun; ":="; body = tactic_expr -> - (name,(it,body)) ] ] - ; - match_pattern: - [ [ IDENT "context"; oid = OPT Constr.ident; - "["; pc = Constr.lconstr_pattern; "]" -> - Subterm (oid, pc) - | pc = Constr.lconstr_pattern -> Term pc ] ] - ; - match_hyps: - [ [ na = name; ":"; mp = match_pattern -> Hyp (na, mp) ] ] - ; - match_context_rule: - [ [ largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; - "=>"; te = tactic_expr -> Pat (largs, mp, te) - | "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; - "]"; "=>"; te = tactic_expr -> Pat (largs, mp, te) - | "_"; "=>"; te = tactic_expr -> All te ] ] - ; - match_context_list: - [ [ mrl = LIST1 match_context_rule SEP "|" -> mrl - | "|"; mrl = LIST1 match_context_rule SEP "|" -> mrl ] ] - ; - match_rule: - [ [ mp = match_pattern; "=>"; te = tactic_expr -> Pat ([],mp,te) - | "_"; "=>"; te = tactic_expr -> All te ] ] - ; - match_list: - [ [ mrl = LIST1 match_rule SEP "|" -> mrl - | "|"; mrl = LIST1 match_rule SEP "|" -> mrl ] ] - ; - - (* Definitions for tactics *) -(* - deftok: - [ [ IDENT "Meta" - | IDENT "Tactic" ] ] - ; -*) - tacdef_body: - [ [ name = identref; it=LIST1 input_fun; ":="; body = tactic_expr -> - (name, TacFun (it, body)) - | name = identref; ":="; body = tactic_expr -> - (name, body) ] ] - ; - tactic: - [ [ tac = tactic_expr -> tac ] ] - ; - Vernac_.command: - [ [ IDENT "Ltac"; - l = LIST1 tacdef_body SEP "with" -> - VernacDeclareTacticDefinition (true, l) ] ] - ; - END diff --git a/parsing/g_module.ml4 b/parsing/g_module.ml4 deleted file mode 100644 index e1937b5f9..000000000 --- a/parsing/g_module.ml4 +++ /dev/null @@ -1,47 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pp -open Ast -open Pcoq -open Prim -open Module -open Util -open Topconstr - -(* Grammar rules for module expressions and types *) - -if !Options.v7 then -GEXTEND Gram - GLOBAL: module_expr module_type; - - module_expr: - [ [ qid = qualid -> CMEident qid - | me1 = module_expr; me2 = module_expr -> CMEapply (me1,me2) - | "("; me = module_expr; ")" -> me -(* ... *) - ] ] - ; - - with_declaration: - [ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.constr -> - CWith_Definition (fqid,c) - | IDENT "Module"; fqid = fullyqualid; ":="; qid = qualid -> - CWith_Module (fqid,qid) - ] ] - ; - - module_type: - [ [ qid = qualid -> CMTEident qid -(* ... *) - | mty = module_type; "with"; decl = with_declaration -> - CMTEwith (mty,decl) ] ] - ; -END diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index bbf00a489..0a3a3c92b 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -8,107 +8,80 @@ (*i $Id$ i*) -open Coqast open Pcoq open Names open Libnames open Topconstr -open Prim -let _ = reset_all_grammars() +let prim_kw = ["{"; "}"; "["; "]"; "("; ")"; "'"] +let _ = List.iter (fun s -> Lexer.add_token("",s)) prim_kw +open Prim open Nametab -let local_id_of_string = id_of_string -let local_make_dirpath = make_dirpath -let local_make_qualid l id' = make_qualid (local_make_dirpath l) id' -let local_make_short_qualid id = make_short_qualid id -let local_make_posint = int_of_string -let local_make_negint n = - int_of_string n -let local_make_path l a = encode_kn (local_make_dirpath l) 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] -GEXTEND Gram - GLOBAL: bigint ident string preident ast - astlist qualid reference dirpath identref name base_ident var; +let local_make_qualid l id = make_qualid (make_dirpath l) id - metaident: - [ [ s = METAIDENT -> Nmeta (loc,s) ] ] - ; - base_ident: - [ [ s = IDENT -> local_id_of_string s ] ] +GEXTEND Gram + GLOBAL: + bigint natural integer identref name ident var preident + fullyqualid qualid reference + ne_string; + preident: + [ [ s = IDENT -> s ] ] ; ident: - [ [ id = base_ident -> id ] ] + [ [ s = IDENT -> id_of_string s ] ] ; - bigint: - [ [ i = INT -> Bigint.of_string i - | "-"; i = INT -> Bigint.neg (Bigint.of_string i) ] ] + var: (* as identref, but interpret as a term identifier in ltac *) + [ [ id = ident -> (loc,id) ] ] ; - field: - [ [ s = FIELD -> local_id_of_string s ] ] + identref: + [ [ id = ident -> (loc,id) ] ] ; - dirpath: - [ [ id = base_ident; l = LIST0 field -> - local_make_dirpath (local_append l id) ] ] + field: + [ [ s = FIELD -> id_of_string s ] ] ; fields: - [ [ id = field; (l,id') = fields -> (local_append l id,id') + [ [ id = field; (l,id') = fields -> (l@[id],id') | id = field -> ([],id) ] ] ; + fullyqualid: + [ [ id = ident; (l,id')=fields -> loc,id::List.rev (id'::l) + | id = ident -> loc,[id] + ] ] + ; basequalid: - [ [ id = base_ident; (l,id')=fields -> local_make_qualid (local_append l id) id' - | id = base_ident -> local_make_short_qualid id + [ [ id = ident; (l,id')=fields -> local_make_qualid (l@[id]) id' + | id = ident -> make_short_qualid id ] ] ; - qualid: - [ [ qid = basequalid -> loc, qid ] ] + name: + [ [ IDENT "_" -> (loc, Anonymous) + | id = ident -> (loc, Name id) ] ] ; reference: - [ [ id = base_ident; (l,id') = fields -> - Qualid (loc, local_make_qualid (local_append l id) id') - | id = base_ident -> Ident (loc,id) + [ [ id = ident; (l,id') = fields -> + Qualid (loc, local_make_qualid (l@[id]) id') + | id = ident -> Ident (loc,id) ] ] ; - string: - [ [ s = STRING -> s ] ] + qualid: + [ [ qid = basequalid -> loc, qid ] ] + ; + ne_string: + [ [ s = STRING -> + if s="" then Util.user_err_loc(loc,"",Pp.str"Empty string"); s + ] ] ; - astpath: - [ [ id = base_ident; (l,a) = fields -> - Path(loc, local_make_path (local_append l id) a) - | id = base_ident -> Nvar(loc, id) - ] ] + integer: + [ [ i = INT -> int_of_string i + | "-"; i = INT -> - int_of_string i ] ] ; - (* ast *) - ast: - [ [ id = metaident -> id - | p = astpath -> p - | s = INT -> Num(loc, local_make_posint s) - | s = STRING -> Str(loc, s) - | "{"; s = METAIDENT; "}" -> Id(loc,s) - | "("; nname = IDENT; l = LIST0 ast; ")" -> Node(loc,nname,l) - | "("; METAIDENT "$LIST"; id = metaident; ")" -> Node(loc,"$LIST",[id]) - | "("; METAIDENT "$STR"; id = metaident; ")" -> Node(loc,"$STR",[id]) - | "("; METAIDENT "$VAR"; id = metaident; ")" -> Node(loc,"$VAR",[id]) - | "("; METAIDENT "$ID"; id = metaident; ")" -> Node(loc,"$ID",[id]) - | "("; METAIDENT "$ABSTRACT"; l = LIST0 ast;")"->Node(loc,"$ABSTRACT",l) - | "("; 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 -> local_make_binding loc a b - -(* - | "["; ido = astidoption; "]"; b = ast -> Slam(loc,ido,b) - | "["; id = METAIDENT; "]"; b = ast -> Smetalam(loc,id,b) -*) - | "'"; a = ast -> Node(loc,"$QUOTE",[a]) ] ] + natural: + [ [ i = INT -> int_of_string i ] ] ; - astlist: - [ [ l = LIST0 ast -> l ] ] + bigint: (* Negative numbers are dealt with specially *) + [ [ i = INT -> (Bigint.of_string i) ] ] ; END diff --git a/parsing/g_primnew.ml4 b/parsing/g_primnew.ml4 deleted file mode 100644 index 0a3a3c92b..000000000 --- a/parsing/g_primnew.ml4 +++ /dev/null @@ -1,87 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(*i $Id$ i*) - -open Pcoq -open Names -open Libnames -open Topconstr - -let prim_kw = ["{"; "}"; "["; "]"; "("; ")"; "'"] -let _ = List.iter (fun s -> Lexer.add_token("",s)) prim_kw - -open Prim -open Nametab - -let local_make_qualid l id = make_qualid (make_dirpath l) id - -GEXTEND Gram - GLOBAL: - bigint natural integer identref name ident var preident - fullyqualid qualid reference - ne_string; - preident: - [ [ s = IDENT -> s ] ] - ; - ident: - [ [ s = IDENT -> id_of_string s ] ] - ; - var: (* as identref, but interpret as a term identifier in ltac *) - [ [ id = ident -> (loc,id) ] ] - ; - identref: - [ [ id = ident -> (loc,id) ] ] - ; - field: - [ [ s = FIELD -> id_of_string s ] ] - ; - fields: - [ [ id = field; (l,id') = fields -> (l@[id],id') - | id = field -> ([],id) - ] ] - ; - fullyqualid: - [ [ id = ident; (l,id')=fields -> loc,id::List.rev (id'::l) - | id = ident -> loc,[id] - ] ] - ; - basequalid: - [ [ id = ident; (l,id')=fields -> local_make_qualid (l@[id]) id' - | id = ident -> make_short_qualid id - ] ] - ; - name: - [ [ IDENT "_" -> (loc, Anonymous) - | id = ident -> (loc, Name id) ] ] - ; - reference: - [ [ id = ident; (l,id') = fields -> - Qualid (loc, local_make_qualid (l@[id]) id') - | id = ident -> Ident (loc,id) - ] ] - ; - qualid: - [ [ qid = basequalid -> loc, qid ] ] - ; - ne_string: - [ [ s = STRING -> - if s="" then Util.user_err_loc(loc,"",Pp.str"Empty string"); s - ] ] - ; - integer: - [ [ i = INT -> int_of_string i - | "-"; i = INT -> - int_of_string i ] ] - ; - natural: - [ [ i = INT -> int_of_string i ] ] - ; - bigint: (* Negative numbers are dealt with specially *) - [ [ i = INT -> (Bigint.of_string i) ] ] - ; -END diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index fb69e3cc0..54d68bbea 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -18,10 +18,9 @@ open Vernacexpr open Prim open Constr -let thm_token = Gram.Entry.create "vernac:thm_token" +let thm_token = G_vernac.thm_token (* Proof commands *) -if !Options.v7 then GEXTEND Gram GLOBAL: command; @@ -35,27 +34,29 @@ GEXTEND Gram | ":"; l = LIST1 IDENT -> l ] ] ; command: - [ [ IDENT "Goal"; c = Constr.constr -> VernacGoal c - | "Proof" -> VernacProof (Tacexpr.TacId "") - | "Proof"; "with"; ta = tactic -> VernacProof ta + [ [ IDENT "Goal"; c = Constr.lconstr -> VernacGoal c + | IDENT "Proof" -> VernacProof (Tacexpr.TacId "") + | IDENT "Proof"; "with"; ta = tactic -> VernacProof ta | IDENT "Abort" -> VernacAbort None | IDENT "Abort"; IDENT "All" -> VernacAbortAll | IDENT "Abort"; id = identref -> VernacAbort (Some id) + | IDENT "Existential"; n = natural; c = constr_body -> + VernacSolveExistential (n,c) | IDENT "Admitted" -> VernacEndProof Admitted - | "Qed" -> VernacEndProof (Proved (true,None)) + | IDENT "Qed" -> VernacEndProof (Proved (true,None)) | IDENT "Save" -> VernacEndProof (Proved (true,None)) - | IDENT "Defined" -> VernacEndProof (Proved (false,None)) - | IDENT "Defined"; id=identref -> - VernacEndProof (Proved (false,Some (id,None))) | IDENT "Save"; tok = thm_token; id = identref -> VernacEndProof (Proved (true,Some (id,Some tok))) | IDENT "Save"; id = identref -> VernacEndProof (Proved (true,Some (id,None))) + | IDENT "Defined" -> VernacEndProof (Proved (false,None)) + | IDENT "Defined"; id=identref -> + VernacEndProof (Proved (false,Some (id,None))) | IDENT "Suspend" -> VernacSuspend | IDENT "Resume" -> VernacResume None | IDENT "Resume"; id = identref -> VernacResume (Some id) | IDENT "Restart" -> VernacRestart - | "Proof"; c = Constr.constr -> VernacExactProof c + | IDENT "Proof"; c = Constr.lconstr -> VernacExactProof c | IDENT "Undo" -> VernacUndo 1 | IDENT "Undo"; n = natural -> VernacUndo n | IDENT "Focus" -> VernacFocus None @@ -63,20 +64,20 @@ GEXTEND Gram | 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 "Implicit"; IDENT "Arguments"; n = OPT natural -> + VernacShow (ShowGoalImplicitly n) | 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 "Proof" -> VernacShow ShowProof | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false) | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true) - | IDENT "Explain"; "Proof"; l = LIST0 integer -> + | IDENT "Show"; IDENT "Match"; id = identref -> VernacShow (ShowMatch id) + | IDENT "Explain"; IDENT "Proof"; l = LIST0 integer -> VernacShow (ExplainProof l) - | IDENT "Explain"; "Proof"; IDENT "Tree"; l = LIST0 integer -> + | IDENT "Explain"; IDENT "Proof"; IDENT "Tree"; l = LIST0 integer -> VernacShow (ExplainTree l) | IDENT "Go"; n = natural -> VernacGo (GoTo n) | IDENT "Go"; IDENT "top" -> VernacGo GoTop @@ -84,22 +85,9 @@ GEXTEND Gram | IDENT "Go"; IDENT "next" -> VernacGo GoNext | IDENT "Guarded" -> VernacCheckGuard (* Hints for Auto and EAuto *) - - | IDENT "HintDestruct"; - local = locality; - dloc = destruct_location; - id = base_ident; - hyptyp = Constr.constr_pattern; - pri = natural; - "["; tac = tactic; "]" -> - VernacHints(local,[],HintsDestruct (id,pri,dloc,hyptyp,tac)) - - | IDENT "Hint"; local = locality; hintname = base_ident; - dbnames = opt_hintbases; ":="; h = hint - -> VernacHints (local,dbnames, h hintname) - - | IDENT "Hints"; local = locality; - (dbnames,h) = hints -> VernacHints (local,dbnames, h) + | IDENT "Hint"; local = locality; h = hint; + dbnames = opt_hintbases -> + VernacHints (local,dbnames, h) (*This entry is not commented, only for debug*) @@ -112,24 +100,27 @@ GEXTEND Gram [ [ IDENT "Local" -> true | -> false ] ] ; 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 = global -> fun name -> HintsUnfold [Some name,qid] - | IDENT "Constructors"; c = global -> fun n -> - HintsConstructors (Some n,[c]) - | IDENT "Extern"; n = natural; c = Constr.constr ; tac = tactic -> - fun name -> HintsExtern (Some name,n,c,tac) ] ] - ; - hints: - [ [ IDENT "Resolve"; l = LIST1 global; dbnames = opt_hintbases -> - (dbnames, - HintsResolve - (List.map (fun qid -> (None, CAppExpl(loc,(None,qid),[]))) l)) - | IDENT "Immediate"; l = LIST1 global; dbnames = opt_hintbases -> - (dbnames, - HintsImmediate - (List.map (fun qid-> (None, CAppExpl (loc,(None,qid),[]))) l)) - | IDENT "Unfold"; l = LIST1 global; dbnames = opt_hintbases -> - (dbnames, HintsUnfold (List.map (fun qid -> (None,qid)) l)) ] ] + [ [ IDENT "Resolve"; lc = LIST1 Constr.constr -> + HintsResolve (List.map (fun c -> (None, c)) lc) + | IDENT "Immediate"; lc = LIST1 Constr.constr -> + HintsImmediate (List.map (fun c -> (None,c)) lc) + | IDENT "Unfold"; lqid = LIST1 global -> + HintsUnfold (List.map (fun g -> (None,g)) lqid) + | IDENT "Constructors"; lc = LIST1 global -> + HintsConstructors (None,lc) + | IDENT "Extern"; n = natural; c = Constr.constr_pattern ; "=>"; + tac = tactic -> + HintsExtern (None,n,c,tac) + | IDENT"Destruct"; + id = ident; ":="; + pri = natural; + dloc = destruct_location; + hyptyp = Constr.constr_pattern; + "=>"; tac = tactic -> + HintsDestruct(id,pri,dloc,hyptyp,tac) ] ] ; - END + constr_body: + [ [ ":="; c = lconstr -> c + | ":"; t = lconstr; ":="; c = lconstr -> CCast(loc,c,Term.DEFAULTcast,t) ] ] + ; +END diff --git a/parsing/g_proofsnew.ml4 b/parsing/g_proofsnew.ml4 deleted file mode 100644 index 5cf5152a3..000000000 --- a/parsing/g_proofsnew.ml4 +++ /dev/null @@ -1,126 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pcoq -open Pp -open Tactic -open Util -open Vernac_ -open Topconstr -open Vernacexpr -open Prim -open Constr - -let thm_token = G_vernacnew.thm_token - -(* Proof commands *) -GEXTEND Gram - GLOBAL: command; - - destruct_location : - [ [ IDENT "Conclusion" -> Tacexpr.ConclLocation () - | discard = [ IDENT "Discardable" -> true | -> false ]; "Hypothesis" - -> Tacexpr.HypLocation discard ] ] - ; - opt_hintbases: - [ [ -> [] - | ":"; l = LIST1 IDENT -> l ] ] - ; - command: - [ [ IDENT "Goal"; c = Constr.lconstr -> VernacGoal c - | IDENT "Proof" -> VernacProof (Tacexpr.TacId "") - | IDENT "Proof"; "with"; ta = tactic -> VernacProof ta - | IDENT "Abort" -> VernacAbort None - | IDENT "Abort"; IDENT "All" -> VernacAbortAll - | IDENT "Abort"; id = identref -> VernacAbort (Some id) - | IDENT "Existential"; n = natural; c = constr_body -> - VernacSolveExistential (n,c) - | IDENT "Admitted" -> VernacEndProof Admitted - | IDENT "Qed" -> VernacEndProof (Proved (true,None)) - | IDENT "Save" -> VernacEndProof (Proved (true,None)) - | IDENT "Save"; tok = thm_token; id = identref -> - VernacEndProof (Proved (true,Some (id,Some tok))) - | IDENT "Save"; id = identref -> - VernacEndProof (Proved (true,Some (id,None))) - | IDENT "Defined" -> VernacEndProof (Proved (false,None)) - | IDENT "Defined"; id=identref -> - VernacEndProof (Proved (false,Some (id,None))) - | IDENT "Suspend" -> VernacSuspend - | IDENT "Resume" -> VernacResume None - | IDENT "Resume"; id = identref -> VernacResume (Some id) - | IDENT "Restart" -> VernacRestart - | IDENT "Proof"; c = Constr.lconstr -> 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 "Implicit"; IDENT "Arguments"; n = OPT natural -> - VernacShow (ShowGoalImplicitly n) - | 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"; IDENT "Proof" -> VernacShow ShowProof - | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false) - | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true) - | IDENT "Show"; IDENT "Match"; id = identref -> VernacShow (ShowMatch id) - | IDENT "Explain"; IDENT "Proof"; l = LIST0 integer -> - VernacShow (ExplainProof l) - | IDENT "Explain"; IDENT "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"; local = locality; h = hint; - dbnames = opt_hintbases -> - VernacHints (local,dbnames, h) - - -(*This entry is not commented, only for debug*) - | IDENT "PrintConstr"; c = Constr.constr -> - VernacExtend ("PrintConstr", - [Genarg.in_gen Genarg.rawwit_constr c]) - ] ]; - - locality: - [ [ IDENT "Local" -> true | -> false ] ] - ; - hint: - [ [ IDENT "Resolve"; lc = LIST1 Constr.constr -> - HintsResolve (List.map (fun c -> (None, c)) lc) - | IDENT "Immediate"; lc = LIST1 Constr.constr -> - HintsImmediate (List.map (fun c -> (None,c)) lc) - | IDENT "Unfold"; lqid = LIST1 global -> - HintsUnfold (List.map (fun g -> (None,g)) lqid) - | IDENT "Constructors"; lc = LIST1 global -> - HintsConstructors (None,lc) - | IDENT "Extern"; n = natural; c = Constr.constr_pattern ; "=>"; - tac = tactic -> - HintsExtern (None,n,c,tac) - | IDENT"Destruct"; - id = ident; ":="; - pri = natural; - dloc = destruct_location; - hyptyp = Constr.constr_pattern; - "=>"; tac = tactic -> - HintsDestruct(id,pri,dloc,hyptyp,tac) ] ] - ; - constr_body: - [ [ ":="; c = lconstr -> c - | ":"; t = lconstr; ":="; c = lconstr -> CCast(loc,c,Term.DEFAULTcast,t) ] ] - ; -END diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4 index 56c9cd67e..180754668 100644 --- a/parsing/g_tactic.ml4 +++ b/parsing/g_tactic.ml4 @@ -9,72 +9,102 @@ (* $Id$ *) open Pp -open Ast open Pcoq open Util open Tacexpr open Rawterm open Genarg + +let compute = Cbv all_flags + +let tactic_kw = [ "->"; "<-" ] +let _ = List.iter (fun s -> Lexer.add_token("",s)) tactic_kw + +(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *) +(* admissible notation "(x t)" *) +let lpar_id_coloneq = + Gram.Entry.of_parser "lpar_id_coloneq" + (fun strm -> + match Stream.npeek 1 strm with + | [("","(")] -> + (match Stream.npeek 2 strm with + | [_; ("IDENT",s)] -> + (match Stream.npeek 3 strm with + | [_; _; ("", ":=")] -> + Stream.junk strm; Stream.junk strm; Stream.junk strm; + Names.id_of_string s + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + +(* idem for (x:=t) and (1:=t) *) +let test_lpar_idnum_coloneq = + Gram.Entry.of_parser "test_lpar_idnum_coloneq" + (fun strm -> + match Stream.npeek 1 strm with + | [("","(")] -> + (match Stream.npeek 2 strm with + | [_; (("IDENT"|"INT"),_)] -> + (match Stream.npeek 3 strm with + | [_; _; ("", ":=")] -> () + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + +(* idem for (x:t) *) +let lpar_id_colon = + Gram.Entry.of_parser "lpar_id_colon" + (fun strm -> + match Stream.npeek 1 strm with + | [("","(")] -> + (match Stream.npeek 2 strm with + | [_; ("IDENT",id)] -> + (match Stream.npeek 3 strm with + | [_; _; ("", ":")] -> + Stream.junk strm; Stream.junk strm; Stream.junk strm; + Names.id_of_string id + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + open Constr open Prim open Tactic -let tactic_kw = - [ "using"; "Orelse"; "Proof"; "Qed"; "And"; "()"; "|-" ] -let _ = - if !Options.v7 then - List.iter (fun s -> Lexer.add_token ("",s)) tactic_kw +let mk_fix_tac (loc,id,bl,ann,ty) = + let n = + match bl,ann with + [([_],_)], None -> 1 + | _, Some x -> + let ids = List.map snd (List.flatten (List.map fst bl)) in + (try list_index (snd x) ids + with Not_found -> error "no such fix variable") + | _ -> error "cannot guess decreasing argument of fix" in + (id,n,Topconstr.CProdN(loc,bl,ty)) -(* Functions overloaded by quotifier *) +let mk_cofix_tac (loc,id,bl,ann,ty) = + let _ = option_app (fun (aloc,_) -> + Util.user_err_loc + (aloc,"Constr:mk_cofix_tac", + Pp.str"Annotation forbidden in cofix expression")) ann in + (id,Topconstr.CProdN(loc,bl,ty)) +(* Functions overloaded by quotifier *) let induction_arg_of_constr c = - try ElimOnIdent (Topconstr.constr_loc c,snd (coerce_to_id c)) + try ElimOnIdent (Topconstr.constr_loc c,snd(coerce_to_id c)) with _ -> ElimOnConstr c -let local_compute = [FBeta;FIota;FDeltaBut [];FZeta] - -let error_oldelim _ = error "OldElim no longer supported" - -let join_to_constr loc c2 = (fst loc), snd (Topconstr.constr_loc c2) - (* Auxiliary grammar rules *) -if !Options.v7 then GEXTEND Gram - GLOBAL: simple_tactic constrarg bindings constr_with_bindings - quantified_hypothesis red_expr int_or_var open_constr casted_open_constr + GLOBAL: simple_tactic constr_with_bindings quantified_hypothesis + bindings red_expr int_or_var open_constr casted_open_constr simple_intropattern; int_or_var: [ [ n = integer -> Genarg.ArgArg n | id = identref -> Genarg.ArgVar id ] ] ; - autoarg_depth: - [ [ n = OPT natural -> n ] ] - ; - autoarg_adding: - [ [ IDENT "Adding" ; "["; l = LIST1 global; "]" -> l | -> [] ] ] - ; - autoarg_destructing: - [ [ IDENT "Destructing" -> true | -> false ] ] - ; - autoarg_usingTDB: - [ [ "Using"; "TDB" -> true | -> false ] ] - ; - autoargs: - [ [ a0 = autoarg_depth; l = autoarg_adding; - a2 = autoarg_destructing; a3 = autoarg_usingTDB -> (a0,l,a2,a3) ] ] - ; - (* Either an hypothesis or a ltac ref (variable or pattern patvar) *) - id_or_ltac_ref: - [ [ id = base_ident -> AI (loc,id) - | "?"; n = natural -> AI (loc,Pattern.patvar_of_int n) ] ] - ; - (* Either a global ref or a ltac ref (variable or pattern patvar) *) - global_or_ltac_ref: - [ [ qid = global -> qid - | "?"; n = natural -> Libnames.Ident (loc,Pattern.patvar_of_int n) ] ] - ; (* An identifier or a quotation meta-variable *) id_or_meta: [ [ id = identref -> AI id @@ -88,14 +118,6 @@ GEXTEND Gram | id = METAIDENT -> MetaId (loc,id) ] ] ; - constrarg: - [ [ IDENT "Inst"; id = identref; "["; c = constr; "]" -> - ConstrContext (id, c) - | IDENT "Eval"; rtc = Tactic.red_expr; "in"; c = constr -> - ConstrEval (rtc,c) - | IDENT "Check"; c = constr -> ConstrTypeOf c - | c = constr -> ConstrTerm c ] ] - ; open_constr: [ [ c = constr -> ((),c) ] ] ; @@ -108,17 +130,24 @@ GEXTEND Gram ] ] ; quantified_hypothesis: - [ [ id = base_ident -> NamedHyp id + [ [ id = ident -> NamedHyp id | n = natural -> AnonHyp n ] ] ; conversion: - [ [ nl = LIST1 integer; c1 = constr; "with"; c2 = constr -> - (Some (nl,c1), c2) - | c1 = constr; "with"; c2 = constr -> (Some ([],c1), c2) - | c = constr -> (None, c) ] ] + [ [ c = constr -> (None, c) + | c1 = constr; "with"; c2 = constr -> (Some ([],c1), c2) + | c1 = constr; "at"; nl = LIST1 integer; "with"; c2 = constr -> + (Some (nl,c1), c2) ] ] + ; + occurrences: + [ [ "at"; nl = LIST1 integer -> nl + | -> [] ] ] ; pattern_occ: - [ [ nl = LIST0 integer; c = constr -> (nl,c) ] ] + [ [ c = constr; nl = occurrences -> (nl,c) ] ] + ; + unfold_occ: + [ [ c = global; nl = occurrences -> (nl,c) ] ] ; intropatterns: [ [ l = LIST0 simple_intropattern -> l ]] @@ -126,22 +155,18 @@ GEXTEND Gram simple_intropattern: [ [ "["; tc = LIST1 intropatterns SEP "|" ; "]" -> IntroOrAndPattern tc | "("; tc = LIST1 simple_intropattern SEP "," ; ")" -> IntroOrAndPattern [tc] - | IDENT "_" -> IntroWildcard - | id = base_ident -> IntroIdentifier id + | "_" -> IntroWildcard + | id = ident -> IntroIdentifier id ] ] ; simple_binding: - [ [ id = base_ident; ":="; c = constr -> (loc, NamedHyp id, c) - | n = natural; ":="; c = constr -> (loc, AnonHyp n, c) ] ] + [ [ "("; id = ident; ":="; c = lconstr; ")" -> (loc, NamedHyp id, c) + | "("; n = natural; ":="; c = lconstr; ")" -> (loc, AnonHyp n, c) ] ] ; bindings: - [ [ c1 = constr; ":="; c2 = constr; bl = LIST0 simple_binding -> - ExplicitBindings - ((join_to_constr loc c2,NamedHyp (snd(coerce_to_id c1)), c2) :: bl) - | n = natural; ":="; c = constr; bl = LIST0 simple_binding -> - ExplicitBindings ((join_to_constr loc c,AnonHyp n, c) :: bl) - | c1 = constr; bl = LIST0 constr -> - ImplicitBindings (c1 :: bl) ] ] + [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> + ExplicitBindings bl + | bl = LIST1 constr -> ImplicitBindings bl ] ] ; constr_with_bindings: [ [ c = constr; l = with_bindings -> (c, l) ] ] @@ -149,85 +174,79 @@ GEXTEND Gram with_bindings: [ [ "with"; bl = bindings -> bl | -> NoBindings ] ] ; - unfold_occ: - [ [ nl = LIST0 integer; c = global_or_ltac_ref -> (nl,c) ] ] - ; red_flag: - [ [ IDENT "Beta" -> FBeta - | IDENT "Delta" -> FDeltaBut [] - | IDENT "Iota" -> FIota - | IDENT "Zeta" -> FZeta - | IDENT "Delta"; "["; idl = LIST1 global_or_ltac_ref; "]" -> FConst idl - | IDENT "Delta"; "-"; "["; idl = LIST1 global_or_ltac_ref; "]" -> FDeltaBut idl + [ [ IDENT "beta" -> FBeta + | IDENT "delta" -> FDeltaBut [] + | IDENT "iota" -> FIota + | IDENT "zeta" -> FZeta + | IDENT "delta"; "["; idl = LIST1 global; "]" -> FConst idl + | IDENT "delta"; "-"; "["; idl = LIST1 global; "]" -> FDeltaBut idl ] ] ; red_tactic: - [ [ IDENT "Red" -> Red false - | IDENT "Hnf" -> Hnf - | IDENT "Simpl"; po = OPT pattern_occ -> Simpl po - | 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 "Vm_compute" -> CbvVm - | IDENT "Unfold"; ul = LIST1 unfold_occ -> Unfold ul - | IDENT "Fold"; cl = LIST1 constr -> Fold cl - | IDENT "Pattern"; pl = LIST1 pattern_occ -> Pattern pl ] ] + [ [ IDENT "red" -> Red false + | IDENT "hnf" -> Hnf + | IDENT "simpl"; po = OPT pattern_occ -> Simpl po + | IDENT "cbv"; s = LIST1 red_flag -> Cbv (make_red_flag s) + | IDENT "lazy"; s = LIST1 red_flag -> Lazy (make_red_flag s) + | IDENT "compute" -> compute + | IDENT "vm_compute" -> CbvVm + | IDENT "unfold"; ul = LIST1 unfold_occ SEP "," -> Unfold ul + | IDENT "fold"; cl = LIST1 constr -> Fold cl + | IDENT "pattern"; pl = LIST1 pattern_occ SEP","-> Pattern pl ] ] ; (* This is [red_tactic] including possible extensions *) red_expr: - [ [ IDENT "Red" -> Red false - | IDENT "Hnf" -> Hnf - | IDENT "Simpl"; po = OPT pattern_occ -> Simpl po - | 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 - | IDENT "Vm_compute" -> CbvVm - | s = IDENT; OPT constr -> ExtraRedExpr s ] ] + [ [ IDENT "red" -> Red false + | IDENT "hnf" -> Hnf + | IDENT "simpl"; po = OPT pattern_occ -> Simpl po + | IDENT "cbv"; s = LIST1 red_flag -> Cbv (make_red_flag s) + | IDENT "lazy"; s = LIST1 red_flag -> Lazy (make_red_flag s) + | IDENT "compute" -> compute + | IDENT "vm_compute" -> CbvVm + | IDENT "unfold"; ul = LIST1 unfold_occ -> Unfold ul + | IDENT "fold"; cl = LIST1 constr -> Fold cl + | IDENT "pattern"; pl = LIST1 pattern_occ -> Pattern pl + | s = IDENT -> ExtraRedExpr s ] ] ; hypident: - [ [ id = id_or_meta -> id,[],(InHyp,ref None) - | "("; "Type"; "of"; id = id_or_meta; ")" -> - id,[],(InHypTypeOnly,ref None) + [ [ id = id_or_meta -> + id,InHyp + | "("; IDENT "type"; IDENT "of"; id = id_or_meta; ")" -> + id,InHypTypeOnly + | "("; IDENT "value"; IDENT "of"; id = id_or_meta; ")" -> + id,InHypValueOnly ] ] ; + hypident_occ: + [ [ (id,l)=hypident; occs=occurrences -> (id,occs,l) ] ] + ; clause: - [ [ "in"; idl = LIST1 hypident -> - {onhyps=Some idl;onconcl=false; concl_occs=[]} - | -> {onhyps=Some[];onconcl=true;concl_occs=[]} ] ] + [ [ "in"; "*"; occs=occurrences -> + {onhyps=None;onconcl=true;concl_occs=occs} + | "in"; "*"; "|-"; (b,occs)=concl_occ -> + {onhyps=None; onconcl=b; concl_occs=occs} + | "in"; hl=LIST0 hypident_occ SEP","; "|-"; (b,occs)=concl_occ -> + {onhyps=Some hl; onconcl=b; concl_occs=occs} + | "in"; hl=LIST0 hypident_occ SEP"," -> + {onhyps=Some hl; onconcl=false; concl_occs=[]} + | -> {onhyps=Some[];onconcl=true; concl_occs=[]} ] ] + ; + concl_occ: + [ [ "*"; occs = occurrences -> (true,occs) + | -> (false, []) ] ] ; simple_clause: [ [ "in"; idl = LIST1 id_or_meta -> idl | -> [] ] ] ; - pattern_occ_hyp_tail_list: - [ [ pl = pattern_occ_hyp_list -> pl - | -> {onhyps=Some[];onconcl=false; concl_occs=[]} ] ] - ; - pattern_occ_hyp_list: - [ [ nl = LIST1 natural; IDENT "Goal" -> - {onhyps=Some[];onconcl=true;concl_occs=nl} - | nl = LIST1 natural; id = id_or_meta; cls = pattern_occ_hyp_tail_list - -> {cls with - onhyps=option_app(fun l -> (id,nl,(InHyp,ref None))::l) - cls.onhyps} - | IDENT "Goal" -> {onhyps=Some[];onconcl=true;concl_occs=[]} - | id = id_or_meta; cls = pattern_occ_hyp_tail_list -> - {cls with - onhyps=option_app(fun l -> (id,[],(InHyp,ref None))::l) - cls.onhyps} ] ] - ; - clause_pattern: - [ [ "in"; p = pattern_occ_hyp_list -> p - | -> {onhyps=None; onconcl=true; concl_occs=[] } ] ] - ; fixdecl: - [ [ id = base_ident; "/"; n = natural; ":"; c = constr -> (id,n,c) ] ] + [ [ "("; id = ident; bl=LIST0 Constr.binder; ann=fixannot; + ":"; ty=lconstr; ")" -> (loc,id,bl,ann,ty) ] ] ; - cofixdecl: - [ [ id = base_ident; ":"; c = constr -> (id,c) ] ] + fixannot: + [ [ "{"; IDENT "struct"; id=name; "}" -> Some id + | -> None ] ] ; hintbases: [ [ "with"; "*" -> None @@ -243,134 +262,141 @@ GEXTEND Gram simple_tactic: [ [ (* Basic tactics *) - IDENT "Intros"; IDENT "until"; id = quantified_hypothesis -> + IDENT "intros"; IDENT "until"; id = quantified_hypothesis -> TacIntrosUntil id - | IDENT "Intros"; pl = intropatterns -> TacIntroPattern pl - | IDENT "Intro"; id = base_ident; IDENT "after"; id2 = identref -> + | IDENT "intros"; pl = intropatterns -> TacIntroPattern pl + | IDENT "intro"; id = ident; IDENT "after"; id2 = identref -> TacIntroMove (Some id, Some id2) - | IDENT "Intro"; IDENT "after"; id2 = identref -> + | IDENT "intro"; IDENT "after"; id2 = identref -> TacIntroMove (None, Some id2) - | IDENT "Intro"; id = base_ident -> TacIntroMove (Some id,None) - | IDENT "Intro" -> TacIntroMove (None, None) + | IDENT "intro"; id = ident -> TacIntroMove (Some id, None) + | IDENT "intro" -> TacIntroMove (None, None) + + | IDENT "assumption" -> TacAssumption + | IDENT "exact"; c = constr -> TacExact c + | IDENT "exact_no_check"; c = constr -> TacExactNoCheck c - | IDENT "Assumption" -> TacAssumption - | IDENT "Exact"; c = constr -> TacExact c - | IDENT "Exact_no_check"; c = constr -> TacExactNoCheck c - - | IDENT "Apply"; cl = constr_with_bindings -> TacApply cl - | IDENT "Elim"; cl = constr_with_bindings; el = OPT eliminator -> + | IDENT "apply"; cl = constr_with_bindings -> TacApply cl + | IDENT "elim"; cl = constr_with_bindings; el = OPT eliminator -> TacElim (cl,el) - | 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 = base_ident; n = natural -> TacFix (Some id,n) - | IDENT "Fix"; id = base_ident; n = natural; "with"; fd = LIST0 fixdecl -> - TacMutualFix (id,n,fd) - | IDENT "Cofix" -> TacCofix None - | IDENT "Cofix"; id = base_ident -> TacCofix (Some id) - | IDENT "Cofix"; id = base_ident; "with"; fd = LIST0 cofixdecl -> - TacMutualCofix (id,fd) + | IDENT "elimtype"; c = constr -> TacElimType c + | IDENT "case"; cl = constr_with_bindings -> TacCase cl + | IDENT "casetype"; c = constr -> TacCaseType c + | "fix"; n = natural -> TacFix (None,n) + | "fix"; id = ident; n = natural -> TacFix (Some id,n) + | "fix"; id = ident; n = natural; "with"; fd = LIST1 fixdecl -> + TacMutualFix (id,n,List.map mk_fix_tac fd) + | "cofix" -> TacCofix None + | "cofix"; id = ident -> TacCofix (Some id) + | "cofix"; id = ident; "with"; fd = LIST1 fixdecl -> + TacMutualCofix (id,List.map mk_cofix_tac fd) - | IDENT "Cut"; c = constr -> TacCut c - | IDENT "Assert"; c = constr -> TacTrueCut (Names.Anonymous,c) - | IDENT "Assert"; c = constr; ":"; t = constr -> - TacTrueCut (Names.Name (snd(coerce_to_id c)),t) - | IDENT "Assert"; c = constr; ":="; b = constr -> - TacForward (false,Names.Name (snd (coerce_to_id c)),b) - | IDENT "Pose"; c = constr; ":="; b = constr -> - TacForward (true,Names.Name (snd(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"; (_,na) = name; ":="; c = constr; p = clause_pattern - -> TacLetTac (na,c,p) -(* | IDENT "Instantiate"; n = natural; c = constr -> - TacInstantiate (n,c,ConclLocation ()) - | IDENT "Instantiate"; n = natural; c = constr; "in"; id = id_or_meta -> - TacInstantiate (n,c,HypLocation(id,InHypTypeOnly)) *) - | IDENT "Specialize"; n = OPT natural; lcb = constr_with_bindings -> + | IDENT "cut"; c = constr -> TacCut c + | IDENT "assert"; id = lpar_id_colon; t = lconstr; ")" -> + TacTrueCut (Names.Name id,t) + | IDENT "assert"; id = lpar_id_coloneq; b = lconstr; ")" -> + TacForward (false,Names.Name id,b) + | IDENT "assert"; c = constr -> TacTrueCut (Names.Anonymous,c) + | IDENT "pose"; id = lpar_id_coloneq; b = lconstr; ")" -> + TacForward (true,Names.Name id,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 "set"; id = lpar_id_coloneq; c = lconstr; ")"; + p = clause -> TacLetTac (Names.Name id,c,p) + | IDENT "set"; c = constr; p = clause -> + TacLetTac (Names.Anonymous,c,p) + (* | IDENT "instantiate"; "("; n = natural; ":="; c = lconstr; ")"; "in"; + hid = hypident -> + let (id,(hloc,_)) = hid in + TacInstantiate (n,c,HypLocation (id,hloc)) + | IDENT "instantiate"; "("; n = natural; ":="; c = lconstr; ")" -> + TacInstantiate (n,c,ConclLocation ()) *) + + | IDENT "specialize"; n = OPT natural; lcb = constr_with_bindings -> TacSpecialize (n,lcb) - | IDENT "LApply"; c = constr -> TacLApply c + | IDENT "lapply"; c = constr -> TacLApply c (* Derived basic tactics *) - | IDENT "Induction"; h = quantified_hypothesis -> TacSimpleInduction (h,ref []) - | IDENT "NewInduction"; c = induction_arg; el = OPT eliminator; - ids = with_names -> TacNewInduction (c,el,(ids,ref [])) - | IDENT "Double"; IDENT "Induction"; h1 = quantified_hypothesis; + | IDENT "simple"; IDENT"induction"; h = quantified_hypothesis -> + TacSimpleInduction h + | IDENT "induction"; c = induction_arg; ids = with_names; + el = OPT eliminator -> TacNewInduction (c,el,ids) + | IDENT "double"; IDENT "induction"; h1 = quantified_hypothesis; h2 = quantified_hypothesis -> TacDoubleInduction (h1,h2) - | IDENT "Destruct"; h = quantified_hypothesis -> TacSimpleDestruct h - | IDENT "NewDestruct"; c = induction_arg; el = OPT eliminator; - ids = with_names -> TacNewDestruct (c,el,(ids,ref [])) - | IDENT "Decompose"; IDENT "Record" ; c = constr -> TacDecomposeAnd c - | IDENT "Decompose"; IDENT "Sum"; c = constr -> TacDecomposeOr c - | IDENT "Decompose"; "["; l = LIST1 global_or_ltac_ref; "]"; c = constr + | IDENT "simple"; IDENT"destruct"; h = quantified_hypothesis -> + TacSimpleDestruct h + | IDENT "destruct"; c = induction_arg; ids = with_names; + el = OPT eliminator -> TacNewDestruct (c,el,ids) + | IDENT "decompose"; IDENT "record" ; c = constr -> TacDecomposeAnd c + | IDENT "decompose"; IDENT "sum"; c = constr -> TacDecomposeOr c + | IDENT "decompose"; "["; l = LIST1 global; "]"; c = constr -> TacDecompose (l,c) (* Automation tactic *) - | IDENT "Trivial"; db = hintbases -> TacTrivial db - | IDENT "Auto"; n = OPT int_or_var; db = hintbases -> TacAuto (n, db) + | IDENT "trivial"; db = hintbases -> TacTrivial db + | IDENT "auto"; n = OPT int_or_var; db = hintbases -> TacAuto (n, db) - | IDENT "AutoTDB"; n = OPT natural -> TacAutoTDB n - | IDENT "CDHyp"; id = identref -> TacDestructHyp (true,id) - | IDENT "DHyp"; id = identref -> TacDestructHyp (false,id) - | IDENT "DConcl" -> TacDestructConcl - | IDENT "SuperAuto"; l = autoargs -> TacSuperAuto l - | IDENT "Auto"; n = OPT int_or_var; IDENT "Decomp"; p = OPT natural -> +(* Obsolete since V8.0 + | IDENT "autotdb"; n = OPT natural -> TacAutoTDB n + | IDENT "cdhyp"; id = identref -> TacDestructHyp (true,id) + | IDENT "dhyp"; id = identref -> TacDestructHyp (false,id) + | IDENT "dconcl" -> TacDestructConcl + | IDENT "superauto"; l = autoargs -> TacSuperAuto l +*) + | IDENT "auto"; n = OPT int_or_var; IDENT "decomp"; p = OPT natural -> TacDAuto (n, p) (* Context management *) - | IDENT "Clear"; "-"; l = LIST1 id_or_ltac_ref -> TacClear (true, l) - | IDENT "Clear"; l = LIST0 id_or_ltac_ref -> TacClear (l=[], l) - | IDENT "ClearBody"; l = LIST1 id_or_ltac_ref -> TacClearBody l - | IDENT "Move"; id1 = id_or_ltac_ref; IDENT "after"; - id2 = id_or_ltac_ref -> TacMove (true,id1,id2) - | IDENT "Rename"; id1 = id_or_ltac_ref; IDENT "into"; - id2 = id_or_ltac_ref -> TacRename (id1,id2) + | IDENT "clear"; "-"; l = LIST1 id_or_meta -> TacClear (true, l) + | IDENT "clear"; l = LIST0 id_or_meta -> TacClear (l=[], l) + | IDENT "clearbody"; l = LIST1 id_or_meta -> TacClearBody l + | IDENT "move"; id1 = id_or_meta; IDENT "after"; id2 = id_or_meta -> + TacMove (true,id1,id2) + | IDENT "rename"; id1 = id_or_meta; IDENT "into"; id2 = id_or_meta -> + TacRename (id1,id2) (* Constructors *) - | IDENT "Left"; bl = with_bindings -> TacLeft bl - | IDENT "Right"; bl = with_bindings -> TacRight bl - | IDENT "Split"; bl = with_bindings -> TacSplit (false,bl) - | IDENT "Exists"; bl = bindings -> TacSplit (true,bl) - | IDENT "Exists" -> TacSplit (true,NoBindings) - | IDENT "Constructor"; n = num_or_meta; l = with_bindings -> + | IDENT "left"; bl = with_bindings -> TacLeft bl + | IDENT "right"; bl = with_bindings -> TacRight bl + | IDENT "split"; bl = with_bindings -> TacSplit (false,bl) + | "exists"; bl = bindings -> TacSplit (true,bl) + | "exists" -> TacSplit (true,NoBindings) + | IDENT "constructor"; n = num_or_meta; l = with_bindings -> TacConstructor (n,l) - | IDENT "Constructor"; t = OPT tactic -> TacAnyConstructor t + | IDENT "constructor"; t = OPT tactic -> TacAnyConstructor t (* Equivalence relations *) - | IDENT "Reflexivity" -> TacReflexivity - | IDENT "Symmetry"; cls = clause -> TacSymmetry cls - | IDENT "Transitivity"; c = constr -> TacTransitivity c + | IDENT "reflexivity" -> TacReflexivity + | IDENT "symmetry"; cls = clause -> TacSymmetry cls + | IDENT "transitivity"; c = constr -> TacTransitivity c (* Equality and inversion *) - | IDENT "Dependent"; k = - [ IDENT "Simple"; IDENT "Inversion" -> SimpleInversion - | IDENT "Inversion" -> FullInversion - | IDENT "Inversion_clear" -> FullInversionClear ]; + | IDENT "dependent"; k = + [ IDENT "simple"; IDENT "inversion" -> SimpleInversion + | IDENT "inversion" -> FullInversion + | IDENT "inversion_clear" -> FullInversionClear ]; hyp = quantified_hypothesis; ids = with_names; co = OPT ["with"; c = constr -> c] -> TacInversion (DepInversion (k,co,ids),hyp) - | IDENT "Simple"; IDENT "Inversion"; + | IDENT "simple"; IDENT "inversion"; hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp) - | IDENT "Inversion"; + | IDENT "inversion"; hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> TacInversion (NonDepInversion (FullInversion, cl, ids), hyp) - | IDENT "Inversion_clear"; + | IDENT "inversion_clear"; hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp) - | IDENT "Inversion"; hyp = quantified_hypothesis; + | IDENT "inversion"; hyp = quantified_hypothesis; "using"; c = constr; cl = simple_clause -> TacInversion (InversionUsing (c,cl), hyp) (* Conversion *) | r = red_tactic; cl = clause -> TacReduce (r, cl) (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) - | IDENT "Change"; (oc,c) = conversion; cl = clause -> TacChange (oc,c,cl) - + | IDENT "change"; (oc,c) = conversion; cl = clause -> TacChange (oc,c,cl) ] ] ; END;; diff --git a/parsing/g_tacticnew.ml4 b/parsing/g_tacticnew.ml4 deleted file mode 100644 index 180754668..000000000 --- a/parsing/g_tacticnew.ml4 +++ /dev/null @@ -1,402 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pp -open Pcoq -open Util -open Tacexpr -open Rawterm -open Genarg - -let compute = Cbv all_flags - -let tactic_kw = [ "->"; "<-" ] -let _ = List.iter (fun s -> Lexer.add_token("",s)) tactic_kw - -(* Hack to parse "(x:=t)" as an explicit argument without conflicts with the *) -(* admissible notation "(x t)" *) -let lpar_id_coloneq = - Gram.Entry.of_parser "lpar_id_coloneq" - (fun strm -> - match Stream.npeek 1 strm with - | [("","(")] -> - (match Stream.npeek 2 strm with - | [_; ("IDENT",s)] -> - (match Stream.npeek 3 strm with - | [_; _; ("", ":=")] -> - Stream.junk strm; Stream.junk strm; Stream.junk strm; - Names.id_of_string s - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - -(* idem for (x:=t) and (1:=t) *) -let test_lpar_idnum_coloneq = - Gram.Entry.of_parser "test_lpar_idnum_coloneq" - (fun strm -> - match Stream.npeek 1 strm with - | [("","(")] -> - (match Stream.npeek 2 strm with - | [_; (("IDENT"|"INT"),_)] -> - (match Stream.npeek 3 strm with - | [_; _; ("", ":=")] -> () - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - -(* idem for (x:t) *) -let lpar_id_colon = - Gram.Entry.of_parser "lpar_id_colon" - (fun strm -> - match Stream.npeek 1 strm with - | [("","(")] -> - (match Stream.npeek 2 strm with - | [_; ("IDENT",id)] -> - (match Stream.npeek 3 strm with - | [_; _; ("", ":")] -> - Stream.junk strm; Stream.junk strm; Stream.junk strm; - Names.id_of_string id - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) - -open Constr -open Prim -open Tactic - -let mk_fix_tac (loc,id,bl,ann,ty) = - let n = - match bl,ann with - [([_],_)], None -> 1 - | _, Some x -> - let ids = List.map snd (List.flatten (List.map fst bl)) in - (try list_index (snd x) ids - with Not_found -> error "no such fix variable") - | _ -> error "cannot guess decreasing argument of fix" in - (id,n,Topconstr.CProdN(loc,bl,ty)) - -let mk_cofix_tac (loc,id,bl,ann,ty) = - let _ = option_app (fun (aloc,_) -> - Util.user_err_loc - (aloc,"Constr:mk_cofix_tac", - Pp.str"Annotation forbidden in cofix expression")) ann in - (id,Topconstr.CProdN(loc,bl,ty)) - -(* Functions overloaded by quotifier *) -let induction_arg_of_constr c = - try ElimOnIdent (Topconstr.constr_loc c,snd(coerce_to_id c)) - with _ -> ElimOnConstr c - -(* Auxiliary grammar rules *) - -GEXTEND Gram - GLOBAL: simple_tactic constr_with_bindings quantified_hypothesis - bindings red_expr int_or_var open_constr casted_open_constr - simple_intropattern; - - int_or_var: - [ [ n = integer -> Genarg.ArgArg n - | id = identref -> Genarg.ArgVar id ] ] - ; - (* An identifier or a quotation meta-variable *) - id_or_meta: - [ [ id = identref -> AI id - - (* This is used in quotations *) - | id = METAIDENT -> MetaId (loc,id) ] ] - ; - (* A number or a quotation meta-variable *) - num_or_meta: - [ [ n = integer -> AI n - | id = METAIDENT -> MetaId (loc,id) - ] ] - ; - open_constr: - [ [ c = constr -> ((),c) ] ] - ; - casted_open_constr: - [ [ c = constr -> ((),c) ] ] - ; - induction_arg: - [ [ n = natural -> ElimOnAnonHyp n - | c = constr -> induction_arg_of_constr c - ] ] - ; - quantified_hypothesis: - [ [ id = ident -> NamedHyp id - | n = natural -> AnonHyp n ] ] - ; - conversion: - [ [ c = constr -> (None, c) - | c1 = constr; "with"; c2 = constr -> (Some ([],c1), c2) - | c1 = constr; "at"; nl = LIST1 integer; "with"; c2 = constr -> - (Some (nl,c1), c2) ] ] - ; - occurrences: - [ [ "at"; nl = LIST1 integer -> nl - | -> [] ] ] - ; - pattern_occ: - [ [ c = constr; nl = occurrences -> (nl,c) ] ] - ; - unfold_occ: - [ [ c = global; nl = occurrences -> (nl,c) ] ] - ; - intropatterns: - [ [ l = LIST0 simple_intropattern -> l ]] - ; - simple_intropattern: - [ [ "["; tc = LIST1 intropatterns SEP "|" ; "]" -> IntroOrAndPattern tc - | "("; tc = LIST1 simple_intropattern SEP "," ; ")" -> IntroOrAndPattern [tc] - | "_" -> IntroWildcard - | id = ident -> IntroIdentifier id - ] ] - ; - simple_binding: - [ [ "("; id = ident; ":="; c = lconstr; ")" -> (loc, NamedHyp id, c) - | "("; n = natural; ":="; c = lconstr; ")" -> (loc, AnonHyp n, c) ] ] - ; - bindings: - [ [ test_lpar_idnum_coloneq; bl = LIST1 simple_binding -> - ExplicitBindings bl - | bl = LIST1 constr -> ImplicitBindings bl ] ] - ; - constr_with_bindings: - [ [ c = constr; l = with_bindings -> (c, l) ] ] - ; - with_bindings: - [ [ "with"; bl = bindings -> bl | -> NoBindings ] ] - ; - red_flag: - [ [ IDENT "beta" -> FBeta - | IDENT "delta" -> FDeltaBut [] - | IDENT "iota" -> FIota - | IDENT "zeta" -> FZeta - | IDENT "delta"; "["; idl = LIST1 global; "]" -> FConst idl - | IDENT "delta"; "-"; "["; idl = LIST1 global; "]" -> FDeltaBut idl - ] ] - ; - red_tactic: - [ [ IDENT "red" -> Red false - | IDENT "hnf" -> Hnf - | IDENT "simpl"; po = OPT pattern_occ -> Simpl po - | IDENT "cbv"; s = LIST1 red_flag -> Cbv (make_red_flag s) - | IDENT "lazy"; s = LIST1 red_flag -> Lazy (make_red_flag s) - | IDENT "compute" -> compute - | IDENT "vm_compute" -> CbvVm - | IDENT "unfold"; ul = LIST1 unfold_occ SEP "," -> Unfold ul - | IDENT "fold"; cl = LIST1 constr -> Fold cl - | IDENT "pattern"; pl = LIST1 pattern_occ SEP","-> Pattern pl ] ] - ; - (* This is [red_tactic] including possible extensions *) - red_expr: - [ [ IDENT "red" -> Red false - | IDENT "hnf" -> Hnf - | IDENT "simpl"; po = OPT pattern_occ -> Simpl po - | IDENT "cbv"; s = LIST1 red_flag -> Cbv (make_red_flag s) - | IDENT "lazy"; s = LIST1 red_flag -> Lazy (make_red_flag s) - | IDENT "compute" -> compute - | IDENT "vm_compute" -> CbvVm - | IDENT "unfold"; ul = LIST1 unfold_occ -> Unfold ul - | IDENT "fold"; cl = LIST1 constr -> Fold cl - | IDENT "pattern"; pl = LIST1 pattern_occ -> Pattern pl - | s = IDENT -> ExtraRedExpr s ] ] - ; - hypident: - [ [ id = id_or_meta -> - id,InHyp - | "("; IDENT "type"; IDENT "of"; id = id_or_meta; ")" -> - id,InHypTypeOnly - | "("; IDENT "value"; IDENT "of"; id = id_or_meta; ")" -> - id,InHypValueOnly - ] ] - ; - hypident_occ: - [ [ (id,l)=hypident; occs=occurrences -> (id,occs,l) ] ] - ; - clause: - [ [ "in"; "*"; occs=occurrences -> - {onhyps=None;onconcl=true;concl_occs=occs} - | "in"; "*"; "|-"; (b,occs)=concl_occ -> - {onhyps=None; onconcl=b; concl_occs=occs} - | "in"; hl=LIST0 hypident_occ SEP","; "|-"; (b,occs)=concl_occ -> - {onhyps=Some hl; onconcl=b; concl_occs=occs} - | "in"; hl=LIST0 hypident_occ SEP"," -> - {onhyps=Some hl; onconcl=false; concl_occs=[]} - | -> {onhyps=Some[];onconcl=true; concl_occs=[]} ] ] - ; - concl_occ: - [ [ "*"; occs = occurrences -> (true,occs) - | -> (false, []) ] ] - ; - simple_clause: - [ [ "in"; idl = LIST1 id_or_meta -> idl - | -> [] ] ] - ; - fixdecl: - [ [ "("; id = ident; bl=LIST0 Constr.binder; ann=fixannot; - ":"; ty=lconstr; ")" -> (loc,id,bl,ann,ty) ] ] - ; - fixannot: - [ [ "{"; IDENT "struct"; id=name; "}" -> Some id - | -> None ] ] - ; - hintbases: - [ [ "with"; "*" -> None - | "with"; l = LIST1 IDENT -> Some l - | -> Some [] ] ] - ; - eliminator: - [ [ "using"; el = constr_with_bindings -> el ] ] - ; - with_names: - [ [ "as"; ipat = simple_intropattern -> Some ipat | -> None ] ] - ; - simple_tactic: - [ [ - (* Basic tactics *) - IDENT "intros"; IDENT "until"; id = quantified_hypothesis -> - TacIntrosUntil id - | IDENT "intros"; pl = intropatterns -> TacIntroPattern pl - | IDENT "intro"; id = ident; IDENT "after"; id2 = identref -> - TacIntroMove (Some id, Some id2) - | IDENT "intro"; IDENT "after"; id2 = identref -> - 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 "exact_no_check"; c = constr -> TacExactNoCheck c - - | IDENT "apply"; cl = constr_with_bindings -> TacApply cl - | IDENT "elim"; cl = constr_with_bindings; el = OPT eliminator -> - TacElim (cl,el) - | IDENT "elimtype"; c = constr -> TacElimType c - | IDENT "case"; cl = constr_with_bindings -> TacCase cl - | IDENT "casetype"; c = constr -> TacCaseType c - | "fix"; n = natural -> TacFix (None,n) - | "fix"; id = ident; n = natural -> TacFix (Some id,n) - | "fix"; id = ident; n = natural; "with"; fd = LIST1 fixdecl -> - TacMutualFix (id,n,List.map mk_fix_tac fd) - | "cofix" -> TacCofix None - | "cofix"; id = ident -> TacCofix (Some id) - | "cofix"; id = ident; "with"; fd = LIST1 fixdecl -> - TacMutualCofix (id,List.map mk_cofix_tac fd) - - | IDENT "cut"; c = constr -> TacCut c - | IDENT "assert"; id = lpar_id_colon; t = lconstr; ")" -> - TacTrueCut (Names.Name id,t) - | IDENT "assert"; id = lpar_id_coloneq; b = lconstr; ")" -> - TacForward (false,Names.Name id,b) - | IDENT "assert"; c = constr -> TacTrueCut (Names.Anonymous,c) - | IDENT "pose"; id = lpar_id_coloneq; b = lconstr; ")" -> - TacForward (true,Names.Name id,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 "set"; id = lpar_id_coloneq; c = lconstr; ")"; - p = clause -> TacLetTac (Names.Name id,c,p) - | IDENT "set"; c = constr; p = clause -> - TacLetTac (Names.Anonymous,c,p) - (* | IDENT "instantiate"; "("; n = natural; ":="; c = lconstr; ")"; "in"; - hid = hypident -> - let (id,(hloc,_)) = hid in - TacInstantiate (n,c,HypLocation (id,hloc)) - | IDENT "instantiate"; "("; n = natural; ":="; c = lconstr; ")" -> - TacInstantiate (n,c,ConclLocation ()) *) - - | IDENT "specialize"; n = OPT natural; lcb = constr_with_bindings -> - TacSpecialize (n,lcb) - | IDENT "lapply"; c = constr -> TacLApply c - - (* Derived basic tactics *) - | IDENT "simple"; IDENT"induction"; h = quantified_hypothesis -> - TacSimpleInduction h - | IDENT "induction"; c = induction_arg; ids = with_names; - el = OPT eliminator -> TacNewInduction (c,el,ids) - | IDENT "double"; IDENT "induction"; h1 = quantified_hypothesis; - h2 = quantified_hypothesis -> TacDoubleInduction (h1,h2) - | IDENT "simple"; IDENT"destruct"; h = quantified_hypothesis -> - TacSimpleDestruct h - | IDENT "destruct"; c = induction_arg; ids = with_names; - el = OPT eliminator -> TacNewDestruct (c,el,ids) - | IDENT "decompose"; IDENT "record" ; c = constr -> TacDecomposeAnd c - | IDENT "decompose"; IDENT "sum"; c = constr -> TacDecomposeOr c - | IDENT "decompose"; "["; l = LIST1 global; "]"; c = constr - -> TacDecompose (l,c) - - (* Automation tactic *) - | IDENT "trivial"; db = hintbases -> TacTrivial db - | IDENT "auto"; n = OPT int_or_var; db = hintbases -> TacAuto (n, db) - -(* Obsolete since V8.0 - | IDENT "autotdb"; n = OPT natural -> TacAutoTDB n - | IDENT "cdhyp"; id = identref -> TacDestructHyp (true,id) - | IDENT "dhyp"; id = identref -> TacDestructHyp (false,id) - | IDENT "dconcl" -> TacDestructConcl - | IDENT "superauto"; l = autoargs -> TacSuperAuto l -*) - | IDENT "auto"; n = OPT int_or_var; IDENT "decomp"; p = OPT natural -> - TacDAuto (n, p) - - (* Context management *) - | IDENT "clear"; "-"; l = LIST1 id_or_meta -> TacClear (true, l) - | IDENT "clear"; l = LIST0 id_or_meta -> TacClear (l=[], l) - | IDENT "clearbody"; l = LIST1 id_or_meta -> TacClearBody l - | IDENT "move"; id1 = id_or_meta; IDENT "after"; id2 = id_or_meta -> - TacMove (true,id1,id2) - | IDENT "rename"; id1 = id_or_meta; IDENT "into"; id2 = id_or_meta -> - TacRename (id1,id2) - - (* Constructors *) - | IDENT "left"; bl = with_bindings -> TacLeft bl - | IDENT "right"; bl = with_bindings -> TacRight bl - | IDENT "split"; bl = with_bindings -> TacSplit (false,bl) - | "exists"; bl = bindings -> TacSplit (true,bl) - | "exists" -> TacSplit (true,NoBindings) - | IDENT "constructor"; n = num_or_meta; l = with_bindings -> - TacConstructor (n,l) - | IDENT "constructor"; t = OPT tactic -> TacAnyConstructor t - - (* Equivalence relations *) - | IDENT "reflexivity" -> TacReflexivity - | IDENT "symmetry"; cls = clause -> TacSymmetry cls - | IDENT "transitivity"; c = constr -> TacTransitivity c - - (* Equality and inversion *) - | IDENT "dependent"; k = - [ IDENT "simple"; IDENT "inversion" -> SimpleInversion - | IDENT "inversion" -> FullInversion - | IDENT "inversion_clear" -> FullInversionClear ]; - hyp = quantified_hypothesis; - ids = with_names; co = OPT ["with"; c = constr -> c] -> - TacInversion (DepInversion (k,co,ids),hyp) - | IDENT "simple"; IDENT "inversion"; - hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> - TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp) - | IDENT "inversion"; - hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> - TacInversion (NonDepInversion (FullInversion, cl, ids), hyp) - | IDENT "inversion_clear"; - hyp = quantified_hypothesis; ids = with_names; cl = simple_clause -> - TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp) - | IDENT "inversion"; hyp = quantified_hypothesis; - "using"; c = constr; cl = simple_clause -> - TacInversion (InversionUsing (c,cl), hyp) - - (* Conversion *) - | r = red_tactic; cl = clause -> TacReduce (r, cl) - (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) - | IDENT "change"; (oc,c) = conversion; cl = clause -> TacChange (oc,c,cl) - ] ] - ; -END;; diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 0b4fc25f1..76ccc5df7 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -8,33 +8,35 @@ (* $Id$ *) +open Pp +open Util open Names open Topconstr open Vernacexpr open Pcoq -open Pp open Tactic -open Util -open Constr -open Vernac_ -open Prim open Decl_kinds - open Genarg +open Extend +open Ppextend +open Goptions -let evar_constr loc = CHole loc +open Prim +open Constr +open Vernac_ +open Module -let class_rawexpr = G_basevernac.class_rawexpr -let thm_token = G_proofs.thm_token +let vernac_kw = [ ";"; ","; ">->"; ":<"; "<:"; "where"; "at" ] +let _ = List.iter (fun s -> Lexer.add_token ("",s)) vernac_kw (* Rem: do not join the different GEXTEND into one, it breaks native *) (* compilation on PowerPC and Sun architectures *) -let filter_com (b,e) = - let (b,e) = unloc (b,e) in - Pp.comments := List.filter (fun ((b',e'),s) -> b'<b || e'>e) !Pp.comments +let check_command = Gram.Entry.create "vernac:check_command" +let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" +let thm_token = Gram.Entry.create "vernac:thm_token" +let def_body = Gram.Entry.create "vernac:def_body" -if !Options.v7 then GEXTEND Gram GLOBAL: vernac gallina_ext; vernac: @@ -44,64 +46,88 @@ GEXTEND Gram | g = gallina_ext; "." -> g | c = command; "." -> c | c = syntax; "." -> c - | n = natural; ":"; tac = Tactic.tactic; "." -> VernacSolve (n,tac,true) - | n = natural; ":"; tac = Tactic.tactic; "!!" -> VernacSolve (n,tac,false) - | n = natural; ":"; v = check_command; "." -> v (Some n) - | "["; l = vernac_list_tail -> VernacList l - - (* For translation from V7 to V8 *) - | IDENT "V7only"; v = vernac -> - filter_com loc; VernacV7only v - | IDENT "V8only"; v = vernac -> VernacV8only v - -(* - (* This is for "Grammar vernac" rules *) - | id = METAIDENT -> VernacVar (Names.id_of_string id) -*) + | "["; l = LIST1 located_vernac; "]"; "." -> VernacList l ] ] ; - - check_command: - [ [ IDENT "Eval"; r = Tactic.red_expr; "in"; c = constr -> - fun g -> VernacCheckMayEval (Some r, g, c) - | IDENT "Check"; c = constr -> - fun g -> VernacCheckMayEval (None, g, c) ] ] - ; vernac: FIRST [ [ IDENT "Time"; v = vernac -> VernacTime v ] ] ; vernac: LAST - [ [ tac = Tactic.tactic; "." -> VernacSolve (1,tac,true) - | tac = Tactic.tactic; "!!" -> VernacSolve (1,tac,false) - | IDENT "Existential"; n = natural; c = constr_body -> - VernacSolveExistential (n,c) - ] ] - ; - constr_body: - [ [ ":="; c = constr; ":"; t = constr -> CCast(loc,c,Term.DEFAULTcast,t) - | ":"; t = constr; ":="; c = constr -> CCast(loc,c,Term.DEFAULTcast,t) - | ":="; c = constr -> c ] ] + [ [ gln = OPT[n=natural; ":" -> n]; + tac = subgoal_command -> tac gln ] ] ; - vernac_list_tail: - [ [ v = located_vernac; l = vernac_list_tail -> v :: l - | "]"; "." -> [] ] ] + subgoal_command: + [ [ c = check_command; "." -> c + | tac = Tactic.tactic; + use_dft_tac = [ "." -> false | "..." -> true ] -> + (fun g -> + let g = match g with Some gl -> gl | _ -> 1 in + VernacSolve(g,tac,use_dft_tac)) ] ] ; 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" + "Keywords Variables/Hypotheses/Parameters expect more than one assumption" | _ -> () +let no_coercion loc (c,x) = + if c then Util.user_err_loc + (loc,"no_coercion",Pp.str"no coercion allowed here"); + x + (* Gallina declarations *) -if !Options.v7 then GEXTEND Gram - GLOBAL: gallina gallina_ext thm_token; + GLOBAL: gallina gallina_ext thm_token def_body; + gallina: + (* Definition, Theorem, Variable, Axiom, ... *) + [ [ thm = thm_token; id = identref; bl = LIST0 binder_let; ":"; + c = lconstr -> + VernacStartTheoremProof (thm, id, (bl, c), false, (fun _ _ -> ())) + | stre = assumption_token; bl = assum_list -> + VernacAssumption (stre, bl) + | stre = assumptions_token; bl = assum_list -> + test_plurial_form bl; + VernacAssumption (stre, bl) + | IDENT "Boxed";"Definition";id = identref; b = def_body -> + VernacDefinition ((Global,Definition true), id, b, (fun _ _ -> ())) + | IDENT "Unboxed";"Definition";id = identref; b = def_body -> + VernacDefinition ((Global,Definition false), id, b, (fun _ _ -> ())) + | (f,d) = def_token; id = identref; b = def_body -> + VernacDefinition (d, id, b, f) + (* Gallina inductive declarations *) + | f = finite_token; + indl = LIST1 inductive_definition SEP "with" -> + VernacInductive (f,indl) + | IDENT "Boxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" -> + VernacFixpoint (recs,true) + | IDENT "Unboxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" -> + VernacFixpoint (recs,false) + | "Fixpoint"; recs = LIST1 rec_definition SEP "with" -> + VernacFixpoint (recs,Options.boxed_definitions()) + | "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" -> + VernacCoFixpoint (corecs,false) + | IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l ] ] + ; + gallina_ext: + [ [ b = record_token; oc = opt_coercion; name = identref; + ps = LIST0 binder_let; ":"; + s = lconstr; ":="; cstr = OPT identref; "{"; + fs = LIST0 record_field SEP ";"; "}" -> + VernacRecord (b,(oc,name),ps,s,cstr,fs) +(* Non port ? + | f = finite_token; s = csort; id = identref; + indpar = LIST0 simple_binder; ":="; lc = constructor_list -> + VernacInductive (f,[id,None,indpar,s,lc]) +*) + ] ] + ; thm_token: [ [ "Theorem" -> Theorem | IDENT "Lemma" -> Lemma @@ -109,105 +135,27 @@ GEXTEND Gram | IDENT "Remark" -> Remark ] ] ; def_token: - [ [ IDENT "Boxed";"Definition" -> - (fun _ _ -> ()), (Global, Definition true) - | IDENT "Unboxed";"Definition" -> - (fun _ _ -> ()), (Global, Definition false) - | "Definition" -> - (fun _ _ -> ()), (Global, Definition false) - (*(Options.boxed_definitions())) *) - | IDENT "Local" -> + [ [ "Definition" -> + (fun _ _ -> ()), (Global, Definition (Options.boxed_definitions())) + | IDENT "Let" -> (fun _ _ -> ()), (Local, Definition (Options.boxed_definitions())) | IDENT "SubClass" -> Class.add_subclass_hook, (Global, SubClass) | IDENT "Local"; IDENT "SubClass" -> - Class.add_subclass_hook, (Local, SubClass) ] ] + Class.add_subclass_hook, (Local, SubClass) ] ] ; assumption_token: [ [ "Hypothesis" -> (Local, Logical) | "Variable" -> (Local, Definitional) | "Axiom" -> (Global, Logical) | "Parameter" -> (Global, Definitional) - | IDENT "Conjecture" -> (Global,Conjectural) ] ] + | IDENT "Conjecture" -> (Global, Conjectural) ] ] ; assumptions_token: [ [ IDENT "Hypotheses" -> (Local, Logical) | IDENT "Variables" -> (Local, Definitional) + | IDENT "Axioms" -> (Global, Logical) | IDENT "Parameters" -> (Global, Definitional) ] ] ; - of_type_with_opt_coercion: - [ [ ":>" -> true - | ":"; ">" -> true - | ":" -> false ] ] - ; - params: - [ [ idl = LIST1 identref SEP ","; coe = of_type_with_opt_coercion; - c = constr -> (coe,(idl,c)) - ] ] - ; - ne_params_list: - [ [ ll = LIST1 params SEP ";" -> ll ] ] - ; - name_comma_list_tail: - [ [ ","; nal = LIST1 name SEP "," -> nal | -> [] ] ] - ; - ident_comma_list_tail: - [ [ ","; nal = LIST1 identref SEP "," -> nal | -> [] ] ] - ; - decl_notation: - [ [ "where"; ntn = STRING; ":="; c = constr; - scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ] - ; - type_option: - [ [ ":"; c = constr -> c - | -> evar_constr loc ] ] - ; - opt_casted_constr: - [ [ c = constr; ":"; t = constr -> CCast(loc,c,Term.DEFAULTcast,t) - | c = constr -> c ] ] - ; - vardecls: - [ [ na = name; nal = name_comma_list_tail; c = type_option - -> LocalRawAssum (na::nal,c) - | na = name; "="; c = opt_casted_constr -> - LocalRawDef (na,c) - | na = name; ":="; c = opt_casted_constr -> - LocalRawDef (na,c) - ] ] - ; - binders: - [ [ "["; bl = LIST1 vardecls SEP ";"; "]" -> bl ] ] - ; - binders_list: - [ [ bls = LIST0 binders -> List.flatten bls ] ] - ; - reduce: - [ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r - | -> None ] ] - ; - def_body: - [ [ bl = binders_list; ":="; red = reduce; c = constr; ":"; t = constr -> - DefineBody (bl, red, c, Some t) - | bl = binders_list; ":"; t = constr; ":="; red = reduce; c = constr -> - DefineBody (bl, red, c, Some t) - | bl = binders_list; ":="; red = reduce; c = constr -> - DefineBody (bl, red, c, None) - | bl = binders_list; ":"; t = constr -> - ProveBody (bl, t) ] ] - ; - gallina: - (* Definition, Theorem, Variable, Axiom, ... *) - [ [ thm = thm_token; id = identref; bl = binders_list; ":"; c = constr -> - VernacStartTheoremProof (thm, id, (bl, c), false, (fun _ _ -> ())) - | (f,d) = def_token; id = identref; b = def_body -> - VernacDefinition (d, id, b, 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) - ] ] - ; - (* Gallina inductive declarations *) finite_token: [ [ "Inductive" -> true | "CoInductive" -> false ] ] @@ -215,204 +163,246 @@ GEXTEND Gram record_token: [ [ IDENT "Record" -> true | IDENT "Structure" -> false ] ] ; - constructor: - [ [ idl = LIST1 identref SEP ","; coe = of_type_with_opt_coercion; - c = constr -> List.map (fun id -> (coe,(id,c))) idl ] ] - ; - constructor_list: - [ [ "|"; l = LIST1 constructor SEP "|" -> List.flatten l - | l = LIST1 constructor SEP "|" -> List.flatten l - | -> [] ] ] - ; - block_old_style: - [ [ ind = oneind_old_style; "with"; indl = block_old_style -> ind :: indl - | ind = oneind_old_style -> [ind] ] ] + (* Simple definitions *) + def_body: + [ [ bl = LIST0 binder_let; ":="; red = reduce; c = lconstr -> + (match c with + CCast(_,c,k,t) -> DefineBody (bl, red, c, Some t) + | _ -> DefineBody (bl, red, c, None)) + | bl = LIST0 binder_let; ":"; t = lconstr; ":="; red = reduce; c = lconstr -> + DefineBody (bl, red, c, Some t) + | bl = LIST0 binder_let; ":"; t = lconstr -> + ProveBody (bl, t) ] ] ; - oneind_old_style: - [ [ id = identref; ":"; c = constr; ":="; lc = constructor_list -> - (id,c,lc) ] ] + reduce: + [ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r + | -> None ] ] ; - oneind: - [ [ id = identref; indpar = simple_binders_list; ":"; c = constr; - ":="; lc = constructor_list; ntn = OPT decl_notation -> + decl_notation: + [ [ OPT [ "where"; ntn = ne_string; ":="; c = constr; + scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ] ] + ; + (* Inductives and records *) + inductive_definition: + [ [ id = identref; indpar = LIST0 binder_let; ":"; c = lconstr; + ":="; lc = constructor_list; ntn = decl_notation -> (id,ntn,indpar,c,lc) ] ] ; - simple_binders_list: - [ [ bl = ne_simple_binders_list -> bl + constructor_list: + [ [ "|"; l = LIST1 constructor SEP "|" -> l + | l = LIST1 constructor SEP "|" -> l | -> [] ] ] ; +(* + csort: + [ [ s = sort -> CSort (loc,s) ] ] + ; +*) opt_coercion: [ [ ">" -> true | -> false ] ] ; - onescheme: - [ [ id = identref; ":="; dep = dep; ind = global; IDENT "Sort"; - s = sort -> (id,dep,ind,s) ] ] - ; - schemes: - [ [ recl = LIST1 onescheme SEP "with" -> recl ] ] - ; - dep: - [ [ IDENT "Induction"; IDENT "for" -> true - | IDENT "Minimality"; IDENT "for" -> false ] ] - ; - onerec: - [ [ id = base_ident; bl = ne_fix_binders; ":"; type_ = constr; - ":="; def = constr; ntn = OPT decl_notation -> - let ni = List.length (List.flatten (List.map fst bl)) - 1 in - let bl = List.map (fun(nal,ty)->LocalRawAssum(nal,ty)) bl in - ((id, ni, bl, type_, def), ntn) ] ] - ; - specifrec: - [ [ l = LIST1 onerec SEP "with" -> l ] ] - ; - onecorec: - [ [ id = base_ident; ":"; c = constr; ":="; def = constr -> - (id,[],c,def) ] ] - ; - specifcorec: - [ [ l = LIST1 onecorec SEP "with" -> l ] ] + (* (co)-fixpoints *) + rec_definition: + [ [ id = ident; bl = LIST1 binder_let; + annot = OPT rec_annotation; type_ = type_cstr; + ":="; def = lconstr; ntn = decl_notation -> + let names = List.map snd (names_of_local_assums bl) in + let ni = + match annot with + Some id -> + (try list_index (Name id) names - 1 + with Not_found -> Util.user_err_loc + (loc,"Fixpoint", + Pp.str "No argument named " ++ Nameops.pr_id id)) + | None -> + if List.length names > 1 then + Util.user_err_loc + (loc,"Fixpoint", + Pp.str "the recursive argument needs to be specified"); + 0 in + ((id, ni, bl, type_, def),ntn) ] ] + ; + corec_definition: + [ [ id = ident; bl = LIST0 binder_let; c = type_cstr; ":="; + def = lconstr -> + (id,bl,c ,def) ] ] + ; + rec_annotation: + [ [ "{"; IDENT "struct"; id=IDENT; "}" -> id_of_string id ] ] + ; + type_cstr: + [ [ ":"; c=lconstr -> c + | -> CHole loc ] ] + ; + (* Inductive schemes *) + scheme: + [ [ id = identref; ":="; dep = dep_scheme; "for"; ind = global; + IDENT "Sort"; s = sort -> + (id,dep,ind,s) ] ] + ; + dep_scheme: + [ [ IDENT "Induction" -> true + | IDENT "Minimality" -> false ] ] + ; + (* Various Binders *) +(* + (* ... without coercions *) + binder_nodef: + [ [ b = binder_let -> + (match b with + LocalRawAssum(l,ty) -> (l,ty) + | LocalRawDef _ -> + Util.user_err_loc + (loc,"fix_param",Pp.str"defined binder not allowed here")) ] ] ; +*) + (* ... with coercions *) record_field: - [ [ id = name; oc = of_type_with_opt_coercion; t = constr -> - (oc,AssumExpr (id,t)) - | id = name; oc = of_type_with_opt_coercion; t = constr; - ":="; b = constr -> - (oc,DefExpr (id,b,Some t)) - | id = name; ":="; b = constr -> - (false,DefExpr (id,b,None)) ] ] - ; - fields: - [ [ fs = LIST0 record_field SEP ";" -> fs ] ] - ; - simple_binders: - [ [ "["; bll = LIST1 vardecls SEP ";"; "]" -> bll ] ] - ; - ne_simple_binders_list: - [ [ bll = LIST1 simple_binders -> (List.flatten bll) ] ] - ; - fix_params: - [ [ idl = LIST1 name SEP ","; ":"; c = constr -> (idl, c) - | idl = LIST1 name SEP "," -> (idl, evar_constr dummy_loc) - ] ] - ; - fix_binders: - [ [ "["; bll = LIST1 fix_params SEP ";"; "]" -> bll ] ] - ; - ne_fix_binders: - [ [ bll = LIST1 fix_binders -> List.flatten bll ] ] + [ [ id = name -> (false,AssumExpr(id,CHole loc)) + | id = name; oc = of_type_with_opt_coercion; t = lconstr -> + (oc,AssumExpr (id,t)) + | id = name; oc = of_type_with_opt_coercion; + t = lconstr; ":="; b = lconstr -> (oc,DefExpr (id,b,Some t)) + | id = name; ":="; b = lconstr -> + match b with + CCast(_,b,_,t) -> (false,DefExpr(id,b,Some t)) + | _ -> (false,DefExpr(id,b,None)) ] ] + ; + assum_list: + [ [ bl = LIST1 assum_coe -> bl | b = simple_assum_coe -> [b] ] ] + ; + assum_coe: + [ [ "("; a = simple_assum_coe; ")" -> a ] ] + ; + simple_assum_coe: + [ [ idl = LIST1 identref; oc = of_type_with_opt_coercion; c = lconstr -> + (oc,(idl,c)) ] ] ; - rec_constructor: - [ [ c = identref -> Some c - | -> None ] ] - ; - gallina_ext: - [ [ IDENT "Mutual"; bl = ne_simple_binders_list ; f = finite_token; - indl = block_old_style -> - let indl' = List.map (fun (id,ar,c) -> (id,None,bl,ar,c)) indl in - VernacInductive (f,indl') - | b = record_token; oc = opt_coercion; name = identref; - ps = simple_binders_list; ":"; - s = constr; ":="; c = rec_constructor; "{"; fs = fields; "}" -> - VernacRecord (b,(oc,name),ps,s,c,fs) - ] ] - ; - gallina: - [ [ IDENT "Mutual"; f = finite_token; indl = LIST1 oneind SEP "with" -> - VernacInductive (f,indl) - | f = finite_token; indl = LIST1 oneind SEP "with" -> - VernacInductive (f,indl) - | IDENT "Boxed"; "Fixpoint"; recs = specifrec -> - VernacFixpoint (recs,true) - | IDENT "Unboxed"; "Fixpoint"; recs = specifrec -> - VernacFixpoint (recs,false) - | "Fixpoint"; recs = specifrec -> - VernacFixpoint (recs,Options.boxed_definitions()) -(* | IDENT "Boxed"; "CoFixpoint"; corecs = specifcorec -> - VernacCoFixpoint (corecs,true) - | IDENT "Unboxed"; "CoFixpoint"; corecs = specifcorec -> - VernacCoFixpoint (corecs,false) *) - | "CoFixpoint"; corecs = specifcorec -> - VernacCoFixpoint (corecs,false (*Options.boxed_definitions()*)) - | IDENT "Scheme"; l = schemes -> VernacScheme l - | f = finite_token; s = csort; id = identref; - indpar = simple_binders_list; ":="; lc = constructor_list -> - VernacInductive (f,[id,None,indpar,s,lc]) ] ] + constructor: + [ [ id = identref; l = LIST0 binder_let; + coe = of_type_with_opt_coercion; c = lconstr -> + (coe,(id,G_constr.mkCProdN loc l c)) + | id = identref; l = LIST0 binder_let -> + (false,(id,G_constr.mkCProdN loc l (CHole loc))) ] ] ; - csort: - [ [ s = sort -> CSort (loc,s) ] ] + of_type_with_opt_coercion: + [ [ ":>" -> true + | ":"; ">" -> true + | ":" -> false ] ] ; +END + + +(* Modules and Sections *) +GEXTEND Gram + GLOBAL: gallina_ext module_expr module_type; + gallina_ext: - [ [ -(* Sections *) - IDENT "Section"; id = identref -> VernacBeginSection id - | IDENT "Chapter"; id = identref -> VernacBeginSection id ] ] + [ [ (* Interactive module declaration *) + IDENT "Module"; export = export_token; id = identref; + bl = LIST0 module_binder; mty_o = OPT of_module_type; + mexpr_o = OPT is_module_expr -> + VernacDefineModule (export, id, bl, mty_o, mexpr_o) + + | IDENT "Module"; "Type"; id = identref; + bl = LIST0 module_binder; mty_o = OPT is_module_type -> + VernacDeclareModuleType (id, bl, mty_o) + + | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref; + bl = LIST0 module_binder; mty_o = of_module_type -> + VernacDeclareModule (export, id, bl, mty_o) + (* Section beginning *) + | IDENT "Section"; id = identref -> VernacBeginSection id + | IDENT "Chapter"; id = identref -> VernacBeginSection id + + (* This end a Section a Module or a Module Type *) + | IDENT "End"; id = identref -> VernacEndSegment id + + (* Requiring an already compiled module *) + | IDENT "Require"; export = export_token; specif = specif_token; + qidl = LIST1 global -> + VernacRequire (export, specif, qidl) + | IDENT "Require"; export = export_token; specif = specif_token; + filename = ne_string -> + VernacRequireFrom (export, specif, filename) + | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl) + | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl) ] ] ; export_token: [ [ IDENT "Import" -> Some false | IDENT "Export" -> Some true | -> None ] ] ; - module_vardecls: - [ [ export = export_token; id = identref; idl = ident_comma_list_tail; ":"; - mty = Module.module_type -> (export,id::idl,mty) ] ] - ; - module_binders: - [ [ "["; bl = LIST1 module_vardecls SEP ";"; "]" -> bl ] ] - ; - module_binders_list: - [ [ bls = LIST0 module_binders -> List.flatten bls ] ] + specif_token: + [ [ IDENT "Implementation" -> Some false + | IDENT "Specification" -> Some true + | -> None ] ] ; of_module_type: - [ [ ":"; mty = Module.module_type -> (mty, true) - | "<:"; mty = Module.module_type -> (mty, false) ] ] + [ [ ":"; mty = module_type -> (mty, true) + | "<:"; mty = module_type -> (mty, false) ] ] ; is_module_type: - [ [ ":="; mty = Module.module_type -> mty ] ] + [ [ ":="; mty = module_type -> mty ] ] ; is_module_expr: - [ [ ":="; mexpr = Module.module_expr -> mexpr ] ] + [ [ ":="; mexpr = module_expr -> mexpr ] ] ; - gallina_ext: - [ [ - (* Interactive module declaration *) - IDENT "Module"; export = export_token; id = identref; - bl = module_binders_list; mty_o = OPT of_module_type; - mexpr_o = OPT is_module_expr -> - VernacDefineModule (export, id, bl, mty_o, mexpr_o) - - | IDENT "Module"; "Type"; id = identref; - bl = module_binders_list; mty_o = OPT is_module_type -> - VernacDeclareModuleType (id, bl, mty_o) - - | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref; - bl = module_binders_list; mty_o = of_module_type -> - VernacDeclareModule (export, id, bl, mty_o) - (* This end a Section a Module or a Module Type *) + (* Module binder *) + module_binder: + [ [ "("; export = export_token; idl = LIST1 identref; ":"; + mty = module_type; ")" -> (export,idl,mty) ] ] + ; - | IDENT "End"; id = identref -> VernacEndSegment id + (* Module expressions *) + module_expr: + [ [ qid = qualid -> CMEident qid + | me1 = module_expr; me2 = module_expr -> CMEapply (me1,me2) + | "("; me = module_expr; ")" -> me +(* ... *) + ] ] + ; + with_declaration: + [ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.lconstr -> + CWith_Definition (fqid,c) + | IDENT "Module"; fqid = fullyqualid; ":="; qid = qualid -> + CWith_Module (fqid,qid) + ] ] + ; + module_type: + [ [ qid = qualid -> CMTEident qid +(* ... *) + | mty = module_type; "with"; decl = with_declaration -> + CMTEwith (mty,decl) ] ] + ; +END +(* Extensions: implicits, coercions, etc. *) +GEXTEND Gram + GLOBAL: gallina_ext; -(* Transparent and Opaque *) - | IDENT "Transparent"; l = LIST1 global -> VernacSetOpacity (false, l) + gallina_ext: + [ [ (* Transparent and Opaque *) + IDENT "Transparent"; l = LIST1 global -> VernacSetOpacity (false, l) | IDENT "Opaque"; l = LIST1 global -> VernacSetOpacity (true, l) -(* Canonical structure *) + (* Canonical structure *) | IDENT "Canonical"; IDENT "Structure"; qid = global -> VernacCanonical qid | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body -> - let s = Ast.coerce_global_to_id qid in + let s = coerce_global_to_id qid in VernacDefinition ((Global,CanonicalStructure),(dummy_loc,s),d, (fun _ -> Recordops.declare_canonical_structure)) -(* Coercions *) + (* Coercions *) | IDENT "Coercion"; qid = global; d = def_body -> - let s = Ast.coerce_global_to_id qid in + let s = coerce_global_to_id qid in VernacDefinition ((Global,Coercion),(dummy_loc,s),d,Class.add_coercion_hook) | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body -> - let s = Ast.coerce_global_to_id qid in + let s = coerce_global_to_id qid in VernacDefinition ((Local,Coercion),(dummy_loc,s),d,Class.add_coercion_hook) | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> @@ -426,114 +416,230 @@ GEXTEND Gram | IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> VernacCoercion (Global, qid, s, t) - | IDENT "Class"; IDENT "Local"; c = global -> - Pp.warning "Class is obsolete"; VernacNop - | IDENT "Class"; c = global -> - Pp.warning "Class is obsolete"; VernacNop -(* Implicit *) -(* - | IDENT "Syntactic"; "Definition"; id = identref; ":="; c = constr; - n = OPT [ "|"; n = natural -> n ] -> - VernacSyntacticDefinition (id,c,n) -*) - | IDENT "Syntactic"; "Definition"; id = ident; ":="; c = constr; - n = OPT [ "|"; n = natural -> n ] -> - let c = match n with - | Some n -> - let l = list_tabulate (fun _ -> (CHole (loc),None)) n in - CApp (loc,(None,c),l) - | None -> c in - VernacSyntacticDefinition (id,c,false,true) - | IDENT "Implicits"; qid = global; "["; l = LIST0 natural; "]" -> - let l = List.map (fun n -> ExplByPos n) l in - VernacDeclareImplicits (qid,Some l) - | IDENT "Implicits"; qid = global -> VernacDeclareImplicits (qid,None) - - | IDENT "Implicit"; ["Variable"; "Type" | IDENT "Variables"; "Type"]; - idl = LIST1 identref SEP ","; ":"; c = constr -> VernacReserve (idl,c) - - (* For compatibility *) - | IDENT "Implicit"; IDENT "Arguments"; IDENT "On" -> - VernacSetOption - (Goptions.SecondaryTable ("Implicit","Arguments"), BoolValue true) - | IDENT "Implicit"; IDENT "Arguments"; IDENT "Off" -> - VernacSetOption - (Goptions.SecondaryTable ("Implicit","Arguments"), BoolValue false) - ] ] + (* Implicit *) + | IDENT "Implicit"; IDENT "Arguments"; qid = global; + pos = OPT [ "["; l = LIST0 ident; "]" -> l ] -> + let pos = option_app (List.map (fun id -> ExplByName id)) pos in + VernacDeclareImplicits (qid,pos) + + | IDENT "Implicit"; ["Type" | IDENT "Types"]; + idl = LIST1 identref; ":"; c = lconstr -> VernacReserve (idl,c) ] ] ; END -(* Modules management *) -if !Options.v7 then GEXTEND Gram - GLOBAL: command; + GLOBAL: command check_command class_rawexpr; - export_token: - [ [ IDENT "Import" -> false - | IDENT "Export" -> true - | -> false ] ] - ; - specif_token: - [ [ IDENT "Implementation" -> Some false - | IDENT "Specification" -> Some true - | -> None ] ] - ; command: - [ [ "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ]; - s = [ s = STRING -> s | s = IDENT -> s ] -> + [ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l + + (* System directory *) + | IDENT "Pwd" -> VernacChdir None + | IDENT "Cd" -> VernacChdir None + | IDENT "Cd"; dir = ne_string -> VernacChdir (Some dir) + + (* Toplevel control *) + | IDENT "Drop" -> VernacToplevelControl Drop + | IDENT "ProtectedLoop" -> VernacToplevelControl ProtectedLoop + | IDENT "Quit" -> VernacToplevelControl Quit + + | IDENT "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ]; + s = [ s = ne_string -> s | s = IDENT -> s ] -> VernacLoad (verbosely, s) -(* | "Compile"; - verbosely = - [ IDENT "Verbose" -> "Verbose" - | -> "" ]; - IDENT "Module"; - only_spec = - [ IDENT "Specification" -> "Specification" - | -> "" ]; - mname = [ 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 global -> - VernacRequire (None, None, qidl) - | IDENT "Require"; export = export_token; specif = specif_token; - qidl = LIST1 global -> VernacRequire (Some export, specif, qidl) -(* | IDENT "Require"; export = export_token; specif = specif_token; - id = identref; filename = STRING -> - VernacRequireFrom (export, specif, id, filename) *) - | IDENT "Require"; export = export_token; specif = specif_token; - filename = STRING -> - VernacRequireFrom (Some export, specif, filename) - | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 STRING -> + | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 ne_string -> VernacDeclareMLModule l - | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl) - | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl) - ] -] + + (* Dump of the universe graph - to file or to stdout *) + | IDENT "Dump"; IDENT "Universes"; fopt = OPT ne_string -> + VernacPrint (PrintUniverses fopt) + + | IDENT "Locate"; l = locatable -> VernacLocate l + + (* Managing load paths *) + | IDENT "Add"; IDENT "LoadPath"; dir = ne_string; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = ne_string; + alias = as_dirpath -> VernacAddLoadPath (true, dir, alias) + | IDENT "Remove"; IDENT "LoadPath"; dir = ne_string -> + VernacRemoveLoadPath dir + + (* For compatibility *) + | IDENT "AddPath"; dir = ne_string; "as"; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "AddRecPath"; dir = ne_string; "as"; alias = as_dirpath -> + VernacAddLoadPath (true, dir, alias) + | IDENT "DelPath"; dir = ne_string -> + VernacRemoveLoadPath dir + + (* Type-Checking (pas dans le refman) *) + | "Type"; c = lconstr -> VernacGlobalCheck c + + (* Printing (careful factorization of entries) *) + | IDENT "Print"; p = printable -> VernacPrint p + | IDENT "Print"; qid = global -> VernacPrint (PrintName qid) + | IDENT "Print"; IDENT "Module"; "Type"; qid = global -> + VernacPrint (PrintModuleType qid) + | IDENT "Print"; IDENT "Module"; qid = global -> + VernacPrint (PrintModule qid) + | IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n) + | IDENT "About"; qid = global -> VernacPrint (PrintAbout qid) + + (* Searching the environment *) + | IDENT "Search"; qid = global; 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) + | IDENT "SearchAbout"; + sl = [ "["; l = LIST1 [ r = global -> SearchRef r + | s = ne_string -> SearchString s ]; "]" -> l + | qid = global -> [SearchRef qid] ]; + l = in_or_out_modules -> + VernacSearch (SearchAbout sl, l) + + | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = ne_string -> + VernacAddMLPath (false, dir) + | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = ne_string -> + VernacAddMLPath (true, dir) + + (* Pour intervenir sur les tables de paramtres *) + | "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 global 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) ] ] ; -END + check_command: (* TODO: rapprocher Eval et Check *) + [ [ IDENT "Eval"; r = Tactic.red_expr; "in"; c = lconstr -> + fun g -> VernacCheckMayEval (Some r, g, c) + | IDENT "Check"; c = lconstr -> + fun g -> VernacCheckMayEval (None, g, c) ] ] + ; + printable: + [ [ IDENT "Term"; qid = global -> PrintName qid + | IDENT "All" -> PrintFullContext + | IDENT "Section"; s = global -> PrintSectionContext s + | IDENT "Grammar"; ent = IDENT -> + (* This should be in "syntax" section but is here for factorization*) + PrintGrammar ("", ent) + | IDENT "LoadPath" -> PrintLoadPath + | IDENT "Modules" -> PrintModules -if !Options.v7 then -GEXTEND Gram - GLOBAL: command; + | IDENT "ML"; IDENT "Path" -> PrintMLLoadPath + | IDENT "ML"; IDENT "Modules" -> PrintMLModules + | IDENT "Graph" -> PrintGraph + | IDENT "Classes" -> PrintClasses + | IDENT "Ltac"; qid = global -> PrintLtac qid + | IDENT "Coercions" -> PrintCoercions + | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr + -> PrintCoercionPaths (s,t) + | IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions + | IDENT "Tables" -> PrintTables +(* Obsolete: was used for cooking V6.3 recipes ?? + | IDENT "Proof"; qid = global -> PrintOpaqueName qid +*) + | IDENT "Hint" -> PrintHintGoal + | IDENT "Hint"; qid = global -> PrintHint qid + | IDENT "Hint"; "*" -> PrintHintDb + | IDENT "HintDb"; s = IDENT -> PrintHintDbName s + | "Rewrite"; IDENT "HintDb"; s = IDENT -> PrintRewriteHintDbName s + | IDENT "Setoids" -> PrintSetoids + | IDENT "Scopes" -> PrintScopes + | IDENT "Scope"; s = IDENT -> PrintScope s + | IDENT "Visibility"; s = OPT IDENT -> PrintVisibility s + | IDENT "Implicit"; qid = global -> PrintImplicit qid ] ] + ; + class_rawexpr: + [ [ IDENT "Funclass" -> FunClass + | IDENT "Sortclass" -> SortClass + | qid = global -> RefClass qid ] ] + ; + locatable: + [ [ qid = global -> LocateTerm qid + | IDENT "File"; f = ne_string -> LocateFile f + | IDENT "Library"; qid = global -> LocateLibrary qid + | IDENT "Module"; qid = global -> LocateModule qid + | s = ne_string -> LocateNotation s ] ] + ; + option_value: + [ [ n = integer -> IntValue n + | s = STRING -> StringValue s ] ] + ; + option_ref_value: + [ [ id = global -> QualidRefValue id + | s = STRING -> StringRefValue s ] ] + ; + as_dirpath: + [ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ] + ; + in_or_out_modules: + [ [ IDENT "inside"; l = LIST1 global -> SearchInside l + | IDENT "outside"; l = LIST1 global -> SearchOutside l + | -> SearchOutside [] ] ] + ; + comment: + [ [ c = constr -> CommentConstr c + | s = STRING -> CommentString s + | n = natural -> CommentInt n ] ] + ; +END; +GEXTEND Gram command: [ [ - (* State management *) IDENT "Write"; IDENT "State"; s = IDENT -> VernacWriteState s - | IDENT "Write"; IDENT "State"; s = STRING -> VernacWriteState s + | IDENT "Write"; IDENT "State"; s = ne_string -> VernacWriteState s | IDENT "Restore"; IDENT "State"; s = IDENT -> VernacRestoreState s - | IDENT "Restore"; IDENT "State"; s = STRING -> VernacRestoreState s + | IDENT "Restore"; IDENT "State"; s = ne_string -> VernacRestoreState s (* Resetting *) | IDENT "Reset"; id = identref -> VernacResetName id | IDENT "Reset"; IDENT "Initial" -> VernacResetInitial | IDENT "Back" -> VernacBack 1 | IDENT "Back"; n = natural -> VernacBack n + | IDENT "BackTo"; n = natural -> VernacBackTo n + | IDENT "Backtrack"; n = natural ; m = natural ; p = natural -> + VernacBacktrack (n,m,p) (* Tactic Debugger *) | IDENT "Debug"; IDENT "On" -> VernacDebug true @@ -542,3 +648,86 @@ GEXTEND Gram ] ]; END ;; + +(* Grammar extensions *) + +GEXTEND Gram + GLOBAL: syntax; + + syntax: + [ [ IDENT "Open"; local = locality; IDENT "Scope"; sc = IDENT -> + VernacOpenCloseScope (local,true,sc) + + | IDENT "Close"; local = locality; IDENT "Scope"; sc = IDENT -> + VernacOpenCloseScope (local,false,sc) + + | IDENT "Delimit"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT -> + VernacDelimiters (sc,key) + + | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with"; + refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl) + + | IDENT "Arguments"; IDENT "Scope"; qid = global; + "["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl) + + | IDENT "Infix"; local = locality; + op = ne_string; ":="; p = global; + modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; + sc = OPT [ ":"; sc = IDENT -> sc ] -> + VernacInfix (local,(op,modl),p,sc) + | IDENT "Notation"; local = locality; id = ident; ":="; c = constr; + b = [ "("; IDENT "only"; IDENT "parsing"; ")" -> true | -> false ] -> + VernacSyntacticDefinition (id,c,local,b) + | IDENT "Notation"; local = locality; s = ne_string; ":="; c = constr; + modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; + sc = OPT [ ":"; sc = IDENT -> sc ] -> + VernacNotation (local,c,(s,modl),sc) + + | IDENT "Tactic"; IDENT "Notation"; n = tactic_level; + pil = LIST1 production_item; ":="; t = Tactic.tactic + -> VernacTacticNotation (n,pil,t) + + | IDENT "Reserved"; IDENT "Notation"; local = locality; s = ne_string; + l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] + -> VernacSyntaxExtension (local,(s,l)) + + (* "Print" "Grammar" should be here but is in "command" entry in order + to factorize with other "Print"-based vernac entries *) + ] ] + ; + tactic_level: + [ [ "("; "at"; IDENT "level"; n = natural; ")" -> n | -> 0 ] ] + ; + locality: + [ [ IDENT "Local" -> true | -> false ] ] + ; + level: + [ [ IDENT "level"; n = natural -> NumLevel n + | IDENT "next"; IDENT "level" -> NextLevel ] ] + ; + syntax_modifier: + [ [ x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev) + | x = IDENT; ","; l = LIST1 IDENT SEP ","; "at"; + lev = level -> SetItemLevel (x::l,lev) + | "at"; IDENT "level"; n = natural -> SetLevel n + | IDENT "left"; IDENT "associativity" -> SetAssoc Gramext.LeftA + | IDENT "right"; IDENT "associativity" -> SetAssoc Gramext.RightA + | IDENT "no"; IDENT "associativity" -> SetAssoc Gramext.NonA + | x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ) + | IDENT "only"; IDENT "parsing" -> SetOnlyParsing + | IDENT "format"; s = [s = STRING -> (loc,s)] -> SetFormat s ] ] + ; + syntax_extension_type: + [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference + | IDENT "bigint" -> ETBigint + ] ] + ; + opt_scope: + [ [ "_" -> None | sc = IDENT -> Some sc ] ] + ; + production_item: + [ [ s = ne_string -> VTerm s + | nt = IDENT; po = OPT [ "("; p = ident; ")" -> p ] -> + VNonTerm (loc,nt,po) ] ] + ; +END diff --git a/parsing/g_vernacnew.ml4 b/parsing/g_vernacnew.ml4 deleted file mode 100644 index 4a0d498d3..000000000 --- a/parsing/g_vernacnew.ml4 +++ /dev/null @@ -1,733 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(* $Id$ *) - -open Pp -open Util -open Names -open Topconstr -open Vernacexpr -open Pcoq -open Tactic -open Decl_kinds -open Genarg -open Extend -open Ppextend -open Goptions - -open Prim -open Constr -open Vernac_ -open Module - -let vernac_kw = [ ";"; ","; ">->"; ":<"; "<:"; "where"; "at" ] -let _ = List.iter (fun s -> Lexer.add_token ("",s)) vernac_kw - -(* Rem: do not join the different GEXTEND into one, it breaks native *) -(* compilation on PowerPC and Sun architectures *) - -let check_command = Gram.Entry.create "vernac:check_command" -let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" -let thm_token = Gram.Entry.create "vernac:thm_token" -let def_body = Gram.Entry.create "vernac:def_body" - -GEXTEND Gram - GLOBAL: vernac gallina_ext; - vernac: - (* Better to parse "." here: in case of failure (e.g. in coerce_to_var), *) - (* "." is still in the stream and discard_to_dot works correctly *) - [ [ g = gallina; "." -> g - | g = gallina_ext; "." -> g - | c = command; "." -> c - | c = syntax; "." -> c - | "["; l = LIST1 located_vernac; "]"; "." -> VernacList l - ] ] - ; - vernac: FIRST - [ [ IDENT "Time"; v = vernac -> VernacTime v ] ] - ; - vernac: LAST - [ [ gln = OPT[n=natural; ":" -> n]; - tac = subgoal_command -> tac gln ] ] - ; - subgoal_command: - [ [ c = check_command; "." -> c - | tac = Tactic.tactic; - use_dft_tac = [ "." -> false | "..." -> true ] -> - (fun g -> - let g = match g with Some gl -> gl | _ -> 1 in - VernacSolve(g,tac,use_dft_tac)) ] ] - ; - 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" - | _ -> () - -let no_coercion loc (c,x) = - if c then Util.user_err_loc - (loc,"no_coercion",Pp.str"no coercion allowed here"); - x - -(* Gallina declarations *) -GEXTEND Gram - GLOBAL: gallina gallina_ext thm_token def_body; - - gallina: - (* Definition, Theorem, Variable, Axiom, ... *) - [ [ thm = thm_token; id = identref; bl = LIST0 binder_let; ":"; - c = lconstr -> - VernacStartTheoremProof (thm, id, (bl, c), false, (fun _ _ -> ())) - | stre = assumption_token; bl = assum_list -> - VernacAssumption (stre, bl) - | stre = assumptions_token; bl = assum_list -> - test_plurial_form bl; - VernacAssumption (stre, bl) - | IDENT "Boxed";"Definition";id = identref; b = def_body -> - VernacDefinition ((Global,Definition true), id, b, (fun _ _ -> ())) - | IDENT "Unboxed";"Definition";id = identref; b = def_body -> - VernacDefinition ((Global,Definition false), id, b, (fun _ _ -> ())) - | (f,d) = def_token; id = identref; b = def_body -> - VernacDefinition (d, id, b, f) - (* Gallina inductive declarations *) - | f = finite_token; - indl = LIST1 inductive_definition SEP "with" -> - VernacInductive (f,indl) - | IDENT "Boxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" -> - VernacFixpoint (recs,true) - | IDENT "Unboxed";"Fixpoint"; recs = LIST1 rec_definition SEP "with" -> - VernacFixpoint (recs,false) - | "Fixpoint"; recs = LIST1 rec_definition SEP "with" -> - VernacFixpoint (recs,Options.boxed_definitions()) - | "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" -> - VernacCoFixpoint (corecs,false) - | IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l ] ] - ; - gallina_ext: - [ [ b = record_token; oc = opt_coercion; name = identref; - ps = LIST0 binder_let; ":"; - s = lconstr; ":="; cstr = OPT identref; "{"; - fs = LIST0 record_field SEP ";"; "}" -> - VernacRecord (b,(oc,name),ps,s,cstr,fs) -(* Non port ? - | f = finite_token; s = csort; id = identref; - indpar = LIST0 simple_binder; ":="; lc = constructor_list -> - VernacInductive (f,[id,None,indpar,s,lc]) -*) - ] ] - ; - thm_token: - [ [ "Theorem" -> Theorem - | IDENT "Lemma" -> Lemma - | IDENT "Fact" -> Fact - | IDENT "Remark" -> Remark ] ] - ; - def_token: - [ [ "Definition" -> - (fun _ _ -> ()), (Global, Definition (Options.boxed_definitions())) - | IDENT "Let" -> - (fun _ _ -> ()), (Local, Definition (Options.boxed_definitions())) - | IDENT "SubClass" -> Class.add_subclass_hook, (Global, SubClass) - | IDENT "Local"; IDENT "SubClass" -> - Class.add_subclass_hook, (Local, SubClass) ] ] - ; - assumption_token: - [ [ "Hypothesis" -> (Local, Logical) - | "Variable" -> (Local, Definitional) - | "Axiom" -> (Global, Logical) - | "Parameter" -> (Global, Definitional) - | IDENT "Conjecture" -> (Global, Conjectural) ] ] - ; - assumptions_token: - [ [ IDENT "Hypotheses" -> (Local, Logical) - | IDENT "Variables" -> (Local, Definitional) - | IDENT "Axioms" -> (Global, Logical) - | IDENT "Parameters" -> (Global, Definitional) ] ] - ; - finite_token: - [ [ "Inductive" -> true - | "CoInductive" -> false ] ] - ; - record_token: - [ [ IDENT "Record" -> true | IDENT "Structure" -> false ] ] - ; - (* Simple definitions *) - def_body: - [ [ bl = LIST0 binder_let; ":="; red = reduce; c = lconstr -> - (match c with - CCast(_,c,k,t) -> DefineBody (bl, red, c, Some t) - | _ -> DefineBody (bl, red, c, None)) - | bl = LIST0 binder_let; ":"; t = lconstr; ":="; red = reduce; c = lconstr -> - DefineBody (bl, red, c, Some t) - | bl = LIST0 binder_let; ":"; t = lconstr -> - ProveBody (bl, t) ] ] - ; - reduce: - [ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r - | -> None ] ] - ; - decl_notation: - [ [ OPT [ "where"; ntn = ne_string; ":="; c = constr; - scopt = OPT [ ":"; sc = IDENT -> sc] -> (ntn,c,scopt) ] ] ] - ; - (* Inductives and records *) - inductive_definition: - [ [ id = identref; indpar = LIST0 binder_let; ":"; c = lconstr; - ":="; lc = constructor_list; ntn = decl_notation -> - (id,ntn,indpar,c,lc) ] ] - ; - constructor_list: - [ [ "|"; l = LIST1 constructor SEP "|" -> l - | l = LIST1 constructor SEP "|" -> l - | -> [] ] ] - ; -(* - csort: - [ [ s = sort -> CSort (loc,s) ] ] - ; -*) - opt_coercion: - [ [ ">" -> true - | -> false ] ] - ; - (* (co)-fixpoints *) - rec_definition: - [ [ id = ident; bl = LIST1 binder_let; - annot = OPT rec_annotation; type_ = type_cstr; - ":="; def = lconstr; ntn = decl_notation -> - let names = List.map snd (names_of_local_assums bl) in - let ni = - match annot with - Some id -> - (try list_index (Name id) names - 1 - with Not_found -> Util.user_err_loc - (loc,"Fixpoint", - Pp.str "No argument named " ++ Nameops.pr_id id)) - | None -> - if List.length names > 1 then - Util.user_err_loc - (loc,"Fixpoint", - Pp.str "the recursive argument needs to be specified"); - 0 in - ((id, ni, bl, type_, def),ntn) ] ] - ; - corec_definition: - [ [ id = ident; bl = LIST0 binder_let; c = type_cstr; ":="; - def = lconstr -> - (id,bl,c ,def) ] ] - ; - rec_annotation: - [ [ "{"; IDENT "struct"; id=IDENT; "}" -> id_of_string id ] ] - ; - type_cstr: - [ [ ":"; c=lconstr -> c - | -> CHole loc ] ] - ; - (* Inductive schemes *) - scheme: - [ [ id = identref; ":="; dep = dep_scheme; "for"; ind = global; - IDENT "Sort"; s = sort -> - (id,dep,ind,s) ] ] - ; - dep_scheme: - [ [ IDENT "Induction" -> true - | IDENT "Minimality" -> false ] ] - ; - (* Various Binders *) -(* - (* ... without coercions *) - binder_nodef: - [ [ b = binder_let -> - (match b with - LocalRawAssum(l,ty) -> (l,ty) - | LocalRawDef _ -> - Util.user_err_loc - (loc,"fix_param",Pp.str"defined binder not allowed here")) ] ] - ; -*) - (* ... with coercions *) - record_field: - [ [ id = name -> (false,AssumExpr(id,CHole loc)) - | id = name; oc = of_type_with_opt_coercion; t = lconstr -> - (oc,AssumExpr (id,t)) - | id = name; oc = of_type_with_opt_coercion; - t = lconstr; ":="; b = lconstr -> (oc,DefExpr (id,b,Some t)) - | id = name; ":="; b = lconstr -> - match b with - CCast(_,b,_,t) -> (false,DefExpr(id,b,Some t)) - | _ -> (false,DefExpr(id,b,None)) ] ] - ; - assum_list: - [ [ bl = LIST1 assum_coe -> bl | b = simple_assum_coe -> [b] ] ] - ; - assum_coe: - [ [ "("; a = simple_assum_coe; ")" -> a ] ] - ; - simple_assum_coe: - [ [ idl = LIST1 identref; oc = of_type_with_opt_coercion; c = lconstr -> - (oc,(idl,c)) ] ] - ; - constructor: - [ [ id = identref; l = LIST0 binder_let; - coe = of_type_with_opt_coercion; c = lconstr -> - (coe,(id,G_constrnew.mkCProdN loc l c)) - | id = identref; l = LIST0 binder_let -> - (false,(id,G_constrnew.mkCProdN loc l (CHole loc))) ] ] - ; - of_type_with_opt_coercion: - [ [ ":>" -> true - | ":"; ">" -> true - | ":" -> false ] ] - ; -END - - -(* Modules and Sections *) -GEXTEND Gram - GLOBAL: gallina_ext module_expr module_type; - - gallina_ext: - [ [ (* Interactive module declaration *) - IDENT "Module"; export = export_token; id = identref; - bl = LIST0 module_binder; mty_o = OPT of_module_type; - mexpr_o = OPT is_module_expr -> - VernacDefineModule (export, id, bl, mty_o, mexpr_o) - - | IDENT "Module"; "Type"; id = identref; - bl = LIST0 module_binder; mty_o = OPT is_module_type -> - VernacDeclareModuleType (id, bl, mty_o) - - | IDENT "Declare"; IDENT "Module"; export = export_token; id = identref; - bl = LIST0 module_binder; mty_o = of_module_type -> - VernacDeclareModule (export, id, bl, mty_o) - (* Section beginning *) - | IDENT "Section"; id = identref -> VernacBeginSection id - | IDENT "Chapter"; id = identref -> VernacBeginSection id - - (* This end a Section a Module or a Module Type *) - | IDENT "End"; id = identref -> VernacEndSegment id - - (* Requiring an already compiled module *) - | IDENT "Require"; export = export_token; specif = specif_token; - qidl = LIST1 global -> - VernacRequire (export, specif, qidl) - | IDENT "Require"; export = export_token; specif = specif_token; - filename = ne_string -> - VernacRequireFrom (export, specif, filename) - | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl) - | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl) ] ] - ; - export_token: - [ [ IDENT "Import" -> Some false - | IDENT "Export" -> Some true - | -> None ] ] - ; - specif_token: - [ [ IDENT "Implementation" -> Some false - | IDENT "Specification" -> Some true - | -> None ] ] - ; - of_module_type: - [ [ ":"; mty = module_type -> (mty, true) - | "<:"; mty = module_type -> (mty, false) ] ] - ; - is_module_type: - [ [ ":="; mty = module_type -> mty ] ] - ; - is_module_expr: - [ [ ":="; mexpr = module_expr -> mexpr ] ] - ; - - (* Module binder *) - module_binder: - [ [ "("; export = export_token; idl = LIST1 identref; ":"; - mty = module_type; ")" -> (export,idl,mty) ] ] - ; - - (* Module expressions *) - module_expr: - [ [ qid = qualid -> CMEident qid - | me1 = module_expr; me2 = module_expr -> CMEapply (me1,me2) - | "("; me = module_expr; ")" -> me -(* ... *) - ] ] - ; - with_declaration: - [ [ "Definition"; fqid = fullyqualid; ":="; c = Constr.lconstr -> - CWith_Definition (fqid,c) - | IDENT "Module"; fqid = fullyqualid; ":="; qid = qualid -> - CWith_Module (fqid,qid) - ] ] - ; - module_type: - [ [ qid = qualid -> CMTEident qid -(* ... *) - | mty = module_type; "with"; decl = with_declaration -> - CMTEwith (mty,decl) ] ] - ; -END - -(* Extensions: implicits, coercions, etc. *) -GEXTEND Gram - GLOBAL: gallina_ext; - - gallina_ext: - [ [ (* Transparent and Opaque *) - IDENT "Transparent"; l = LIST1 global -> VernacSetOpacity (false, l) - | IDENT "Opaque"; l = LIST1 global -> VernacSetOpacity (true, l) - - (* Canonical structure *) - | IDENT "Canonical"; IDENT "Structure"; qid = global -> - VernacCanonical qid - | IDENT "Canonical"; IDENT "Structure"; qid = global; d = def_body -> - let s = coerce_global_to_id qid in - VernacDefinition - ((Global,CanonicalStructure),(dummy_loc,s),d, - (fun _ -> Recordops.declare_canonical_structure)) - - (* Coercions *) - | IDENT "Coercion"; qid = global; d = def_body -> - let s = coerce_global_to_id qid in - VernacDefinition ((Global,Coercion),(dummy_loc,s),d,Class.add_coercion_hook) - | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body -> - let s = coerce_global_to_id qid in - VernacDefinition ((Local,Coercion),(dummy_loc,s),d,Class.add_coercion_hook) - | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = identref; - ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> - VernacIdentityCoercion (Local, f, s, t) - | IDENT "Identity"; IDENT "Coercion"; f = identref; ":"; - s = class_rawexpr; ">->"; t = class_rawexpr -> - VernacIdentityCoercion (Global, f, s, t) - | IDENT "Coercion"; IDENT "Local"; qid = global; ":"; - s = class_rawexpr; ">->"; t = class_rawexpr -> - VernacCoercion (Local, qid, s, t) - | IDENT "Coercion"; qid = global; ":"; s = class_rawexpr; ">->"; - t = class_rawexpr -> - VernacCoercion (Global, qid, s, t) - - (* Implicit *) - | IDENT "Implicit"; IDENT "Arguments"; qid = global; - pos = OPT [ "["; l = LIST0 ident; "]" -> l ] -> - let pos = option_app (List.map (fun id -> ExplByName id)) pos in - VernacDeclareImplicits (qid,pos) - - | IDENT "Implicit"; ["Type" | IDENT "Types"]; - idl = LIST1 identref; ":"; c = lconstr -> VernacReserve (idl,c) ] ] - ; -END - -GEXTEND Gram - GLOBAL: command check_command class_rawexpr; - - command: - [ [ IDENT "Comments"; l = LIST0 comment -> VernacComments l - - (* System directory *) - | IDENT "Pwd" -> VernacChdir None - | IDENT "Cd" -> VernacChdir None - | IDENT "Cd"; dir = ne_string -> VernacChdir (Some dir) - - (* Toplevel control *) - | IDENT "Drop" -> VernacToplevelControl Drop - | IDENT "ProtectedLoop" -> VernacToplevelControl ProtectedLoop - | IDENT "Quit" -> VernacToplevelControl Quit - - | IDENT "Load"; verbosely = [ IDENT "Verbose" -> true | -> false ]; - s = [ s = ne_string -> s | s = IDENT -> s ] -> - VernacLoad (verbosely, s) - | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 ne_string -> - VernacDeclareMLModule l - - (* Dump of the universe graph - to file or to stdout *) - | IDENT "Dump"; IDENT "Universes"; fopt = OPT ne_string -> - VernacPrint (PrintUniverses fopt) - - | IDENT "Locate"; l = locatable -> VernacLocate l - - (* Managing load paths *) - | IDENT "Add"; IDENT "LoadPath"; dir = ne_string; alias = as_dirpath -> - VernacAddLoadPath (false, dir, alias) - | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = ne_string; - alias = as_dirpath -> VernacAddLoadPath (true, dir, alias) - | IDENT "Remove"; IDENT "LoadPath"; dir = ne_string -> - VernacRemoveLoadPath dir - - (* For compatibility *) - | IDENT "AddPath"; dir = ne_string; "as"; alias = as_dirpath -> - VernacAddLoadPath (false, dir, alias) - | IDENT "AddRecPath"; dir = ne_string; "as"; alias = as_dirpath -> - VernacAddLoadPath (true, dir, alias) - | IDENT "DelPath"; dir = ne_string -> - VernacRemoveLoadPath dir - - (* Type-Checking (pas dans le refman) *) - | "Type"; c = lconstr -> VernacGlobalCheck c - - (* Printing (careful factorization of entries) *) - | IDENT "Print"; p = printable -> VernacPrint p - | IDENT "Print"; qid = global -> VernacPrint (PrintName qid) - | IDENT "Print"; IDENT "Module"; "Type"; qid = global -> - VernacPrint (PrintModuleType qid) - | IDENT "Print"; IDENT "Module"; qid = global -> - VernacPrint (PrintModule qid) - | IDENT "Inspect"; n = natural -> VernacPrint (PrintInspect n) - | IDENT "About"; qid = global -> VernacPrint (PrintAbout qid) - - (* Searching the environment *) - | IDENT "Search"; qid = global; 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) - | IDENT "SearchAbout"; - sl = [ "["; l = LIST1 [ r = global -> SearchRef r - | s = ne_string -> SearchString s ]; "]" -> l - | qid = global -> [SearchRef qid] ]; - l = in_or_out_modules -> - VernacSearch (SearchAbout sl, l) - - | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = ne_string -> - VernacAddMLPath (false, dir) - | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = ne_string -> - VernacAddMLPath (true, dir) - - (* Pour intervenir sur les tables de paramtres *) - | "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 global 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) ] ] - ; - check_command: (* TODO: rapprocher Eval et Check *) - [ [ IDENT "Eval"; r = Tactic.red_expr; "in"; c = lconstr -> - fun g -> VernacCheckMayEval (Some r, g, c) - | IDENT "Check"; c = lconstr -> - fun g -> VernacCheckMayEval (None, g, c) ] ] - ; - printable: - [ [ IDENT "Term"; qid = global -> PrintName qid - | IDENT "All" -> PrintFullContext - | IDENT "Section"; s = global -> PrintSectionContext s - | IDENT "Grammar"; ent = IDENT -> - (* This should be in "syntax" section but is here for factorization*) - PrintGrammar ("", ent) - | IDENT "LoadPath" -> PrintLoadPath - | IDENT "Modules" -> PrintModules - - | IDENT "ML"; IDENT "Path" -> PrintMLLoadPath - | IDENT "ML"; IDENT "Modules" -> PrintMLModules - | IDENT "Graph" -> PrintGraph - | IDENT "Classes" -> PrintClasses - | IDENT "Ltac"; qid = global -> PrintLtac qid - | IDENT "Coercions" -> PrintCoercions - | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr - -> PrintCoercionPaths (s,t) - | IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions - | IDENT "Tables" -> PrintTables -(* Obsolete: was used for cooking V6.3 recipes ?? - | IDENT "Proof"; qid = global -> PrintOpaqueName qid -*) - | IDENT "Hint" -> PrintHintGoal - | IDENT "Hint"; qid = global -> PrintHint qid - | IDENT "Hint"; "*" -> PrintHintDb - | IDENT "HintDb"; s = IDENT -> PrintHintDbName s - | "Rewrite"; IDENT "HintDb"; s = IDENT -> PrintRewriteHintDbName s - | IDENT "Setoids" -> PrintSetoids - | IDENT "Scopes" -> PrintScopes - | IDENT "Scope"; s = IDENT -> PrintScope s - | IDENT "Visibility"; s = OPT IDENT -> PrintVisibility s - | IDENT "Implicit"; qid = global -> PrintImplicit qid ] ] - ; - class_rawexpr: - [ [ IDENT "Funclass" -> FunClass - | IDENT "Sortclass" -> SortClass - | qid = global -> RefClass qid ] ] - ; - locatable: - [ [ qid = global -> LocateTerm qid - | IDENT "File"; f = ne_string -> LocateFile f - | IDENT "Library"; qid = global -> LocateLibrary qid - | IDENT "Module"; qid = global -> LocateModule qid - | s = ne_string -> LocateNotation s ] ] - ; - option_value: - [ [ n = integer -> IntValue n - | s = STRING -> StringValue s ] ] - ; - option_ref_value: - [ [ id = global -> QualidRefValue id - | s = STRING -> StringRefValue s ] ] - ; - as_dirpath: - [ [ d = OPT [ "as"; d = dirpath -> d ] -> d ] ] - ; - in_or_out_modules: - [ [ IDENT "inside"; l = LIST1 global -> SearchInside l - | IDENT "outside"; l = LIST1 global -> SearchOutside l - | -> SearchOutside [] ] ] - ; - comment: - [ [ c = constr -> CommentConstr c - | s = STRING -> CommentString s - | n = natural -> CommentInt n ] ] - ; -END; - -GEXTEND Gram - command: - [ [ -(* State management *) - IDENT "Write"; IDENT "State"; s = IDENT -> VernacWriteState s - | IDENT "Write"; IDENT "State"; s = ne_string -> VernacWriteState s - | IDENT "Restore"; IDENT "State"; s = IDENT -> VernacRestoreState s - | IDENT "Restore"; IDENT "State"; s = ne_string -> VernacRestoreState s - -(* Resetting *) - | IDENT "Reset"; id = identref -> VernacResetName id - | IDENT "Reset"; IDENT "Initial" -> VernacResetInitial - | IDENT "Back" -> VernacBack 1 - | IDENT "Back"; n = natural -> VernacBack n - | IDENT "BackTo"; n = natural -> VernacBackTo n - | IDENT "Backtrack"; n = natural ; m = natural ; p = natural -> - VernacBacktrack (n,m,p) - -(* Tactic Debugger *) - | IDENT "Debug"; IDENT "On" -> VernacDebug true - | IDENT "Debug"; IDENT "Off" -> VernacDebug false - - ] ]; - END -;; - -(* Grammar extensions *) - -GEXTEND Gram - GLOBAL: syntax; - - syntax: - [ [ IDENT "Open"; local = locality; IDENT "Scope"; sc = IDENT -> - VernacOpenCloseScope (local,true,sc) - - | IDENT "Close"; local = locality; IDENT "Scope"; sc = IDENT -> - VernacOpenCloseScope (local,false,sc) - - | IDENT "Delimit"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT -> - VernacDelimiters (sc,key) - - | IDENT "Bind"; IDENT "Scope"; sc = IDENT; "with"; - refl = LIST1 class_rawexpr -> VernacBindScope (sc,refl) - - | IDENT "Arguments"; IDENT "Scope"; qid = global; - "["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl) - - | IDENT "Infix"; local = locality; - op = ne_string; ":="; p = global; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ] -> - VernacInfix (local,(op,modl),p,sc) - | IDENT "Notation"; local = locality; id = ident; ":="; c = constr; - b = [ "("; IDENT "only"; IDENT "parsing"; ")" -> true | -> false ] -> - VernacSyntacticDefinition (id,c,local,b) - | IDENT "Notation"; local = locality; s = ne_string; ":="; c = constr; - modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ] -> - VernacNotation (local,c,(s,modl),sc) - - | IDENT "Tactic"; IDENT "Notation"; n = tactic_level; - pil = LIST1 production_item; ":="; t = Tactic.tactic - -> VernacTacticNotation (n,pil,t) - - | IDENT "Reserved"; IDENT "Notation"; local = locality; s = ne_string; - l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] - -> VernacSyntaxExtension (local,(s,l)) - - (* "Print" "Grammar" should be here but is in "command" entry in order - to factorize with other "Print"-based vernac entries *) - ] ] - ; - tactic_level: - [ [ "("; "at"; IDENT "level"; n = natural; ")" -> n | -> 0 ] ] - ; - locality: - [ [ IDENT "Local" -> true | -> false ] ] - ; - level: - [ [ IDENT "level"; n = natural -> NumLevel n - | IDENT "next"; IDENT "level" -> NextLevel ] ] - ; - syntax_modifier: - [ [ x = IDENT; "at"; lev = level -> SetItemLevel ([x],lev) - | x = IDENT; ","; l = LIST1 IDENT SEP ","; "at"; - lev = level -> SetItemLevel (x::l,lev) - | "at"; IDENT "level"; n = natural -> SetLevel n - | IDENT "left"; IDENT "associativity" -> SetAssoc Gramext.LeftA - | IDENT "right"; IDENT "associativity" -> SetAssoc Gramext.RightA - | IDENT "no"; IDENT "associativity" -> SetAssoc Gramext.NonA - | x = IDENT; typ = syntax_extension_type -> SetEntryType (x,typ) - | IDENT "only"; IDENT "parsing" -> SetOnlyParsing - | IDENT "format"; s = [s = STRING -> (loc,s)] -> SetFormat s ] ] - ; - syntax_extension_type: - [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference - | IDENT "bigint" -> ETBigint - ] ] - ; - opt_scope: - [ [ "_" -> None | sc = IDENT -> Some sc ] ] - ; - production_item: - [ [ s = ne_string -> VTerm s - | nt = IDENT; po = OPT [ "("; p = ident; ")" -> p ] -> - VNonTerm (loc,nt,po) ] ] - ; -END diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml index e6de302f6..36470c13d 100644 --- a/parsing/ppconstr.ml +++ b/parsing/ppconstr.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id$ *) +(* $Id$ *) (*i*) open Util @@ -21,3 +21,719 @@ open Term open Pattern (*i*) +let sep_p = fun _ -> str"." +let sep_v = fun _ -> str"," ++ spc() +let sep_pp = fun _ -> str":" +let sep_bar = fun _ -> spc() ++ str"| " +let pr_tight_coma () = str "," ++ cut () + +let latom = 0 +let lannot = 100 +let lprod = 200 +let llambda = 200 +let lif = 200 +let lletin = 200 +let lfix = 200 +let larrow = 90 +let lcast = 100 +let larg = 9 +let lapp = 10 +let lposint = 0 +let lnegint = 35 (* must be consistent with Notation "- x" *) +let ltop = (200,E) +let lproj = 1 +let lsimple = (1,E) + +let prec_less child (parent,assoc) = + if parent < 0 && child = lprod then true + else + let parent = abs parent in + match assoc with + | E -> (<=) child parent + | L -> (<) child parent + | Prec n -> child<=n + | Any -> true + +let env_assoc_value v env = + try List.nth env (v-1) + with Not_found -> anomaly ("Inconsistent environment for pretty-print rule") + +let decode_constrlist_value = function + | CAppExpl (_,_,l) -> l + | CApp (_,_,l) -> List.map fst l + | _ -> anomaly "Ill-formed list argument of notation" + +let decode_patlist_value = function + | CPatCstr (_,_,l) -> l + | _ -> anomaly "Ill-formed list argument of notation" + +open Notation + +let rec print_hunk n decode pr env = function + | UnpMetaVar (e,prec) -> pr (n,prec) (env_assoc_value e env) + | UnpListMetaVar (e,prec,sl) -> + prlist_with_sep (fun () -> prlist (print_hunk n decode pr env) sl) + (pr (n,prec)) (decode (env_assoc_value e env)) + | UnpTerminal s -> str s + | UnpBox (b,sub) -> ppcmd_of_box b (prlist (print_hunk n decode pr env) sub) + | UnpCut cut -> ppcmd_of_cut cut + +let pr_notation_gen decode pr s env = + let unpl, level = find_notation_printing_rule s in + prlist (print_hunk level decode pr env) unpl, level + +let pr_notation = pr_notation_gen decode_constrlist_value +let pr_patnotation = pr_notation_gen decode_patlist_value + +let pr_delimiters key strm = + strm ++ str ("%"^key) + +let surround p = hov 1 (str"(" ++ p ++ str")") + +let pr_located pr ((b,e),x) = + if Options.do_translate() && (b,e)<>dummy_loc then + let (b,e) = unloc (b,e) in + comment b ++ pr x ++ comment e + else pr x + +let pr_com_at n = + if Options.do_translate() && n <> 0 then comment n + else mt() + +let pr_with_comments loc pp = pr_located (fun x -> x) (loc,pp) + +let pr_sep_com sep f c = pr_with_comments (constr_loc c) (sep() ++ f c) + +open Rawterm + +let pr_opt pr = function + | None -> mt () + | Some x -> spc() ++ pr x + +let pr_optc pr = function + | None -> mt () + | Some x -> pr_sep_com spc pr x + +let pr_universe = Univ.pr_uni + +let pr_sort = function + | RProp Term.Null -> str "Prop" + | RProp Term.Pos -> str "Set" + | RType u -> str "Type" ++ pr_opt pr_universe u + +let pr_expl_args pr (a,expl) = + match expl with + | None -> pr (lapp,L) a + | Some (_,ExplByPos n) -> + anomaly("Explicitation by position not implemented") + | Some (_,ExplByName id) -> + str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")" + +let pr_opt_type pr = function + | CHole _ -> mt () + | t -> cut () ++ str ":" ++ pr t + +let pr_opt_type_spc pr = function + | CHole _ -> mt () + | t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t + +let pr_id = pr_id + +let pr_name = function + | Anonymous -> str"_" + | Name id -> pr_id id + +let pr_lident (b,_ as loc,id) = + if loc <> dummy_loc then + let (b,_) = unloc loc in + pr_located pr_id (make_loc (b,b+String.length(string_of_id id)),id) + else pr_id id + +let pr_lname = function + (loc,Name id) -> pr_lident (loc,id) + | lna -> pr_located pr_name lna + +let pr_or_var pr = function + | Genarg.ArgArg x -> pr x + | Genarg.ArgVar (loc,s) -> pr_lident (loc,s) + +let las = lapp +let lpator = 100 + +let rec pr_patt sep inh p = + let (strm,prec) = match p with + | CPatAlias (_,p,id) -> + pr_patt mt (las,E) p ++ str " as " ++ pr_id id, las + | CPatCstr (_,c,[]) -> pr_reference c, latom + | CPatCstr (_,c,args) -> + pr_reference c ++ prlist (pr_patt spc (lapp,L)) args, lapp + | CPatAtom (_,None) -> str "_", latom + | CPatAtom (_,Some r) -> pr_reference r, latom + | CPatOr (_,pl) -> + hov 0 (prlist_with_sep pr_bar (pr_patt spc (lpator,L)) pl), lpator + | CPatNotation (_,"( _ )",[p]) -> + pr_patt (fun()->str"(") (max_int,E) p ++ str")", latom + | CPatNotation (_,s,env) -> pr_patnotation (pr_patt mt) s env + | CPatNumeral (_,i) -> Bigint.pr_bigint i, latom + | CPatDelimiters (_,k,p) -> pr_delimiters k (pr_patt mt lsimple p), 1 + in + let loc = cases_pattern_loc p in + pr_with_comments loc + (sep() ++ if prec_less prec inh then strm else surround strm) + +let pr_patt = pr_patt mt + + +let pr_eqn pr (loc,pl,rhs) = + spc() ++ hov 4 + (pr_with_comments loc + (str "| " ++ + hov 0 (prlist_with_sep sep_v (pr_patt ltop) pl ++ str " =>") ++ + pr_sep_com spc (pr ltop) rhs)) + +let begin_of_binder = function + LocalRawDef((loc,_),_) -> fst (unloc loc) + | LocalRawAssum((loc,_)::_,_) -> fst (unloc loc) + | _ -> assert false + +let begin_of_binders = function + | b::_ -> begin_of_binder b + | _ -> 0 + +let pr_binder many pr (nal,t) = + match t with + | CHole _ -> prlist_with_sep spc pr_lname nal + | _ -> + let s = prlist_with_sep spc pr_lname nal ++ str" : " ++ pr t in + hov 1 (if many then surround s else s) + +let pr_binder_among_many pr_c = function + | LocalRawAssum (nal,t) -> + pr_binder true pr_c (nal,t) + | LocalRawDef (na,c) -> + let c,topt = match c with + | CCast(_,c,_,t) -> c, t + | _ -> c, CHole dummy_loc in + hov 1 (surround + (pr_lname na ++ pr_opt_type pr_c topt ++ + str":=" ++ cut() ++ pr_c c)) + +let pr_undelimited_binders pr_c = + prlist_with_sep spc (pr_binder_among_many pr_c) + +let pr_delimited_binders kw pr_c bl = + let n = begin_of_binders bl in + match bl with + | [LocalRawAssum (nal,t)] -> + pr_com_at n ++ kw() ++ pr_binder false pr_c (nal,t) + | LocalRawAssum _ :: _ as bdl -> + pr_com_at n ++ kw() ++ pr_undelimited_binders pr_c bdl + | _ -> assert false + +let pr_let_binder pr x a = + hov 0 (hov 0 (pr_name x ++ brk(0,1) ++ str ":=") ++ + pr_sep_com (fun () -> brk(0,1)) (pr ltop) a) + +let rec extract_prod_binders = function +(* | CLetIn (loc,na,b,c) as x -> + let bl,c = extract_prod_binders c in + if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*) + | CProdN (loc,[],c) -> + extract_prod_binders c + | CProdN (loc,(nal,t)::bl,c) -> + let bl,c = extract_prod_binders (CProdN(loc,bl,c)) in + LocalRawAssum (nal,t) :: bl, c + | c -> [], c + +let rec extract_lam_binders = function +(* | CLetIn (loc,na,b,c) as x -> + let bl,c = extract_lam_binders c in + if bl = [] then [], x else LocalRawDef (na,b) :: bl, c*) + | CLambdaN (loc,[],c) -> + extract_lam_binders c + | CLambdaN (loc,(nal,t)::bl,c) -> + let bl,c = extract_lam_binders (CLambdaN(loc,bl,c)) in + LocalRawAssum (nal,t) :: bl, c + | c -> [], c + +let pr_global vars ref = pr_global_env vars ref + +let split_lambda = function + | CLambdaN (loc,[[na],t],c) -> (na,t,c) + | CLambdaN (loc,([na],t)::bl,c) -> (na,t,CLambdaN(loc,bl,c)) + | CLambdaN (loc,(na::nal,t)::bl,c) -> (na,t,CLambdaN(loc,(nal,t)::bl,c)) + | _ -> anomaly "ill-formed fixpoint body" + +let rename na na' t c = + match (na,na') with + | (_,Name id), (_,Name id') -> (na',t,replace_vars_constr_expr [id,id'] c) + | (_,Name id), (_,Anonymous) -> (na,t,c) + | _ -> (na',t,c) + +let split_product na' = function + | CArrow (loc,t,c) -> (na',t,c) + | CProdN (loc,[[na],t],c) -> rename na na' t c + | CProdN (loc,([na],t)::bl,c) -> rename na na' t (CProdN(loc,bl,c)) + | CProdN (loc,(na::nal,t)::bl,c) -> + rename na na' t (CProdN(loc,(nal,t)::bl,c)) + | _ -> anomaly "ill-formed fixpoint body" + +let merge_binders (na1,ty1) cofun (na2,ty2) codom = + let na = + match snd na1, snd na2 with + Anonymous, Name id -> + if occur_var_constr_expr id cofun then + failwith "avoid capture" + else na2 + | Name id, Anonymous -> + if occur_var_constr_expr id codom then + failwith "avoid capture" + else na1 + | Anonymous, Anonymous -> na1 + | Name id1, Name id2 -> + if id1 <> id2 then failwith "not same name" else na1 in + let ty = + match ty1, ty2 with + CHole _, _ -> ty2 + | _, CHole _ -> ty1 + | _ -> + Constrextern.check_same_type ty1 ty2; + ty2 in + (LocalRawAssum ([na],ty), codom) + +let rec strip_domain bvar cofun c = + match c with + | CArrow(loc,a,b) -> + merge_binders bvar cofun ((dummy_loc,Anonymous),a) b + | CProdN(loc,[([na],ty)],c') -> + merge_binders bvar cofun (na,ty) c' + | CProdN(loc,([na],ty)::bl,c') -> + merge_binders bvar cofun (na,ty) (CProdN(loc,bl,c')) + | CProdN(loc,(na::nal,ty)::bl,c') -> + merge_binders bvar cofun (na,ty) (CProdN(loc,(nal,ty)::bl,c')) + | _ -> failwith "not a product" + +(* Note: binder sharing is lost *) +let rec strip_domains (nal,ty) cofun c = + match nal with + [] -> assert false + | [na] -> + let bnd, c' = strip_domain (na,ty) cofun c in + ([bnd],None,c') + | na::nal -> + let f = CLambdaN(dummy_loc,[(nal,ty)],cofun) in + let bnd, c1 = strip_domain (na,ty) f c in + (try + let bl, rest, c2 = strip_domains (nal,ty) cofun c1 in + (bnd::bl, rest, c2) + with Failure _ -> ([bnd],Some (nal,ty), c1)) + +(* Re-share binders *) +let rec factorize_binders = function + | ([] | [_] as l) -> l + | LocalRawAssum (nal,ty) as d :: (LocalRawAssum (nal',ty')::l as l') -> + (try + let _ = Constrextern.check_same_type ty ty' in + factorize_binders (LocalRawAssum (nal@nal',ty)::l) + with _ -> + d :: factorize_binders l') + | d :: l -> d :: factorize_binders l + +(* Extract lambdas when a type constraint occurs *) +let rec extract_def_binders c ty = + match c with + | CLambdaN(loc,bvar::lams,b) -> + (try + let f = CLambdaN(loc,lams,b) in + let bvar', rest, ty' = strip_domains bvar f ty in + let c' = + match rest, lams with + None,[] -> b + | None, _ -> f + | Some bvar,_ -> CLambdaN(loc,bvar::lams,b) in + let (bl,c2,ty2) = extract_def_binders c' ty' in + (factorize_binders (bvar'@bl), c2, ty2) + with Failure _ -> + ([],c,ty)) + | _ -> ([],c,ty) + +let rec split_fix n typ def = + if n = 0 then ([],typ,def) + else + let (na,_,def) = split_lambda def in + let (na,t,typ) = split_product na typ in + let (bl,typ,def) = split_fix (n-1) typ def in + (LocalRawAssum ([na],t)::bl,typ,def) + +let pr_recursive_decl pr pr_dangling dangling_with_for id bl annot t c = + let pr_body = + if dangling_with_for then pr_dangling else pr in + pr_id id ++ str" " ++ + hov 0 (pr_undelimited_binders (pr ltop) bl ++ annot) ++ + pr_opt_type_spc pr t ++ str " :=" ++ + pr_sep_com (fun () -> brk(1,2)) (pr_body ltop) c + +let pr_fixdecl pr prd dangling_with_for (id,n,bl,t,c) = + let annot = + let ids = names_of_local_assums bl in + if List.length ids > 1 then + spc() ++ str "{struct " ++ pr_name (snd (List.nth ids n)) ++ str"}" + else mt() in + pr_recursive_decl pr prd dangling_with_for id bl annot t c + +let pr_cofixdecl pr prd dangling_with_for (id,bl,t,c) = + pr_recursive_decl pr prd dangling_with_for id bl (mt()) t c + +let pr_recursive pr_decl id = function + | [] -> anomaly "(co)fixpoint with no definition" + | [d1] -> pr_decl false d1 + | dl -> + prlist_with_sep (fun () -> fnl() ++ str "with ") + (pr_decl true) dl ++ + fnl() ++ str "for " ++ pr_id id + +let pr_arg pr x = spc () ++ pr x + +let is_var id = function + | CRef (Ident (_,id')) when id=id' -> true + | _ -> false + +let tm_clash = function + | (CRef (Ident (_,id)), Some (CApp (_,_,nal))) + when List.exists (function CRef (Ident (_,id')),_ -> id=id' | _ -> false) + nal + -> Some id + | (CRef (Ident (_,id)), Some (CAppExpl (_,_,nal))) + when List.exists (function CRef (Ident (_,id')) -> id=id' | _ -> false) + nal + -> Some id + | _ -> None + +let pr_case_item pr (tm,(na,indnalopt)) = + hov 0 (pr (lcast,E) tm ++ +(* + (match na with + | Name id when not (is_var id tm) -> spc () ++ str "as " ++ pr_id id + | Anonymous when tm_clash (tm,indnalopt) <> None -> + (* hide [tm] name to avoid conflicts *) + spc () ++ str "as _" (* ++ pr_id (out_some (tm_clash (tm,indnalopt)))*) + | _ -> mt ()) ++ +*) + (match na with (* Decision of printing "_" or not moved to constrextern.ml *) + | Some na -> spc () ++ str "as " ++ pr_name na + | None -> mt ()) ++ + (match indnalopt with + | None -> mt () +(* + | Some (_,ind,nal) -> + spc () ++ str "in " ++ + hov 0 (pr_reference ind ++ prlist (pr_arg pr_name) nal)) +*) + | Some t -> spc () ++ str "in " ++ pr lsimple t)) + +let pr_case_type pr po = + match po with + | None | Some (CHole _) -> mt() + | Some p -> + spc() ++ hov 2 (str "return" ++ pr_sep_com spc (pr lsimple) p) + +let pr_return_type pr po = pr_case_type pr po + +let pr_simple_return_type pr na po = + (match na with + | Some (Name id) -> + spc () ++ str "as " ++ pr_id id + | _ -> mt ()) ++ + pr_case_type pr po + +let pr_proj pr pr_app a f l = + hov 0 (pr lsimple a ++ cut() ++ str ".(" ++ pr_app pr f l ++ str ")") + +let pr_appexpl pr f l = + hov 2 ( + str "@" ++ pr_reference f ++ + prlist (pr_sep_com spc (pr (lapp,L))) l) + +let pr_app pr a l = + hov 2 ( + pr (lapp,L) a ++ + prlist (fun a -> spc () ++ pr_expl_args pr a) l) + +let rec pr sep inherited a = + let (strm,prec) = match a with + | CRef r -> pr_reference r, latom + | CFix (_,id,fix) -> + hov 0 (str"fix " ++ + pr_recursive + (pr_fixdecl (pr mt) (pr_dangling_with_for mt)) (snd id) fix), + lfix + | CCoFix (_,id,cofix) -> + hov 0 (str "cofix " ++ + pr_recursive + (pr_cofixdecl (pr mt) (pr_dangling_with_for mt)) (snd id) cofix), + lfix + | CArrow (_,a,b) -> + hov 0 (pr mt (larrow,L) a ++ str " ->" ++ + pr (fun () ->brk(1,0)) (-larrow,E) b), + larrow + | CProdN _ -> + let (bl,a) = extract_prod_binders a in + hov 0 ( + hov 2 (pr_delimited_binders (fun () -> str"forall" ++ spc()) + (pr mt ltop) bl) ++ + str "," ++ pr spc ltop a), + lprod + | CLambdaN _ -> + let (bl,a) = extract_lam_binders a in + hov 0 ( + hov 2 (pr_delimited_binders (fun () -> str"fun" ++ spc()) + (pr mt ltop) bl) ++ + + str " =>" ++ pr spc ltop a), + llambda + | CLetIn (_,(_,Name x),(CFix(_,(_,x'),[_])|CCoFix(_,(_,x'),[_]) as fx), b) + when x=x' -> + hv 0 ( + hov 2 (str "let " ++ pr mt ltop fx ++ str " in") ++ + pr spc ltop b), + lletin + | CLetIn (_,x,a,b) -> + hv 0 ( + hov 2 (str "let " ++ pr_lname x ++ str " :=" ++ + pr spc ltop a ++ str " in") ++ + pr spc ltop b), + lletin + | CAppExpl (_,(Some i,f),l) -> + let l1,l2 = list_chop i l in + let c,l1 = list_sep_last l1 in + let p = pr_proj (pr mt) pr_appexpl c f l1 in + if l2<>[] then + p ++ prlist (pr spc (lapp,L)) l2, lapp + else + p, lproj + | CAppExpl (_,(None,Ident (_,var)),[t]) + | CApp (_,(_,CRef(Ident(_,var))),[t,None]) + when var = Topconstr.ldots_var -> + hov 0 (str ".." ++ pr spc (latom,E) t ++ spc () ++ str ".."), larg + | CAppExpl (_,(None,f),l) -> pr_appexpl (pr mt) f l, lapp + | CApp (_,(Some i,f),l) -> + let l1,l2 = list_chop i l in + let c,l1 = list_sep_last l1 in + assert (snd c = None); + let p = pr_proj (pr mt) pr_app (fst c) f l1 in + if l2<>[] then + p ++ prlist (fun a -> spc () ++ pr_expl_args (pr mt) a) l2, lapp + else + p, lproj + | CApp (_,(None,a),l) -> pr_app (pr mt) a l, lapp + | CCases (_,rtntypopt,c,eqns) -> + v 0 + (hv 0 (str "match" ++ brk (1,2) ++ + hov 0 ( + prlist_with_sep sep_v + (pr_case_item (pr_dangling_with_for mt)) c + ++ pr_case_type (pr_dangling_with_for mt) rtntypopt) ++ + spc () ++ str "with") ++ + prlist (pr_eqn (pr mt)) eqns ++ spc() ++ str "end"), + latom + | CLetTuple (_,nal,(na,po),c,b) -> + hv 0 ( + str "let " ++ + hov 0 (str "(" ++ + prlist_with_sep sep_v pr_name nal ++ + str ")" ++ + pr_simple_return_type (pr mt) na po ++ str " :=" ++ + pr spc ltop c ++ str " in") ++ + pr spc ltop b), + lletin + | CIf (_,c,(na,po),b1,b2) -> + (* On force les parenthses autour d'un "if" sous-terme (mme si le + parsing est lui plus tolrant) *) + hv 0 ( + hov 1 (str "if " ++ pr mt ltop c ++ pr_simple_return_type (pr mt) na po) ++ + spc () ++ + hov 0 (str "then" ++ pr (fun () -> brk (1,1)) ltop b1) ++ spc () ++ + hov 0 (str "else" ++ pr (fun () -> brk (1,1)) ltop b2)), + lif + + | CHole _ -> str "_", latom + | CEvar (_,n) -> str (Evd.string_of_existential n), latom + | CPatVar (_,(_,p)) -> str "?" ++ pr_patvar p, latom + | CSort (_,s) -> pr_sort s, latom + | CCast (_,a,_,b) -> + hv 0 (pr mt (lcast,L) a ++ cut () ++ str ":" ++ pr mt (-lcast,E) b), + lcast + | CNotation (_,"( _ )",[t]) -> + pr (fun()->str"(") (max_int,L) t ++ str")", latom + | CNotation (_,s,env) -> pr_notation (pr mt) s env + | CNumeral (_,p) -> + Bigint.pr_bigint p, + (if Bigint.is_pos_or_zero p then lposint else lnegint) + | CDelimiters (_,sc,a) -> pr_delimiters sc (pr mt lsimple a), 1 + | CDynamic _ -> str "<dynamic>", latom + in + let loc = constr_loc a in + pr_with_comments loc + (sep() ++ if prec_less prec inherited then strm else surround strm) + +and pr_dangling_with_for sep inherited a = + match a with + | (CFix (_,_,[_])|CCoFix(_,_,[_])) -> pr sep (latom,E) a + | _ -> pr sep inherited a + +let pr = pr mt + +let rec abstract_constr_expr c = function + | [] -> c + | LocalRawDef (x,b)::bl -> mkLetInC(x,b,abstract_constr_expr c bl) + | LocalRawAssum (idl,t)::bl -> + List.fold_right (fun x b -> mkLambdaC([x],t,b)) idl + (abstract_constr_expr c bl) + +let rec prod_constr_expr c = function + | [] -> c + | LocalRawDef (x,b)::bl -> mkLetInC(x,b,prod_constr_expr c bl) + | LocalRawAssum (idl,t)::bl -> + List.fold_right (fun x b -> mkProdC([x],t,b)) idl + (prod_constr_expr c bl) + +let rec strip_context n iscast t = + if n = 0 then + [], if iscast then match t with CCast (_,c,_,_) -> c | _ -> t else t + else match t with + | CLambdaN (loc,(nal,t)::bll,c) -> + let n' = List.length nal in + if n' > n then + let nal1,nal2 = list_chop n nal in + [LocalRawAssum (nal1,t)], CLambdaN (loc,(nal2,t)::bll,c) + else + let bl', c = strip_context (n-n') iscast + (if bll=[] then c else CLambdaN (loc,bll,c)) in + LocalRawAssum (nal,t) :: bl', c + | CProdN (loc,(nal,t)::bll,c) -> + let n' = List.length nal in + if n' > n then + let nal1,nal2 = list_chop n nal in + [LocalRawAssum (nal1,t)], CProdN (loc,(nal2,t)::bll,c) + else + let bl', c = strip_context (n-n') iscast + (if bll=[] then c else CProdN (loc,bll,c)) in + LocalRawAssum (nal,t) :: bl', c + | CArrow (loc,t,c) -> + let bl', c = strip_context (n-1) iscast c in + LocalRawAssum ([loc,Anonymous],t) :: bl', c + | CCast (_,c,_,_) -> strip_context n false c + | CLetIn (_,na,b,c) -> + let bl', c = strip_context (n-1) iscast c in + LocalRawDef (na,b) :: bl', c + | _ -> anomaly "ppconstrnew: strip_context" + +let pr_constr_env env c = pr lsimple c +let pr_lconstr_env env c = pr ltop c +let pr_constr c = pr_constr_env (Global.env()) c +let pr_lconstr c = pr_lconstr_env (Global.env()) c + +let pr_binders = pr_undelimited_binders (pr ltop) + +let pr_lconstr_env_n env iscast bl c = bl, pr ltop c +let pr_type c = pr ltop c + +let transf_pattern env c = + if Options.do_translate() then + Constrextern.extern_rawconstr (Termops.vars_of_env env) + (Constrintern.for_grammar + (Constrintern.intern_gen false ~allow_soapp:true Evd.empty env) c) + else c + +let pr_pattern c = pr lsimple (transf_pattern (Global.env()) c) + +let pr_rawconstr_env env c = + pr_constr (Constrextern.extern_rawconstr (Termops.vars_of_env env) c) +let pr_lrawconstr_env env c = + pr_lconstr (Constrextern.extern_rawconstr (Termops.vars_of_env env) c) + +let pr_cases_pattern = pr_patt ltop + +let pr_pattern_occ prc = function + ([],c) -> prc c + | (nl,c) -> hov 1 (prc c ++ spc() ++ str"at " ++ + hov 0 (prlist_with_sep spc int nl)) + +let pr_unfold_occ pr_ref = function + ([],qid) -> pr_ref qid + | (nl,qid) -> hov 1 (pr_ref qid ++ spc() ++ str"at " ++ + hov 0 (prlist_with_sep spc int nl)) + +let pr_qualid qid = str (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_metaid id = str"?" ++ pr_id id + +let pr_red_expr (pr_constr,pr_lconstr,pr_ref) = function + | Red false -> str "red" + | Hnf -> str "hnf" + | Simpl o -> str "simpl" ++ pr_opt (pr_pattern_occ pr_constr) o + | Cbv f -> + if f = {rBeta=true;rIota=true;rZeta=true;rDelta=true;rConst=[]} then + str "compute" + else + hov 1 (str "cbv" ++ pr_red_flag pr_ref f) + | Lazy f -> + hov 1 (str "lazy" ++ pr_red_flag pr_ref f) + | Unfold l -> + hov 1 (str "unfold" ++ spc() ++ + prlist_with_sep pr_coma (pr_unfold_occ pr_ref) l) + | Fold l -> hov 1 (str "fold" ++ prlist (pr_arg pr_constr) l) + | Pattern l -> + hov 1 (str "pattern" ++ + pr_arg (prlist_with_sep pr_coma (pr_pattern_occ pr_constr)) l) + + | Red true -> error "Shouldn't be accessible from user" + | ExtraRedExpr s -> str s + | CbvVm -> str "vm_compute" + +let rec pr_may_eval test prc prlc pr2 = function + | ConstrEval (r,c) -> + hov 0 + (str "eval" ++ brk (1,1) ++ + pr_red_expr (prc,prlc,pr2) r ++ + str " in" ++ spc() ++ prc c) + | ConstrContext ((_,id),c) -> + hov 0 + (str "context " ++ pr_id id ++ spc () ++ + str "[" ++ prlc c ++ str "]") + | ConstrTypeOf c -> hov 1 (str "type of" ++ spc() ++ prc c) + | ConstrTerm c when test c -> h 0 (str "(" ++ prc c ++ str ")") + | ConstrTerm c -> prc c + +let pr_may_eval a = pr_may_eval (fun _ -> false) a + +(** constr printers *) + +let pr_term_env env c = pr lsimple (Constrextern.extern_constr false env c) +let pr_lterm_env env c = pr ltop (Constrextern.extern_constr false env c) +let pr_term c = pr_term_env (Global.env()) c +let pr_lterm c = pr_lterm_env (Global.env()) c + +let pr_constr_pattern_env env c = + pr lsimple (Constrextern.extern_pattern env Termops.empty_names_context c) + +let pr_constr_pattern t = + pr lsimple + (Constrextern.extern_pattern (Global.env()) Termops.empty_names_context t) diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli index 5bbeecc2c..e6af0e369 100644 --- a/parsing/ppconstr.mli +++ b/parsing/ppconstr.mli @@ -1,3 +1,4 @@ + (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) @@ -5,6 +6,82 @@ (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) - + (*i $Id$ i*) +open Pp +open Environ +open Term +open Libnames +open Pcoq +open Rawterm +open Topconstr +open Names +open Util +open Genarg + +val extract_lam_binders : + constr_expr -> local_binder list * constr_expr +val extract_prod_binders : + constr_expr -> local_binder list * constr_expr +val extract_def_binders : + constr_expr -> constr_expr -> + local_binder list * constr_expr * constr_expr +val split_fix : + int -> constr_expr -> constr_expr -> + local_binder list * constr_expr * constr_expr +val pr_binders : local_binder list -> std_ppcmds + +val prec_less : int -> int * Ppextend.parenRelation -> bool + +val pr_global : Idset.t -> global_reference -> std_ppcmds + +val pr_tight_coma : unit -> std_ppcmds +val pr_located : + ('a -> std_ppcmds) -> 'a located -> std_ppcmds +val pr_lident : identifier located -> std_ppcmds +val pr_lname : name located -> std_ppcmds + +val pr_with_comments : loc -> std_ppcmds -> std_ppcmds +val pr_com_at : int -> std_ppcmds +val pr_sep_com : + (unit -> std_ppcmds) -> + (constr_expr -> std_ppcmds) -> + constr_expr -> std_ppcmds +val pr_opt : ('a -> std_ppcmds) -> 'a option -> std_ppcmds +val pr_id : identifier -> std_ppcmds +val pr_name : name -> std_ppcmds +val pr_qualid : qualid -> std_ppcmds +val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds +val pr_metaid : identifier -> std_ppcmds +val pr_red_expr : + ('a -> std_ppcmds) * ('a -> std_ppcmds) * ('b -> std_ppcmds) -> + ('a,'b) red_expr_gen -> std_ppcmds + +val pr_sort : rawsort -> std_ppcmds +val pr_pattern : Tacexpr.pattern_expr -> std_ppcmds +val pr_constr : constr_expr -> std_ppcmds +val pr_lconstr : constr_expr -> std_ppcmds +val pr_constr_env : env -> constr_expr -> std_ppcmds +val pr_lconstr_env : env -> constr_expr -> std_ppcmds +val pr_type : constr_expr -> std_ppcmds +val pr_cases_pattern : cases_pattern_expr -> std_ppcmds +val pr_may_eval : + ('a -> std_ppcmds) -> ('a -> std_ppcmds) -> ('b -> std_ppcmds) -> ('a,'b) may_eval + -> std_ppcmds +val abstract_constr_expr : constr_expr -> local_binder list -> constr_expr +val prod_constr_expr : constr_expr -> local_binder list -> constr_expr + + +val pr_rawconstr_env : env -> rawconstr -> std_ppcmds +val pr_lrawconstr_env : env -> rawconstr -> std_ppcmds + +(** constr printers *) + +val pr_term_env : env -> constr -> std_ppcmds +val pr_lterm_env : env -> constr -> std_ppcmds +val pr_term : constr -> std_ppcmds +val pr_lterm : constr -> std_ppcmds + +val pr_constr_pattern_env : env -> Pattern.constr_pattern -> std_ppcmds +val pr_constr_pattern : Pattern.constr_pattern -> std_ppcmds diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml index 8360e280a..4609d4b2e 100644 --- a/parsing/pptactic.ml +++ b/parsing/pptactic.ml @@ -19,12 +19,9 @@ open Genarg open Libnames open Pattern open Ppextend +open Ppconstr -let pr_red_expr = Ppconstrnew.pr_red_expr -let pr_may_eval = Ppconstrnew.pr_may_eval -let pr_sort = Ppconstrnew.pr_sort let pr_global x = Nametab.pr_global_env Idset.empty x -let pr_opt = Ppconstrnew.pr_opt type grammar_terminals = string option list @@ -273,8 +270,730 @@ let pr_extend_gen prgen lev s l = str s ++ prlist prgen l ++ str " (* Generic printer *)" let pr_raw_extend prc prlc prtac = - pr_extend_gen (pr_raw_generic prc prlc prtac Ppconstrnew.pr_reference) + pr_extend_gen (pr_raw_generic prc prlc prtac pr_reference) let pr_glob_extend prc prlc prtac = pr_extend_gen (pr_glob_generic prc prlc prtac) let pr_extend prc prlc prtac = pr_extend_gen (pr_generic prc prlc prtac) + +(**********************************************************************) +(* The tactic printer *) + +let sep_v = fun _ -> str"," ++ spc() + +let strip_prod_binders_expr n ty = + let rec strip_ty acc n ty = + match ty with + Topconstr.CProdN(_,bll,a) -> + let nb = + List.fold_left (fun i (nal,_) -> i + List.length nal) 0 bll in + if nb >= n then (List.rev (bll@acc), a) + else strip_ty (bll@acc) (n-nb) a + | Topconstr.CArrow(_,a,b) -> + if n=1 then + (List.rev (([(dummy_loc,Anonymous)],a)::acc), b) + else strip_ty (([(dummy_loc,Anonymous)],a)::acc) (n-1) b + | _ -> error "Cannot translate fix tactic: not enough products" in + strip_ty [] n ty + + +(* In new syntax only double quote char is escaped by repeating it *) +let rec escape_string s = + let rec escape_at s i = + if i<0 then s + else if s.[i] == '"' then + let s' = String.sub s 0 i^"\""^String.sub s i (String.length s - i) in + escape_at s' (i-1) + else escape_at s (i-1) in + escape_at s (String.length s - 1) + +let qstring s = str ("\""^escape_string s^"\"") +let qsnew = qstring + +let pr_ltac_or_var pr = function + | ArgArg x -> pr x + | ArgVar (loc,id) -> pr_with_comments loc (pr_id id) + +let pr_arg pr x = spc () ++ pr x + +let pr_ltac_constant sp = + pr_qualid (Nametab.shortest_qualid_of_tactic sp) + +let pr_evaluable_reference_env env = function + | EvalVarRef id -> pr_id id + | EvalConstRef sp -> + Nametab.pr_global_env (Termops.vars_of_env env) (Libnames.ConstRef sp) + +let pr_inductive env ind = + Nametab.pr_global_env (Termops.vars_of_env env) (Libnames.IndRef ind) + +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_esubst prc l = + let pr_qhyp = function + (_,AnonHyp n,c) -> str "(" ++ int n ++ str" := " ++ prc c ++ str ")" + | (_,NamedHyp id,c) -> + str "(" ++ pr_id id ++ str" := " ++ prc c ++ str ")" + in + prlist_with_sep spc pr_qhyp l + +let pr_bindings_gen for_ex prlc prc = function + | ImplicitBindings l -> + spc () ++ + hv 2 ((if for_ex then mt() else str "with" ++ spc ()) ++ + prlist_with_sep spc prc l) + | ExplicitBindings l -> + spc () ++ + hv 2 ((if for_ex then mt() else str "with" ++ spc ()) ++ + pr_esubst prlc l) + | NoBindings -> mt () + +let pr_bindings prlc prc = pr_bindings_gen false prlc prc + +let pr_with_bindings prlc prc (c,bl) = + hov 1 (prc c ++ pr_bindings prlc prc bl) + +let pr_with_constr prc = function + | None -> mt () + | Some c -> spc () ++ hov 1 (str "with" ++ spc () ++ prc c) + +(* Translator copy of pr_intro_pattern based on a translating "pr_id" *) +let rec pr_intro_pattern = function + | IntroOrAndPattern pll -> pr_case_intro_pattern pll + | IntroWildcard -> str "_" + | IntroIdentifier id -> pr_id id +and pr_case_intro_pattern = function + | [_::_ as pl] -> + str "(" ++ hov 0 (prlist_with_sep pr_coma pr_intro_pattern pl) ++ str ")" + | pll -> + str "[" ++ + hv 0 (prlist_with_sep pr_bar + (fun l -> hov 0 (prlist_with_sep spc pr_intro_pattern l)) pll) + ++ str "]" + +let pr_with_names = function + | None -> mt () + | Some ipat -> spc () ++ hov 1 (str "as" ++ spc () ++ pr_intro_pattern ipat) + +let pr_occs pp = function + [] -> pp + | nl -> hov 1 (pp ++ spc() ++ str"at " ++ + hov 0 (prlist_with_sep spc int nl)) + +let pr_hyp_location pr_id = function + | id, occs, InHyp -> spc () ++ pr_occs (pr_id id) occs + | id, occs, InHypTypeOnly -> + spc () ++ pr_occs (str "(type of " ++ pr_id id ++ str ")") occs + | id, occs, InHypValueOnly -> + spc () ++ pr_occs (str "(value of " ++ pr_id id ++ str ")") occs + +let pr_in pp = spc () ++ hov 0 (str "in" ++ pp) + +let pr_simple_clause pr_id = function + | [] -> mt () + | l -> pr_in (spc () ++ prlist_with_sep spc pr_id l) + +let pr_clauses pr_id = function + { onhyps=None; onconcl=true; concl_occs=nl } -> + pr_in (pr_occs (str " *") nl) + | { onhyps=None; onconcl=false } -> pr_in (str " * |-") + | { onhyps=Some l; onconcl=true; concl_occs=nl } -> + pr_in (prlist_with_sep (fun () -> str",") (pr_hyp_location pr_id) l + ++ pr_occs (str" |- *") nl) + | { onhyps=Some l; onconcl=false } -> + pr_in (prlist_with_sep (fun()->str",") (pr_hyp_location pr_id) l) + +let pr_clause_pattern pr_id = function + | (None, []) -> mt () + | (glopt,l) -> + str " in" ++ + prlist + (fun (id,nl) -> prlist (pr_arg int) nl + ++ spc () ++ pr_id id) l ++ + pr_opt (fun nl -> prlist_with_sep spc int nl ++ str " Goal") glopt + +let pr_induction_arg prc = function + | ElimOnConstr c -> prc c + | ElimOnIdent (loc,id) -> pr_with_comments loc (pr_id id) + | ElimOnAnonHyp n -> int n + +let pr_induction_kind = function + | SimpleInversion -> str "simple inversion" + | FullInversion -> str "inversion" + | FullInversionClear -> str "inversion_clear" + +let pr_lazy lz = if lz then str "lazy " else mt () + +let pr_match_pattern pr_pat = function + | Term a -> pr_pat a + | Subterm (None,a) -> str "context [" ++ pr_pat a ++ str "]" + | Subterm (Some id,a) -> + str "context " ++ pr_id id ++ str "[" ++ pr_pat a ++ str "]" + +let pr_match_hyps pr_pat = function + | Hyp (nal,mp) -> pr_lname nal ++ str ":" ++ pr_match_pattern pr_pat mp + +let pr_match_rule m pr pr_pat = function + | Pat ([],mp,t) when m -> + pr_match_pattern pr_pat mp ++ + spc () ++ str "=>" ++ brk (1,4) ++ pr t + | Pat (rl,mp,t) -> + prlist_with_sep (fun () -> str",") (pr_match_hyps pr_pat) rl ++ + spc () ++ str "|-" ++ spc () ++ pr_match_pattern pr_pat mp ++ spc () ++ + str "=>" ++ brk (1,4) ++ pr t + | All t -> str "_" ++ spc () ++ str "=>" ++ brk (1,4) ++ 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) -> + hov 0 (str k ++ pr_lident id ++ str " :=" ++ brk (1,1) ++ + pr (TacArg t)) + | (id,Some c,t) -> + hv 0 (str k ++ pr_lident id ++ str" :" ++ brk(1,2) ++ + pr c ++ + str " :=" ++ brk (1,1) ++ pr (TacArg t)) + +let pr_let_clauses pr = function + | hd::tl -> + hv 0 + (pr_let_clause "let " pr hd ++ + prlist (fun t -> spc () ++ pr_let_clause "with " pr t) tl) + | [] -> anomaly "LetIn must declare at least one binding" + +let pr_rec_clause pr (id,(l,t)) = + hov 0 + (pr_lident id ++ prlist pr_funvar l ++ str " :=") ++ spc () ++ pr t + +let pr_rec_clauses pr l = + prlist_with_sep (fun () -> fnl () ++ str "with ") (pr_rec_clause pr) l + +let pr_seq_body pr tl = + hv 0 (str "[ " ++ + prlist_with_sep (fun () -> spc () ++ str "| ") pr tl ++ + str " ]") + +let pr_hintbases = function + | None -> spc () ++ str "with *" + | Some [] -> mt () + | Some l -> + spc () ++ hov 2 (str "with" ++ prlist (fun s -> spc () ++ str s) l) + +let pr_autoarg_adding = function + | [] -> mt () + | l -> + spc () ++ str "adding [" ++ + hv 0 (prlist_with_sep spc pr_reference 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_tacarg_using_rule pr_gen = function + | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_tacarg_using_rule pr_gen (l,al) + | Egrammar.TacNonTerm _ :: l, a :: al -> pr_gen a ++ pr_tacarg_using_rule pr_gen (l,al) + | [], [] -> mt () + | _ -> failwith "Inconsistent arguments of extended tactic" + +let pr_then () = str ";" + +let ltop = (5,E) +let lseq = 5 +let ltactical = 3 +let lorelse = 2 +let llet = 1 +let lfun = 1 +let labstract = 3 +let lmatch = 1 +let latom = 0 +let lcall = 1 +let leval = 1 +let ltatom = 1 + +let level_of (n,p) = match p with E -> n | L -> n-1 | Prec n -> n | Any -> lseq + +open Closure + +let make_pr_tac + (pr_tac_level,pr_constr,pr_lconstr,pr_pat, + pr_cst,pr_ind,pr_ref,pr_ident, + pr_extend,strip_prod_binders) = + +let pr_bindings env = + pr_bindings (pr_lconstr env) (pr_constr env) in +let pr_ex_bindings env = + pr_bindings_gen true (pr_lconstr env) (pr_constr env) in +let pr_with_bindings env = + pr_with_bindings (pr_lconstr env) (pr_constr env) in +let pr_eliminator env cb = + str "using" ++ pr_arg (pr_with_bindings env) cb in +let pr_extend env = + pr_extend (pr_constr env) (pr_lconstr env) (pr_tac_level env) in +let pr_red_expr env = + pr_red_expr (pr_constr env,pr_lconstr env,pr_cst env) in + +let pr_constrarg env c = spc () ++ pr_constr env c in +let pr_lconstrarg env c = spc () ++ pr_lconstr env c in +let pr_intarg n = spc () ++ int n in + +let pr_binder_fix env (nal,t) = +(* match t with + | CHole _ -> spc() ++ prlist_with_sep spc (pr_lname) nal + | _ ->*) + let s = + prlist_with_sep spc (pr_lname) nal ++ str ":" ++ + pr_lconstr env t in + spc() ++ hov 1 (str"(" ++ s ++ str")") in + +let pr_fix_tac env (id,n,c) = + let rec set_nth_name avoid n = function + (nal,ty)::bll -> + if n <= List.length nal then + match list_chop (n-1) nal with + _, (_,Name id) :: _ -> id, (nal,ty)::bll + | bef, (loc,Anonymous) :: aft -> + let id = next_ident_away_from (id_of_string"y") avoid in + id, ((bef@(loc,Name id)::aft, ty)::bll) + | _ -> assert false + else + let (id,bll') = set_nth_name avoid (n-List.length nal) bll in + (id,(nal,ty)::bll') + | [] -> assert false in + let (bll,ty) = strip_prod_binders n c in + let names = + List.fold_left + (fun ln (nal,_) -> List.fold_left + (fun ln na -> match na with (_,Name id) -> id::ln | _ -> ln) + ln nal) + [] bll in + let idarg,bll = set_nth_name names n bll in + let annot = + if List.length names = 1 then mt() + else spc() ++ str"{struct " ++ pr_id idarg ++ str"}" in + hov 1 (str"(" ++ pr_id id ++ + prlist (pr_binder_fix env) bll ++ annot ++ str" :" ++ + pr_lconstrarg env ty ++ str")") in +(* spc() ++ + hov 0 (pr_id id ++ pr_intarg n ++ str":" ++ pr_constrarg + env c) +*) +let pr_cofix_tac env (id,c) = + hov 1 (str"(" ++ pr_id id ++ str" :" ++ pr_lconstrarg env c ++ str")") in + + + (* Printing tactics as arguments *) +let rec pr_atom0 env = 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" + | TacReflexivity -> str "reflexivity" + | t -> str "(" ++ pr_atom1 env t ++ str ")" + + (* Main tactic printer *) +and pr_atom1 env = function + | TacAutoTDB _ | TacDestructHyp _ | TacDestructConcl + | TacSuperAuto _ | TacExtend (_, + ("GTauto"|"GIntuition"|"TSimplif"| + "LinearIntuition"),_) -> + errorlabstrm "Obsolete V8" (str "Tactic is not ported to V8.0") + | TacExtend (loc,s,l) -> + pr_with_comments loc (pr_extend env 1 s l) + | TacAlias (loc,s,l,_) -> + pr_with_comments loc (pr_extend env 1 s (List.map snd l)) + + (* Basic tactics *) + | TacIntroPattern [] as t -> pr_atom0 env t + | TacIntroPattern (_::_ as p) -> + hov 1 (str "intros" ++ 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 env 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_lident id2) + | TacAssumption as t -> pr_atom0 env t + | TacExact c -> hov 1 (str "exact" ++ pr_constrarg env c) + | TacExactNoCheck c -> hov 1 (str "exact_no_check" ++ pr_constrarg env c) + | TacApply cb -> hov 1 (str "apply" ++ spc () ++ pr_with_bindings env cb) + | TacElim (cb,cbo) -> + hov 1 (str "elim" ++ pr_arg (pr_with_bindings env) cb ++ + pr_opt (pr_eliminator env) cbo) + | TacElimType c -> hov 1 (str "elimtype" ++ pr_constrarg env c) + | TacCase cb -> hov 1 (str "case" ++ spc () ++ pr_with_bindings env cb) + | TacCaseType c -> hov 1 (str "casetype" ++ pr_constrarg env c) + | TacFix (ido,n) -> hov 1 (str "fix" ++ pr_opt pr_id ido ++ pr_intarg n) + | TacMutualFix (id,n,l) -> + hov 1 (str "fix" ++ spc () ++ pr_id id ++ pr_intarg n ++ spc() ++ + str"with " ++ prlist_with_sep spc (pr_fix_tac env) l) + | TacCofix ido -> hov 1 (str "cofix" ++ pr_opt pr_id ido) + | TacMutualCofix (id,l) -> + hov 1 (str "cofix" ++ spc () ++ pr_id id ++ spc() ++ + str"with " ++ prlist_with_sep spc (pr_cofix_tac env) l) + | TacCut c -> hov 1 (str "cut" ++ pr_constrarg env c) + | TacTrueCut (Anonymous,c) -> + hov 1 (str "assert" ++ pr_constrarg env c) + | TacTrueCut (Name id,c) -> + hov 1 (str "assert" ++ spc () ++ + hov 1 (str"(" ++ pr_id id ++ str " :" ++ + pr_lconstrarg env c ++ str")")) + | TacForward (false,na,c) -> + hov 1 (str "assert" ++ spc () ++ + hov 1 (str"(" ++ pr_name na ++ str " :=" ++ + pr_lconstrarg env c ++ str")")) + | TacForward (true,Anonymous,c) -> + hov 1 (str "pose" ++ pr_constrarg env c) + | TacForward (true,Name id,c) -> + hov 1 (str "pose" ++ spc() ++ + hov 1 (str"(" ++ pr_id id ++ str " :=" ++ + pr_lconstrarg env c ++ str")")) + | TacGeneralize l -> + hov 1 (str "generalize" ++ spc () ++ + prlist_with_sep spc (pr_constr env) l) + | TacGeneralizeDep c -> + hov 1 (str "generalize" ++ spc () ++ str "dependent" ++ + pr_constrarg env c) + | TacLetTac (Anonymous,c,cl) -> + hov 1 (str "set" ++ pr_constrarg env c) ++ pr_clauses pr_ident cl + | TacLetTac (Name id,c,cl) -> + hov 1 (str "set" ++ spc () ++ + hov 1 (str"(" ++ pr_id id ++ str " :=" ++ + pr_lconstrarg env c ++ str")") ++ + pr_clauses pr_ident cl) +(* | TacInstantiate (n,c,ConclLocation ()) -> + hov 1 (str "instantiate" ++ spc() ++ + hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++ + pr_lconstrarg env c ++ str ")" )) + | TacInstantiate (n,c,HypLocation (id,hloc)) -> + hov 1 (str "instantiate" ++ spc() ++ + hov 1 (str"(" ++ pr_arg int n ++ str" :=" ++ + pr_lconstrarg env c ++ str ")" ) + ++ str "in" ++ pr_hyp_location pr_ident (id,[],(hloc,ref None))) +*) + (* Derived basic tactics *) + | TacSimpleInduction h -> + hov 1 (str "simple induction" ++ pr_arg pr_quantified_hypothesis h) + | TacNewInduction (h,e,ids) -> + hov 1 (str "induction" ++ spc () ++ + pr_induction_arg (pr_constr env) h ++ pr_with_names ids ++ + pr_opt (pr_eliminator env) e) + | TacSimpleDestruct h -> + hov 1 (str "simple destruct" ++ pr_arg pr_quantified_hypothesis h) + | TacNewDestruct (h,e,ids) -> + hov 1 (str "destruct" ++ spc () ++ + pr_induction_arg (pr_constr env) h ++ pr_with_names ids ++ + pr_opt (pr_eliminator env) e) + | 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_constrarg env c) + | TacDecomposeOr c -> + hov 1 (str "decompose sum" ++ pr_constrarg env c) + | TacDecompose (l,c) -> + hov 1 (str "decompose" ++ spc () ++ + hov 0 (str "[" ++ prlist_with_sep spc (pr_ind env) l + ++ str "]" ++ pr_constrarg env c)) + | TacSpecialize (n,c) -> + hov 1 (str "specialize" ++ spc () ++ pr_opt int n ++ + pr_with_bindings env c) + | TacLApply c -> + hov 1 (str "lapply" ++ pr_constrarg env c) + + (* Automation tactics *) + | TacTrivial (Some []) as x -> pr_atom0 env x + | TacTrivial db -> hov 0 (str "trivial" ++ pr_hintbases db) + | TacAuto (None,Some []) as x -> pr_atom0 env x + | TacAuto (n,db) -> + hov 0 (str "auto" ++ pr_opt (pr_or_var int) n ++ pr_hintbases db) + | TacDAuto (n,p) -> + hov 1 (str "auto" ++ pr_opt (pr_or_var int) n ++ str "decomp" ++ pr_opt int p) + + (* Context management *) + | TacClear (keep,l) -> + hov 1 (str "clear" ++ spc () ++ (if keep then str "- " else mt ()) ++ + prlist_with_sep spc pr_ident l) + | TacClearBody l -> + hov 1 (str "clearbody" ++ spc () ++ prlist_with_sep spc pr_ident l) + | TacMove (b,id1,id2) -> + (* Rem: only b = true is available for users *) + assert b; + hov 1 + (str "move" ++ brk (1,1) ++ pr_ident id1 ++ spc () ++ + str "after" ++ brk (1,1) ++ pr_ident id2) + | TacRename (id1,id2) -> + hov 1 + (str "rename" ++ brk (1,1) ++ pr_ident id1 ++ spc () ++ + str "into" ++ brk (1,1) ++ pr_ident id2) + + (* Constructors *) + | TacLeft l -> hov 1 (str "left" ++ pr_bindings env l) + | TacRight l -> hov 1 (str "right" ++ pr_bindings env l) + | TacSplit (false,l) -> hov 1 (str "split" ++ pr_bindings env l) + | TacSplit (true,l) -> hov 1 (str "exists" ++ pr_ex_bindings env l) + | TacAnyConstructor (Some t) -> + hov 1 (str "constructor" ++ pr_arg (pr_tac_level env (latom,E)) t) + | TacAnyConstructor None as t -> pr_atom0 env t + | TacConstructor (n,l) -> + hov 1 (str "constructor" ++ pr_or_metaid pr_intarg n ++ pr_bindings env l) + + (* Conversion *) + | TacReduce (r,h) -> + hov 1 (pr_red_expr env r ++ + pr_clauses pr_ident h) + | TacChange (occ,c,h) -> + hov 1 (str "change" ++ brk (1,1) ++ + (match occ with + None -> mt() + | Some([],c1) -> hov 1 (pr_constr env c1 ++ spc() ++ str "with ") + | Some(ocl,c1) -> + hov 1 (pr_constr env c1 ++ spc() ++ + str "at " ++ prlist_with_sep spc int ocl) ++ spc() ++ + str "with ") ++ + pr_constr env c ++ pr_clauses pr_ident h) + + (* Equivalence relations *) + | TacReflexivity as x -> pr_atom0 env x + | TacSymmetry cls -> str "symmetry " ++ pr_clauses pr_ident cls + | TacTransitivity c -> str "transitivity" ++ pr_constrarg env c + + (* Equality and inversion *) + | TacInversion (DepInversion (k,c,ids),hyp) -> + hov 1 (str "dependent " ++ pr_induction_kind k ++ spc () ++ + pr_quantified_hypothesis hyp ++ + pr_with_names ids ++ pr_with_constr (pr_constr env) c) + | TacInversion (NonDepInversion (k,cl,ids),hyp) -> + hov 1 (pr_induction_kind k ++ spc () ++ + pr_quantified_hypothesis hyp ++ + pr_with_names ids ++ pr_simple_clause pr_ident cl) + | TacInversion (InversionUsing (c,cl),hyp) -> + hov 1 (str "inversion" ++ spc() ++ pr_quantified_hypothesis hyp ++ + spc () ++ str "using" ++ spc () ++ pr_constr env c ++ + pr_simple_clause pr_ident cl) + +in + +let rec pr_tac env inherited tac = + let (strm,prec) = match tac with + | TacAbstract (t,None) -> + str "abstract " ++ pr_tac env (labstract,L) t, labstract + | TacAbstract (t,Some s) -> + hov 0 + (str "abstract (" ++ pr_tac env (labstract,L) t ++ str")" ++ spc () ++ + str "using " ++ pr_id s), + labstract + | TacLetRecIn (l,t) -> + hv 0 + (str "let rec " ++ pr_rec_clauses (pr_tac env ltop) l ++ str " in" ++ + fnl () ++ pr_tac env (llet,E) t), + llet + | TacLetIn (llc,u) -> + v 0 + (hv 0 (pr_let_clauses (pr_tac env ltop) llc + ++ str " in") ++ + fnl () ++ pr_tac env (llet,E) u), + llet + | TacMatch (lz,t,lrul) -> + hov 0 (pr_lazy lz ++ str "match " ++ pr_tac env ltop t ++ str " with" + ++ prlist + (fun r -> fnl () ++ str "| " ++ + pr_match_rule true (pr_tac env ltop) pr_pat r) + lrul + ++ fnl() ++ str "end"), + lmatch + | TacMatchContext (lz,lr,lrul) -> + hov 0 (pr_lazy lz ++ + str (if lr then "match reverse goal with" else "match goal with") + ++ prlist + (fun r -> fnl () ++ str "| " ++ + pr_match_rule false (pr_tac env ltop) pr_pat r) + lrul + ++ fnl() ++ str "end"), + lmatch + | TacFun (lvar,body) -> +(* let env = List.fold_right (option_fold_right Idset.add) lvar env in*) + hov 2 (str "fun" ++ + prlist pr_funvar lvar ++ str " =>" ++ spc () ++ + pr_tac env (lfun,E) body), + lfun + | TacThens (t,tl) -> + hov 1 (pr_tac env (lseq,E) t ++ pr_then () ++ spc () ++ + pr_seq_body (pr_tac env ltop) tl), + lseq + | TacThen (t1,t2) -> + hov 1 (pr_tac env (lseq,E) t1 ++ pr_then () ++ spc () ++ + pr_tac env (lseq,L) t2), + lseq + | TacTry t -> + hov 1 (str "try" ++ spc () ++ pr_tac env (ltactical,E) t), + ltactical + | TacDo (n,t) -> + hov 1 (str "do " ++ pr_or_var int n ++ spc () ++ + pr_tac env (ltactical,E) t), + ltactical + | TacRepeat t -> + hov 1 (str "repeat" ++ spc () ++ pr_tac env (ltactical,E) t), + ltactical + | TacProgress t -> + hov 1 (str "progress" ++ spc () ++ pr_tac env (ltactical,E) t), + ltactical + | TacInfo t -> + hov 1 (str "info" ++ spc () ++ pr_tac env (ltactical,E) t), + ltactical + | TacOrelse (t1,t2) -> + hov 1 (pr_tac env (lorelse,L) t1 ++ str " ||" ++ brk (1,1) ++ + pr_tac env (lorelse,E) t2), + lorelse + | TacFail (ArgArg 0,"") -> str "fail", latom + | TacFail (n,s) -> + str "fail" ++ (if n=ArgArg 0 then mt () else pr_arg (pr_or_var int) n) ++ + (if s="" then mt() else (spc() ++ qstring s)), latom + | TacFirst tl -> + str "first" ++ spc () ++ pr_seq_body (pr_tac env ltop) tl, llet + | TacSolve tl -> + str "solve" ++ spc () ++ pr_seq_body (pr_tac env ltop) tl, llet + | TacId "" -> str "idtac", latom + | TacId s -> str "idtac" ++ (qstring s), latom + | TacAtom (loc,TacAlias (_,s,l,_)) -> + pr_with_comments loc + (pr_extend env (level_of inherited) s (List.map snd l)), + latom + | TacAtom (loc,t) -> + pr_with_comments loc (hov 1 (pr_atom1 env t)), ltatom + | TacArg(Tacexp e) -> pr_tac_level env (latom,E) e, latom + | TacArg(ConstrMayEval (ConstrTerm c)) -> + str "constr:" ++ pr_constr env c, latom + | TacArg(ConstrMayEval c) -> + pr_may_eval (pr_constr env) (pr_lconstr env) (pr_cst env) c, leval + | TacArg(TacFreshId sopt) -> str "fresh" ++ pr_opt qstring sopt, latom + | TacArg(Integer n) -> int n, latom + | TacArg(TacCall(loc,f,l)) -> + pr_with_comments loc + (hov 1 (pr_ref f ++ spc () ++ + prlist_with_sep spc (pr_tacarg env) l)), + lcall + | TacArg a -> pr_tacarg env a, latom + in + if prec_less prec inherited then strm + else str"(" ++ strm ++ str")" + +and pr_tacarg env = function + | TacDynamic (loc,t) -> + pr_with_comments loc (str ("<dynamic ["^(Dyn.tag t)^"]>")) + | MetaIdArg (loc,s) -> pr_with_comments loc (str ("$" ^ s)) + | IntroPattern ipat -> str "ipattern:" ++ pr_intro_pattern ipat + | TacVoid -> str "()" + | Reference r -> pr_ref r + | ConstrMayEval c -> + pr_may_eval (pr_constr env) (pr_lconstr env) (pr_cst env) c + | TacFreshId sopt -> str "fresh" ++ pr_opt qstring sopt + | TacExternal (_,com,req,la) -> + str "external" ++ spc() ++ qstring com ++ spc() ++ qstring req ++ + spc() ++ prlist_with_sep spc (pr_tacarg env) la + | (TacCall _|Tacexp _|Integer _) as a -> + str "ltac:" ++ pr_tac env (latom,E) (TacArg a) + +in (pr_tac, pr_match_rule) + +let strip_prod_binders_rawterm n (ty,_) = + let rec strip_ty acc n ty = + if n=0 then (List.rev acc, (ty,None)) else + match ty with + Rawterm.RProd(loc,na,a,b) -> + strip_ty (([dummy_loc,na],(a,None))::acc) (n-1) b + | _ -> error "Cannot translate fix tactic: not enough products" in + strip_ty [] n ty + +let strip_prod_binders_constr n ty = + let rec strip_ty acc n ty = + if n=0 then (List.rev acc, ty) else + match Term.kind_of_term ty with + Term.Prod(na,a,b) -> + strip_ty (([dummy_loc,na],a)::acc) (n-1) b + | _ -> error "Cannot translate fix tactic: not enough products" in + strip_ty [] n ty + +let drop_env f _env = f + +let rec raw_printers = + (pr_raw_tactic_level, + drop_env pr_constr, + drop_env pr_lconstr, + pr_pattern, + drop_env pr_reference, + drop_env pr_reference, + pr_reference, + pr_or_metaid pr_lident, + pr_raw_extend, + strip_prod_binders_expr) + +and pr_raw_tactic_level env n (t:raw_tactic_expr) = + fst (make_pr_tac raw_printers) env n t + +and pr_raw_match_rule env t = + snd (make_pr_tac raw_printers) env t + +let pr_and_constr_expr pr (c,_) = pr c + +let rec glob_printers = + (pr_glob_tactic_level, + (fun env -> pr_and_constr_expr (pr_rawconstr_env env)), + (fun env -> pr_and_constr_expr (pr_lrawconstr_env env)), + (fun c -> pr_constr_pattern_env (Global.env()) c), + (fun env -> pr_or_var (pr_and_short_name (pr_evaluable_reference_env env))), + (fun env -> pr_or_var (pr_inductive env)), + pr_ltac_or_var (pr_located pr_ltac_constant), + pr_lident, + pr_glob_extend, + strip_prod_binders_rawterm) + +and pr_glob_tactic_level env n (t:glob_tactic_expr) = + fst (make_pr_tac glob_printers) env n t + +and pr_glob_match_rule env t = + snd (make_pr_tac glob_printers) env t + +let ((pr_tactic_level:Environ.env -> tolerability -> Proof_type.tactic_expr -> std_ppcmds),_) = + make_pr_tac + (pr_glob_tactic_level, + pr_term_env, + pr_lterm_env, + pr_constr_pattern, + pr_evaluable_reference_env, + pr_inductive, + pr_ltac_constant, + pr_id, + pr_extend, + strip_prod_binders_constr) + +let pr_raw_tactic env = pr_raw_tactic_level env ltop +let pr_glob_tactic env = pr_glob_tactic_level env ltop +let pr_tactic env = pr_tactic_level env ltop + +let _ = Tactic_debug.set_tactic_printer + (fun x -> pr_glob_tactic (Global.env()) x) + +let _ = Tactic_debug.set_match_pattern_printer + (fun env hyp -> + pr_match_pattern + (Printer.pr_pattern_env env (Termops.names_of_rel_context env)) hyp) + +let _ = Tactic_debug.set_match_rule_printer + (fun rl -> + pr_match_rule false (pr_glob_tactic (Global.env())) Printer.pr_pattern rl) diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli index fa835fff1..60c41736b 100644 --- a/parsing/pptactic.mli +++ b/parsing/pptactic.mli @@ -16,6 +16,7 @@ open Proof_type open Topconstr open Rawterm open Ppextend +open Environ val pr_or_var : ('a -> std_ppcmds) -> 'a or_var -> std_ppcmds val pr_or_metaid : ('a -> std_ppcmds) -> 'a or_metaid -> std_ppcmds @@ -75,3 +76,15 @@ val pr_extend : (Term.constr -> std_ppcmds) -> (Term.constr -> std_ppcmds) -> (tolerability -> glob_tactic_expr -> std_ppcmds) -> int -> string -> closed_generic_argument list -> std_ppcmds + +val qsnew : string -> std_ppcmds + +val pr_intro_pattern : intro_pattern_expr -> std_ppcmds + +val pr_raw_tactic : env -> raw_tactic_expr -> std_ppcmds + +val pr_raw_tactic_level : env -> tolerability -> raw_tactic_expr -> std_ppcmds + +val pr_glob_tactic : env -> glob_tactic_expr -> std_ppcmds + +val pr_tactic : env -> Proof_type.tactic_expr -> std_ppcmds diff --git a/parsing/ppvernac.ml b/parsing/ppvernac.ml new file mode 100644 index 000000000..b8a73b37f --- /dev/null +++ b/parsing/ppvernac.ml @@ -0,0 +1,888 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(* $Id$ *) + +open Pp +open Names +open Nameops +open Nametab +open Util +open Extend +open Vernacexpr +open Ppconstr +open Pptactic +open Rawterm +open Genarg +open Pcoq +open Libnames +open Ppextend +open Topconstr +open Decl_kinds +open Tacinterp + +let pr_spc_type = pr_sep_com spc pr_type + +let pr_lident (b,_ as loc,id) = + if loc <> dummy_loc then + let (b,_) = unloc loc in + pr_located pr_id (make_loc (b,b+String.length(string_of_id id)),id) + else pr_id id + +let string_of_fqid fqid = + String.concat "." (List.map string_of_id fqid) + +let pr_fqid fqid = str (string_of_fqid fqid) + +let pr_lfqid (_,_ as loc,fqid) = + if loc <> dummy_loc then + let (b,_) = unloc loc in + pr_located pr_fqid (make_loc (b,b+String.length(string_of_fqid fqid)),fqid) + else + pr_fqid fqid + +let pr_lname = function + (loc,Name id) -> pr_lident (loc,id) + | lna -> pr_located pr_name lna + +let pr_ltac_id = Nameops.pr_id + +let pr_module = Libnames.pr_reference + +let pr_import_module = Libnames.pr_reference + +let sep_end () = str"." + +(* Warning: [pr_raw_tactic] globalises and fails if globalisation fails *) + +let pr_raw_tactic_env l env t = + pr_glob_tactic env (Tacinterp.glob_tactic_env l env t) + +let pr_gen env t = + pr_raw_generic + pr_constr + pr_lconstr + (pr_raw_tactic_level env) pr_reference t + +let pr_raw_tactic tac = + pr_glob_tactic (Global.env()) (Tacinterp.glob_tactic tac) + +let rec extract_signature = function + | [] -> [] + | Egrammar.TacNonTerm (_,(_,t),_) :: l -> t :: extract_signature l + | _::l -> extract_signature l + +let rec match_vernac_rule tys = function + [] -> raise Not_found + | pargs::rls -> + if extract_signature pargs = tys then pargs + else match_vernac_rule tys rls + +let sep = fun _ -> spc() +let sep_p = fun _ -> str"." +let sep_v = fun _ -> str"," +let sep_v2 = fun _ -> str"," ++ spc() +let sep_pp = fun _ -> str":" + +let pr_ne_sep sep pr = function + [] -> mt() + | l -> sep() ++ pr l + +let pr_entry_prec = function + | Some Gramext.LeftA -> str"LEFTA " + | Some Gramext.RightA -> str"RIGHTA " + | Some Gramext.NonA -> str"NONA " + | None -> mt() + +let pr_prec = function + | Some Gramext.LeftA -> str", left associativity" + | Some Gramext.RightA -> str", right associativity" + | Some Gramext.NonA -> str", no associativity" + | None -> mt() + +let pr_set_entry_type = function + | ETIdent -> str"ident" + | ETReference -> str"global" + | ETPattern -> str"pattern" + | ETConstr _ -> str"constr" + | ETOther (_,e) -> str e + | ETBigint -> str "bigint" + | ETConstrList _ -> failwith "Internal entry type" + +let strip_meta id = + let s = string_of_id id in + if s.[0]='$' then id_of_string (String.sub s 1 (String.length s - 1)) + else id + +let pr_production_item = function + | VNonTerm (loc,nt,Some p) -> str nt ++ str"(" ++ pr_id (strip_meta p) ++ str")" + | VNonTerm (loc,nt,None) -> str nt + | VTerm s -> qsnew s + +let pr_comment pr_c = function + | CommentConstr c -> pr_c c + | CommentString s -> qsnew s + | CommentInt n -> int n + +let pr_in_out_modules = function + | SearchInside l -> spc() ++ str"inside" ++ spc() ++ prlist_with_sep sep pr_module l + | SearchOutside [] -> mt() + | SearchOutside l -> spc() ++ str"outside" ++ spc() ++ prlist_with_sep sep pr_module l + +let pr_search_about = function + | SearchRef r -> pr_reference r + | SearchString s -> qsnew s + +let pr_search a b pr_p = match a with + | SearchHead qid -> str"Search" ++ spc() ++ pr_reference qid ++ pr_in_out_modules b + | SearchPattern c -> str"SearchPattern" ++ spc() ++ pr_p c ++ pr_in_out_modules b + | SearchRewrite c -> str"SearchRewrite" ++ spc() ++ pr_p c ++ pr_in_out_modules b + | SearchAbout sl -> str"SearchAbout" ++ spc() ++ str "[" ++ prlist_with_sep spc pr_search_about sl ++ str "]" ++ pr_in_out_modules b + +let pr_locality local = if local then str "Local " else str "" + +let pr_explanation imps = function + | ExplByPos n -> pr_id (Impargs.name_of_implicit (List.nth imps (n-1))) + | ExplByName id -> pr_id id + +let pr_class_rawexpr = function + | FunClass -> str"Funclass" + | SortClass -> str"Sortclass" + | RefClass qid -> pr_reference qid + +let pr_option_ref_value = function + | QualidRefValue id -> pr_reference id + | StringRefValue s -> qsnew s + +let pr_printoption a b = match a with + | Goptions.PrimaryTable table -> str table ++ pr_opt (prlist_with_sep sep pr_option_ref_value) b + | Goptions.SecondaryTable (table,field) -> str table ++ spc() ++ str field ++ pr_opt (prlist_with_sep sep pr_option_ref_value) b + +let pr_set_option a b = + let pr_opt_value = function + | IntValue n -> spc() ++ int n + | StringValue s -> spc() ++ str s + | BoolValue b -> mt() + in pr_printoption a None ++ pr_opt_value b + +let pr_topcmd _ = str"(* <Warning> : No printer for toplevel commands *)" + +let pr_destruct_location = function + | Tacexpr.ConclLocation () -> str"Conclusion" + | Tacexpr.HypLocation b -> if b then str"Discardable Hypothesis" else str"Hypothesis" + +let pr_opt_hintbases l = match l with + | [] -> mt() + | _ as z -> str":" ++ spc() ++ prlist_with_sep sep str z + +let pr_hints local db h pr_c pr_pat = + let opth = pr_opt_hintbases db in + let pph = + match h with + | HintsResolve l -> + str "Resolve " ++ + prlist_with_sep sep pr_c (List.map snd l) + | HintsImmediate l -> + str"Immediate" ++ spc() ++ + prlist_with_sep sep pr_c (List.map snd l) + | HintsUnfold l -> + str "Unfold " ++ + prlist_with_sep sep pr_reference (List.map snd l) + | HintsConstructors (n,c) -> + str"Constructors" ++ spc() ++ + prlist_with_sep spc pr_reference c + | HintsExtern (name,n,c,tac) -> + str "Extern" ++ spc() ++ int n ++ spc() ++ pr_pat c ++ str" =>" ++ + spc() ++ pr_raw_tactic tac + | HintsDestruct(name,i,loc,c,tac) -> + str "Destruct " ++ pr_id name ++ str" :=" ++ spc() ++ + hov 0 (int i ++ spc() ++ pr_destruct_location loc ++ spc() ++ + pr_c c ++ str " =>") ++ spc() ++ + pr_raw_tactic tac in + hov 2 (str"Hint "++pr_locality local ++ pph ++ opth) + +let pr_with_declaration pr_c = function + | CWith_Definition (id,c) -> + let p = pr_c c in + str"Definition" ++ spc() ++ pr_lfqid id ++ str" := " ++ p + | CWith_Module (id,qid) -> + str"Module" ++ spc() ++ pr_lfqid id ++ str" := " ++ + pr_located pr_qualid qid + +let rec pr_module_type pr_c = function + | CMTEident qid -> spc () ++ pr_located pr_qualid qid + | CMTEwith (mty,decl) -> + let m = pr_module_type pr_c mty in + let p = pr_with_declaration pr_c decl in + m ++ spc() ++ str"with" ++ spc() ++ p + +let pr_of_module_type prc (mty,b) = + str (if b then ":" else "<:") ++ + pr_module_type prc mty + +let pr_require_token = function + | Some true -> str "Export " + | Some false -> str "Import " + | None -> mt() + +let pr_module_vardecls pr_c (export,idl,mty) = + let m = pr_module_type pr_c mty in + (* Update the Nametab for interpreting the body of module/modtype *) + let lib_dir = Lib.library_dp() in + List.iter (fun (_,id) -> + Declaremods.process_module_bindings [id] + [make_mbid lib_dir (string_of_id id), + Modintern.interp_modtype (Global.env()) mty]) idl; + (* Builds the stream *) + spc() ++ + hov 1 (str"(" ++ pr_require_token export ++ + prlist_with_sep spc pr_lident idl ++ str":" ++ m ++ str")") + +let pr_module_binders l pr_c = + (* Effet de bord complexe pour garantir la declaration des noms des + modules parametres dans la Nametab des l'appel de pr_module_binders + malgre l'aspect paresseux des streams *) + let ml = List.map (pr_module_vardecls pr_c) l in + prlist (fun id -> id) ml + +let pr_module_binders_list l pr_c = pr_module_binders l pr_c + +let rec pr_module_expr = function + | CMEident qid -> pr_located pr_qualid qid + | CMEapply (me1,(CMEident _ as me2)) -> + pr_module_expr me1 ++ spc() ++ pr_module_expr me2 + | CMEapply (me1,me2) -> + pr_module_expr me1 ++ spc() ++ + hov 1 (str"(" ++ pr_module_expr me2 ++ str")") + +let pr_type_option pr_c = function + | CHole loc -> mt() + | _ as c -> brk(0,2) ++ str":" ++ pr_c c + +let pr_decl_notation prc = + pr_opt (fun (ntn,c,scopt) -> fnl () ++ + str "where " ++ qsnew ntn ++ str " := " ++ prc c ++ + pr_opt (fun sc -> str ": " ++ str sc) scopt) + +let pr_vbinders l = + hv 0 (pr_binders l) + +let pr_binders_arg = + pr_ne_sep spc pr_binders + +let pr_and_type_binders_arg bl = + pr_binders_arg bl + +let pr_onescheme (id,dep,ind,s) = + hov 0 (pr_lident id ++ str" :=") ++ spc() ++ + hov 0 ((if dep then str"Induction for" else str"Minimality for") + ++ spc() ++ pr_reference ind) ++ spc() ++ + hov 0 (str"Sort" ++ spc() ++ pr_sort s) + +let begin_of_inductive = function + [] -> 0 + | (_,((loc,_),_))::_ -> fst (unloc loc) + +let pr_class_rawexpr = function + | FunClass -> str"Funclass" + | SortClass -> str"Sortclass" + | RefClass qid -> pr_reference qid + +let pr_assumption_token many = function + | (Local,Logical) -> + str (if many then "Hypotheses" else "Hypothesis") + | (Local,Definitional) -> + str (if many then "Variables" else "Variable") + | (Global,Logical) -> + str (if many then "Axioms" else "Axiom") + | (Global,Definitional) -> + str (if many then "Parameters" else "Parameter") + | (Global,Conjectural) -> str"Conjecture" + | (Local,Conjectural) -> + anomaly "Don't know how to translate a local conjecture" + +let pr_params pr_c (xl,(c,t)) = + hov 2 (prlist_with_sep sep pr_lident xl ++ spc() ++ + (if c then str":>" else str":" ++ + spc() ++ pr_c t)) + +let rec factorize = function + | [] -> [] + | (c,(idl,t))::l -> + match factorize l with + | (xl,t')::l' when t' = (c,t) -> (idl@xl,t')::l' + | l' -> (idl,(c,t))::l' + +let pr_ne_params_list pr_c l = + match factorize l with + | [p] -> pr_params pr_c p + | l -> + prlist_with_sep spc + (fun p -> hov 1 (str "(" ++ pr_params pr_c p ++ str ")")) l +(* + prlist_with_sep pr_semicolon (pr_params pr_c) +*) + +let pr_thm_token = function + | Theorem -> str"Theorem" + | Lemma -> str"Lemma" + | Fact -> str"Fact" + | Remark -> str"Remark" + +let pr_syntax_modifier = function + | SetItemLevel (l,NextLevel) -> + prlist_with_sep sep_v2 str l ++ + spc() ++ str"at next level" + | SetItemLevel (l,NumLevel n) -> + prlist_with_sep sep_v2 str l ++ + spc() ++ str"at level" ++ spc() ++ int n + | SetLevel n -> str"at level" ++ spc() ++ int n + | SetAssoc Gramext.LeftA -> str"left associativity" + | SetAssoc Gramext.RightA -> str"right associativity" + | SetAssoc Gramext.NonA -> str"no associativity" + | SetEntryType (x,typ) -> str x ++ spc() ++ pr_set_entry_type typ + | SetOnlyParsing -> str"only parsing" + | SetFormat s -> str"format " ++ pr_located qsnew s + +let pr_syntax_modifiers = function + | [] -> mt() + | l -> spc() ++ + hov 1 (str"(" ++ prlist_with_sep sep_v2 pr_syntax_modifier l ++ str")") + +let print_level n = + if n <> 0 then str " (at level " ++ int n ++ str ")" else mt () + +let pr_grammar_tactic_rule n (_,pil,t) = + hov 2 (str "Tactic Notation" ++ print_level n ++ spc() ++ + hov 0 (prlist_with_sep sep pr_production_item pil ++ + spc() ++ str":=" ++ spc() ++ pr_raw_tactic t)) + +let pr_box b = let pr_boxkind = function + | PpHB n -> str"h" ++ spc() ++ int n + | PpVB n -> str"v" ++ spc() ++ int n + | PpHVB n -> str"hv" ++ spc() ++ int n + | PpHOVB n -> str"hov" ++ spc() ++ int n + | PpTB -> str"t" +in str"<" ++ pr_boxkind b ++ str">" + +let pr_paren_reln_or_extern = function + | None,L -> str"L" + | None,E -> str"E" + | Some pprim,Any -> qsnew pprim + | Some pprim,Prec p -> qsnew pprim ++ spc() ++ str":" ++ spc() ++ int p + | _ -> mt() +(* +let rec pr_next_hunks = function + | UNP_FNL -> str"FNL" + | UNP_TAB -> str"TAB" + | RO c -> qsnew c + | UNP_BOX (b,ll) -> str"[" ++ pr_box b ++ prlist_with_sep sep pr_next_hunks ll ++ str"]" + | UNP_BRK (n,m) -> str"[" ++ int n ++ spc() ++ int m ++ str"]" + | UNP_TBRK (n,m) -> str"[ TBRK" ++ int n ++ spc() ++ int m ++ str"]" + | PH (e,None,_) -> print_ast e + | PH (e,Some ext,pr) -> print_ast e ++ spc() ++ str":" ++ spc() ++ pr_paren_reln_or_extern (Some ext,pr) + | UNP_SYMBOLIC _ -> mt() + +let pr_unparsing u = + str "[ " ++ prlist_with_sep sep pr_next_hunks u ++ str " ]" + +let pr_astpat a = str"<<" ++ print_ast a ++ str">>" + +let pr_syntax_rule (nm,s,u) = str nm ++ spc() ++ str"[" ++ pr_astpat s ++ str"]" ++ spc() ++ str"->" ++ spc() ++ pr_unparsing u + +let pr_syntax_entry (p,rl) = + str"level" ++ spc() ++ int p ++ str" :" ++ fnl() ++ + prlist_with_sep (fun _ -> fnl() ++ str"| ") pr_syntax_rule rl +*) +let pr_vernac_solve (i,env,tac,deftac) = + (if i = 1 then mt() else int i ++ str ": ") ++ + pr_glob_tactic env tac + ++ (try if deftac & Pfedit.get_end_tac() <> None then str ".." else mt () + with UserError _|Stdpp.Exc_located _ -> mt()) + +(**************************************) +(* Pretty printer for vernac commands *) +(**************************************) +let make_pr_vernac pr_constr pr_lconstr = + +let pr_constrarg c = spc () ++ pr_constr c in +let pr_lconstrarg c = spc () ++ pr_lconstr c in +let pr_intarg n = spc () ++ int n in + +let rec pr_vernac = function + + (* Proof management *) + | VernacAbortAll -> str "Abort All" + | VernacRestart -> str"Restart" + | VernacSuspend -> str"Suspend" + | VernacUnfocus -> str"Unfocus" + | VernacGoal c -> str"Goal" ++ pr_lconstrarg c + | VernacAbort id -> str"Abort" ++ pr_opt pr_lident id + | VernacResume id -> str"Resume" ++ pr_opt pr_lident id + | VernacUndo i -> if i=1 then str"Undo" else str"Undo" ++ pr_intarg i + | VernacBacktrack (i,j,k) -> + str "Backtrack" ++ spc() ++ prlist_with_sep sep int [i;j;k] + | VernacFocus i -> str"Focus" ++ pr_opt int i + | VernacGo g -> + let pr_goable = function + | GoTo i -> int i + | GoTop -> str"top" + | GoNext -> str"next" + | GoPrev -> str"prev" + in str"Go" ++ spc() ++ pr_goable g + | VernacShow s -> + let pr_showable = function + | ShowGoal n -> str"Show" ++ pr_opt int n + | ShowGoalImplicitly n -> str"Show Implicit Arguments" ++ pr_opt int n + | ShowProof -> str"Show Proof" + | ShowNode -> str"Show Node" + | ShowScript -> str"Show Script" + | ShowExistentials -> str"Show Existentials" + | ShowTree -> str"Show Tree" + | ShowProofNames -> str"Show Conjectures" + | ShowIntros b -> str"Show " ++ (if b then str"Intros" else str"Intro") + | ShowMatch id -> str"Show Match " ++ pr_lident id + | ExplainProof l -> str"Explain Proof" ++ spc() ++ prlist_with_sep sep int l + | ExplainTree l -> str"Explain Proof Tree" ++ spc() ++ prlist_with_sep sep int l + in pr_showable s + | VernacCheckGuard -> str"Guarded" + | VernacDebug b -> pr_topcmd b + + (* Resetting *) + | VernacResetName id -> str"Reset" ++ spc() ++ pr_lident id + | VernacResetInitial -> str"Reset Initial" + | VernacBack i -> if i=1 then str"Back" else str"Back" ++ pr_intarg i + | VernacBackTo i -> str"BackTo" ++ pr_intarg i + + (* State management *) + | VernacWriteState s -> str"Write State" ++ spc () ++ qsnew s + | VernacRestoreState s -> str"Restore State" ++ spc() ++ qsnew s + + (* Control *) + | VernacList l -> + hov 2 (str"[" ++ spc() ++ + prlist (fun v -> pr_located pr_vernac v ++ sep_end () ++ fnl()) l + ++ spc() ++ str"]") + | VernacLoad (f,s) -> str"Load" ++ if f then (spc() ++ str"Verbose" + ++ spc()) else spc() ++ qsnew s + | VernacTime v -> str"Time" ++ spc() ++ pr_vernac v + | VernacVar id -> pr_lident id + + (* Syntax *) + | VernacTacticNotation (n,r,e) -> pr_grammar_tactic_rule n ("",r,e) + | VernacOpenCloseScope (local,opening,sc) -> + str (if opening then "Open " else "Close ") ++ pr_locality local ++ + str "Scope" ++ spc() ++ str sc + | VernacDelimiters (sc,key) -> + str"Delimit Scope" ++ spc () ++ str sc ++ + spc() ++ str "with " ++ str key + | VernacBindScope (sc,cll) -> + str"Bind Scope" ++ spc () ++ str sc ++ + spc() ++ str "with " ++ prlist_with_sep spc pr_class_rawexpr cll + | VernacArgumentsScope (q,scl) -> let pr_opt_scope = function + | None -> str"_" + | Some sc -> str sc in + str"Arguments Scope" ++ spc() ++ pr_reference q ++ spc() ++ str"[" ++ prlist_with_sep sep pr_opt_scope scl ++ str"]" + | VernacInfix (local,(s,mv),q,sn) -> (* A Verifier *) + hov 0 (hov 0 (str"Infix " ++ pr_locality local + ++ qsnew s ++ str " :=" ++ spc() ++ pr_reference q) ++ + pr_syntax_modifiers mv ++ + (match sn with + | None -> mt() + | Some sc -> spc() ++ str":" ++ spc() ++ str sc)) + | VernacNotation (local,c,(s,l),opt) -> + let ps = + let n = String.length s in + if n > 2 & s.[0] = '\'' & s.[n-1] = '\'' + then + let s' = String.sub s 1 (n-2) in + if String.contains s' '\'' then qsnew s else str s' + else qsnew s in + hov 2( str"Notation" ++ spc() ++ pr_locality local ++ ps ++ + str " :=" ++ pr_constrarg c ++ pr_syntax_modifiers l ++ + (match opt with + | None -> mt() + | Some sc -> str" :" ++ spc() ++ str sc)) + | VernacSyntaxExtension (local,(s,l)) -> + str"Reserved Notation" ++ spc() ++ pr_locality local ++ qsnew s ++ + pr_syntax_modifiers l + + (* Gallina *) + | VernacDefinition (d,id,b,f) -> (* A verifier... *) + let pr_def_token = function + | Local, Coercion -> str"Coercion Local" + | Global, Coercion -> str"Coercion" + | Local, Definition _ -> str"Let" + | Global, Definition b -> + if b then str"Boxed Definition" + else str"Definition" + | Local, SubClass -> str"Local SubClass" + | Global, SubClass -> str"SubClass" + | Global, CanonicalStructure -> str"Canonical Structure" + | Local, CanonicalStructure -> + anomaly "Don't know how to translate a local canonical structure" in + let pr_reduce = function + | None -> mt() + | Some r -> + str"Eval" ++ spc() ++ + pr_red_expr (pr_constr, pr_lconstr, pr_reference) r ++ + str" in" ++ spc() in + let pr_def_body = function + | DefineBody (bl,red,body,d) -> + let ty = match d with + | None -> mt() + | Some ty -> spc() ++ str":" ++ pr_sep_com spc pr_lconstr ty + in + (pr_binders_arg bl,ty,Some (pr_reduce red ++ pr_lconstr body)) + | ProveBody (bl,t) -> + (pr_binders_arg bl, str" :" ++ pr_spc_type t, None) in + let (binds,typ,c) = pr_def_body b in + hov 2 (pr_def_token d ++ spc() ++ pr_lident id ++ binds ++ typ ++ + (match c with + | None -> mt() + | Some cc -> str" :=" ++ spc() ++ cc)) + + | VernacStartTheoremProof (ki,id,(bl,c),b,d) -> + hov 1 (pr_thm_token ki ++ spc() ++ pr_lident id ++ spc() ++ + (match bl with + | [] -> mt() + | _ -> pr_binders bl ++ spc()) + ++ str":" ++ pr_spc_type c) + | VernacEndProof Admitted -> str"Admitted" + | VernacEndProof (Proved (opac,o)) -> (match o with + | None -> if opac then str"Qed" else str"Defined" + | Some (id,th) -> (match th with + | None -> (if opac then str"Save" else str"Defined") ++ spc() ++ pr_lident id + | Some tok -> str"Save" ++ spc() ++ pr_thm_token tok ++ spc() ++ pr_lident id)) + | VernacExactProof c -> + hov 2 (str"Proof" ++ pr_lconstrarg c) + | VernacAssumption (stre,l) -> + let n = List.length (List.flatten (List.map fst (List.map snd l))) in + hov 2 + (pr_assumption_token (n > 1) stre ++ spc() ++ + pr_ne_params_list pr_type l) + | VernacInductive (f,l) -> + + let pr_constructor (coe,(id,c)) = + hov 2 (pr_lident id ++ str" " ++ + (if coe then str":>" else str":") ++ + pr_sep_com spc pr_type c) in + let pr_constructor_list l = match l with + | [] -> mt() + | _ -> + pr_com_at (begin_of_inductive l) ++ + fnl() ++ + str (if List.length l = 1 then " " else " | ") ++ + prlist_with_sep (fun _ -> fnl() ++ str" | ") pr_constructor l in + let pr_oneind key (id,ntn,indpar,s,lc) = + hov 0 ( + str key ++ spc() ++ + pr_lident id ++ pr_and_type_binders_arg indpar ++ spc() ++ str":" ++ + spc() ++ pr_type s ++ + str" :=") ++ pr_constructor_list lc ++ + pr_decl_notation pr_constr ntn in + + hov 1 (pr_oneind (if f then "Inductive" else "CoInductive") (List.hd l)) + ++ + (prlist (fun ind -> fnl() ++ hov 1 (pr_oneind "with" ind)) (List.tl l)) + + + | VernacFixpoint (recs,b) -> + let name_of_binder = function + | LocalRawAssum (nal,_) -> nal + | LocalRawDef (_,_) -> [] in + let pr_onerec = function + | (id,n,bl,type_,def),ntn -> + let (bl',def,type_) = + if Options.do_translate() then extract_def_binders def type_ + else ([],def,type_) in + let bl = bl @ bl' in + let ids = List.flatten (List.map name_of_binder bl) in + let name = + try snd (List.nth ids n) + with Failure _ -> + warn (str "non-printable fixpoint \""++pr_id id++str"\""); + Anonymous in + let annot = + if List.length ids > 1 then + spc() ++ str "{struct " ++ pr_name name ++ str"}" + else mt() in + pr_id id ++ pr_binders_arg bl ++ annot ++ spc() + ++ pr_type_option (fun c -> spc() ++ pr_type c) type_ + ++ str" :=" ++ brk(1,1) ++ pr_lconstr def ++ + pr_decl_notation pr_constr ntn + in + let start = if b then "Boxed Fixpoint" else "Fixpoint" in + hov 1 (str start ++ spc() ++ + prlist_with_sep (fun _ -> fnl() ++ fnl() ++ str"with ") pr_onerec recs) + + | VernacCoFixpoint (corecs,b) -> + let pr_onecorec (id,bl,c,def) = + let (bl',def,c) = + if Options.do_translate() then extract_def_binders def c + else ([],def,c) in + let bl = bl @ bl' in + pr_id id ++ spc() ++ pr_binders bl ++ spc() ++ str":" ++ + spc() ++ pr_type c ++ + str" :=" ++ brk(1,1) ++ pr_lconstr def in + let start = if b then "Boxed CoFixpoint" else "CoFixpoint" in + hov 1 (str start ++ spc() ++ + prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onecorec corecs) + | VernacScheme l -> + hov 2 (str"Scheme" ++ spc() ++ + prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onescheme l) + + (* Gallina extensions *) + | VernacRecord (b,(oc,name),ps,s,c,fs) -> + let pr_record_field = function + | (oc,AssumExpr (id,t)) -> + hov 1 (pr_lname id ++ + (if oc then str" :>" else str" :") ++ spc() ++ + pr_type t) + | (oc,DefExpr(id,b,opt)) -> (match opt with + | Some t -> + hov 1 (pr_lname id ++ + (if oc then str" :>" else str" :") ++ spc() ++ + pr_type t ++ str" :=" ++ pr_lconstr b) + | None -> + hov 1 (pr_lname id ++ str" :=" ++ spc() ++ + pr_lconstr b)) in + hov 2 + (str (if b then "Record" else "Structure") ++ + (if oc then str" > " else str" ") ++ pr_lident name ++ + pr_and_type_binders_arg ps ++ str" :" ++ spc() ++ pr_type s ++ + str" := " ++ + (match c with + | None -> mt() + | Some sc -> pr_lident sc) ++ + spc() ++ str"{" ++ + hv 0 (prlist_with_sep pr_semicolon pr_record_field fs ++ str"}")) + | VernacBeginSection id -> hov 2 (str"Section" ++ spc () ++ pr_lident id) + | VernacEndSegment id -> hov 2 (str"End" ++ spc() ++ pr_lident id) + | VernacRequire (exp,spe,l) -> hov 2 + (str "Require" ++ spc() ++ pr_require_token exp ++ + (match spe with + | None -> mt() + | Some flag -> + (if flag then str"Specification" else str"Implementation") ++ + spc ()) ++ + prlist_with_sep sep pr_module l) + | VernacImport (f,l) -> + (if f then str"Export" else str"Import") ++ spc() ++ + prlist_with_sep sep pr_import_module l + | VernacCanonical q -> str"Canonical Structure" ++ spc() ++ pr_reference q + | VernacCoercion (s,id,c1,c2) -> + hov 1 ( + str"Coercion" ++ (match s with | Local -> spc() ++ + str"Local" ++ spc() | Global -> spc()) ++ + pr_reference id ++ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ + spc() ++ str">->" ++ spc() ++ pr_class_rawexpr c2) + | VernacIdentityCoercion (s,id,c1,c2) -> + hov 1 ( + str"Identity Coercion" ++ (match s with | Local -> spc() ++ + str"Local" ++ spc() | Global -> spc()) ++ pr_lident id ++ + spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ spc() ++ str">->" ++ + spc() ++ pr_class_rawexpr c2) + + (* Modules and Module Types *) + | VernacDefineModule (export,m,bl,ty,bd) -> + let b = pr_module_binders_list bl pr_lconstr in + hov 2 (str"Module" ++ spc() ++ pr_require_token export ++ + pr_lident m ++ b ++ + pr_opt (pr_of_module_type pr_lconstr) ty ++ + pr_opt (fun me -> str ":= " ++ pr_module_expr me) bd) + | VernacDeclareModule (export,id,bl,m1) -> + let b = pr_module_binders_list bl pr_lconstr in + hov 2 (str"Declare Module" ++ spc() ++ pr_require_token export ++ + pr_lident id ++ b ++ + pr_of_module_type pr_lconstr m1) + | VernacDeclareModuleType (id,bl,m) -> + let b = pr_module_binders_list bl pr_lconstr in + hov 2 (str"Module Type " ++ pr_lident id ++ b ++ + pr_opt (fun mt -> str ":= " ++ pr_module_type pr_lconstr mt) m) + + (* Solving *) + | VernacSolve (i,tac,deftac) -> + (* Normally shunted by vernac.ml *) + let env = + try snd (Pfedit.get_goal_context i) + with UserError _ -> Global.env() in + let tac = + Options.with_option Options.translate_syntax + (Constrintern.for_grammar (Tacinterp.glob_tactic_env [] env)) tac in + pr_vernac_solve (i,env,tac,deftac) + + | VernacSolveExistential (i,c) -> + str"Existential " ++ int i ++ pr_lconstrarg c + + (* Auxiliary file and library management *) + | VernacRequireFrom (exp,spe,f) -> hov 2 + (str"Require" ++ spc() ++ pr_require_token exp ++ + (match spe with + | None -> mt() + | Some false -> str"Implementation" ++ spc() + | Some true -> str"Specification" ++ spc ()) ++ + qsnew f) + | VernacAddLoadPath (fl,s,d) -> hov 2 + (str"Add" ++ + (if fl then str" Rec " else spc()) ++ + str"LoadPath" ++ spc() ++ qsnew s ++ + (match d with + | None -> mt() + | Some dir -> spc() ++ str"as" ++ spc() ++ pr_dirpath dir)) + | VernacRemoveLoadPath s -> str"Remove LoadPath" ++ qsnew s + | VernacAddMLPath (fl,s) -> + str"Add" ++ (if fl then str" Rec " else spc()) ++ str"ML Path" ++ qsnew s + | VernacDeclareMLModule l -> + hov 2 (str"Declare ML Module" ++ spc() ++ prlist_with_sep sep qsnew l) + | VernacChdir s -> str"Cd" ++ pr_opt qsnew s + + (* Commands *) + | VernacDeclareTacticDefinition (rc,l) -> + let pr_tac_body (id, body) = + let idl, body = + match body with + | Tacexpr.TacFun (idl,b) -> idl,b + | _ -> [], body in + pr_located pr_ltac_id id ++ + prlist (function None -> str " _" + | Some id -> spc () ++ pr_id id) idl + ++ str" :=" ++ brk(1,1) ++ + let idl = List.map out_some (List.filter (fun x -> not (x=None)) idl)in + pr_raw_tactic_env + (idl @ List.map snd (List.map fst l)) + (Global.env()) + body in + hov 1 + (((*if !Options.p1 then + (if rc then str "Recursive " else mt()) ++ + str "Tactic Definition " else*) + (* Rec by default *) str "Ltac ") ++ + prlist_with_sep (fun () -> fnl() ++ str"with ") pr_tac_body l) + | VernacHints (local,dbnames,h) -> + pr_hints local dbnames h pr_constr pr_pattern + | VernacSyntacticDefinition (id,c,local,onlyparsing) -> + hov 2 + (str"Notation " ++ pr_locality local ++ pr_id id ++ str" :=" ++ + pr_constrarg c ++ + pr_syntax_modifiers (if onlyparsing then [SetOnlyParsing] else [])) + | VernacDeclareImplicits (q,None) -> + hov 2 (str"Implicit Arguments" ++ spc() ++ pr_reference q) + | VernacDeclareImplicits (q,Some l) -> + let r = Nametab.global q in + Impargs.declare_manual_implicits r l; + let imps = Impargs.implicits_of_global r in + hov 1 (str"Implicit Arguments" ++ spc() ++ pr_reference q ++ spc() ++ + str"[" ++ prlist_with_sep sep (pr_explanation imps) l ++ str"]") + | VernacReserve (idl,c) -> + hov 1 (str"Implicit Type" ++ + str (if List.length idl > 1 then "s " else " ") ++ + prlist_with_sep spc pr_lident idl ++ str " :" ++ spc () ++ pr_type c) + | VernacSetOpacity (fl,l) -> + hov 1 ((if fl then str"Opaque" else str"Transparent") ++ + spc() ++ prlist_with_sep sep pr_reference l) + | VernacUnsetOption na -> + hov 1 (str"Unset" ++ spc() ++ pr_printoption na None) + | VernacSetOption (na,v) -> hov 2 (str"Set" ++ spc() ++ pr_set_option na v) + | VernacAddOption (na,l) -> hov 2 (str"Add" ++ spc() ++ pr_printoption na (Some l)) + | VernacRemoveOption (na,l) -> hov 2 (str"Remove" ++ spc() ++ pr_printoption na (Some l)) + | VernacMemOption (na,l) -> hov 2 (str"Test" ++ spc() ++ pr_printoption na (Some l)) + | VernacPrintOption na -> hov 2 (str"Test" ++ spc() ++ pr_printoption na None) + | VernacCheckMayEval (r,io,c) -> + let pr_mayeval r c = match r with + | Some r0 -> + hov 2 (str"Eval" ++ spc() ++ + pr_red_expr (pr_constr,pr_lconstr,pr_reference) r0 ++ + spc() ++ str"in" ++ spc () ++ pr_constr c) + | None -> hov 2 (str"Check" ++ spc() ++ pr_constr c) + in + (if io = None then mt() else int (out_some io) ++ str ": ") ++ + pr_mayeval r c + | VernacGlobalCheck c -> hov 2 (str"Type" ++ pr_constrarg c) + | VernacPrint p -> + let pr_printable = function + | PrintFullContext -> str"Print All" + | PrintSectionContext s -> + str"Print Section" ++ spc() ++ Libnames.pr_reference s + | PrintGrammar (uni,ent) -> + msgerrnl (str "warning: no direct translation of Print Grammar entry"); + str"Print Grammar" ++ spc() ++ str ent + | PrintLoadPath -> str"Print LoadPath" + | PrintModules -> str"Print Modules" + | PrintMLLoadPath -> str"Print ML Path" + | PrintMLModules -> str"Print ML Modules" + | PrintGraph -> str"Print Graph" + | PrintClasses -> str"Print Classes" + | PrintLtac qid -> str"Print Ltac" ++ spc() ++ pr_reference qid + | PrintCoercions -> str"Print Coercions" + | PrintCoercionPaths (s,t) -> str"Print Coercion Paths" ++ spc() ++ pr_class_rawexpr s ++ spc() ++ pr_class_rawexpr t + | PrintCanonicalConversions -> str"Print Canonical Structures" + | PrintTables -> str"Print Tables" + | PrintOpaqueName qid -> str"Print Term" ++ spc() ++ pr_reference qid + | PrintHintGoal -> str"Print Hint" + | PrintHint qid -> str"Print Hint" ++ spc() ++ pr_reference qid + | PrintHintDb -> str"Print Hint *" + | PrintHintDbName s -> str"Print HintDb" ++ spc() ++ str s + | PrintRewriteHintDbName s -> str"Print Rewrite HintDb" ++ spc() ++ str s + | PrintUniverses fopt -> str"Dump Universes" ++ pr_opt str fopt + | PrintName qid -> str"Print" ++ spc() ++ pr_reference qid + | PrintModuleType qid -> str"Print Module Type" ++ spc() ++ pr_reference qid + | PrintModule qid -> str"Print Module" ++ spc() ++ pr_reference qid + | PrintInspect n -> str"Inspect" ++ spc() ++ int n + | PrintSetoids -> str"Print Setoids" + | PrintScopes -> str"Print Scopes" + | PrintScope s -> str"Print Scope" ++ spc() ++ str s + | PrintVisibility s -> str"Print Visibility" ++ pr_opt str s + | PrintAbout qid -> str"About" ++ spc() ++ pr_reference qid + | PrintImplicit qid -> str"Print Implicit" ++ spc() ++ pr_reference qid + in pr_printable p + | VernacSearch (sea,sea_r) -> pr_search sea sea_r pr_pattern + | VernacLocate loc -> + let pr_locate =function + | LocateTerm qid -> pr_reference qid + | LocateFile f -> str"File" ++ spc() ++ qsnew f + | LocateLibrary qid -> str"Library" ++ spc () ++ pr_module qid + | LocateModule qid -> str"Module" ++ spc () ++ pr_module qid + | LocateNotation s -> qsnew s + in str"Locate" ++ spc() ++ pr_locate loc + | VernacComments l -> + hov 2 + (str"Comments" ++ spc() ++ prlist_with_sep sep (pr_comment pr_constr) l) + | VernacNop -> mt() + + (* Toplevel control *) + | VernacToplevelControl exn -> pr_topcmd exn + + (* For extension *) + | VernacExtend (s,c) -> pr_extend s c + | VernacProof Tacexpr.TacId _ -> str "Proof" + | VernacProof te -> str "Proof with" ++ spc() ++ pr_raw_tactic te + +and pr_extend s cl = + let pr_arg a = + try pr_gen (Global.env()) a + with Failure _ -> str ("<error in "^s^">") in + try + let rls = List.assoc s (Egrammar.get_extend_vernac_grammars()) in + let rl = match_vernac_rule (List.map Genarg.genarg_tag cl) rls in + let (pp,_) = + List.fold_left + (fun (strm,args) pi -> + match pi with + Egrammar.TacNonTerm _ -> + (strm ++ pr_gen (Global.env()) (List.hd args), + List.tl args) + | Egrammar.TacTerm s -> (strm ++ spc() ++ str s, args)) + (mt(),cl) rl in + hov 1 pp + with Not_found -> + hov 1 (str ("TODO("^s) ++ prlist_with_sep sep pr_arg cl ++ str ")") + +in pr_vernac + +let pr_vernac v = make_pr_vernac pr_constr pr_lconstr v ++ sep_end () diff --git a/parsing/ppvernac.mli b/parsing/ppvernac.mli new file mode 100644 index 000000000..e76f1b579 --- /dev/null +++ b/parsing/ppvernac.mli @@ -0,0 +1,31 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(*i $Id$ i*) + +open Pp +open Genarg +open Vernacexpr +open Names +open Nameops +open Nametab +open Util +open Ppconstr +open Pptactic +open Rawterm +open Pcoq +open Libnames +open Ppextend +open Topconstr + +val sep_end : unit -> std_ppcmds + +val pr_vernac : vernac_expr -> std_ppcmds + +val pr_vernac_solve : + int * Environ.env * Tacexpr.glob_tactic_expr * bool -> std_ppcmds diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml index 76c87f2c8..821679b4f 100644 --- a/parsing/prettyp.ml +++ b/parsing/prettyp.ml @@ -526,38 +526,6 @@ let print_impargs ref = (if has_impl then print_impl_args impl else (str "No implicit arguments" ++ fnl ())) -let print_local_context () = - let env = Lib.contents_after None in - let rec print_var_rec = function - | [] -> (mt ()) - | (oname,Lib.Leaf lobj)::rest -> - if "VARIABLE" = object_tag lobj then - let d = get_variable (basename (fst oname)) in - (print_var_rec rest ++ - print_named_decl d) - else - print_var_rec rest - | _::rest -> print_var_rec rest - - and print_last_const = function - | (oname,Lib.Leaf lobj)::rest -> - (match object_tag lobj with - | "CONSTANT" -> - let kn = constant_of_kn (snd oname) in - let {const_body=val_0;const_type=typ} = - Global.lookup_constant kn in - (print_last_const rest ++ - print_basename kn ++str" = " ++ - print_typed_body (val_0,typ)) - | "INDUCTIVE" -> - let kn = snd oname in - (print_last_const rest ++print_mutual kn ++ fnl ()) - | "VARIABLE" -> (mt ()) - | _ -> print_last_const rest) - | _ -> (mt ()) - in - (print_var_rec env ++ print_last_const env) - let unfold_head_fconst = let rec unfrec k = match kind_of_term k with | Const cst -> constant_value (Global.env ()) cst diff --git a/parsing/prettyp.mli b/parsing/prettyp.mli index 4c52420d0..f23cc38ac 100644 --- a/parsing/prettyp.mli +++ b/parsing/prettyp.mli @@ -42,7 +42,6 @@ val build_inductive : mutual_inductive -> int -> val print_mutual : mutual_inductive -> std_ppcmds val print_name : reference -> std_ppcmds val print_opaque_name : reference -> std_ppcmds -val print_local_context : unit -> std_ppcmds val print_about : reference -> std_ppcmds val print_impargs : reference -> std_ppcmds diff --git a/parsing/printer.ml b/parsing/printer.ml index a87415d95..7efea20a8 100644 --- a/parsing/printer.ml +++ b/parsing/printer.ml @@ -25,6 +25,7 @@ open Evd open Proof_type open Refiner open Pfedit +open Ppconstr let emacs_str s = if !Options.print_emacs then s else "" @@ -33,13 +34,13 @@ let emacs_str s = if !Options.print_emacs then s else "" (* [at_top] means ids of env must be avoided in bound variables *) let prterm_core at_top env t = - Ppconstrnew.pr_lconstr (Constrextern.extern_constr at_top env t) + pr_lconstr (Constrextern.extern_constr at_top env t) let prtype_core at_top env t = - Ppconstrnew.pr_lconstr (Constrextern.extern_type at_top env t) + pr_lconstr (Constrextern.extern_type at_top env t) let pr_cases_pattern t = - Ppconstrnew.pr_cases_pattern (Constrextern.extern_cases_pattern Idset.empty t) + pr_cases_pattern (Constrextern.extern_cases_pattern Idset.empty t) let pr_pattern_env tenv env t = - Ppconstrnew.pr_constr (Constrextern.extern_pattern tenv env t) + pr_constr (Constrextern.extern_pattern tenv env t) (**********************************************************************) (* Derived printers *) @@ -70,7 +71,7 @@ let pr_evaluable_reference ref = pr_global ref' let pr_rawterm t = - Ppconstrnew.pr_lconstr (Constrextern.extern_rawconstr Idset.empty t) + pr_lconstr (Constrextern.extern_rawconstr Idset.empty t) open Pattern @@ -284,10 +285,10 @@ let pr_nth_open_subgoal n = (* Elementary tactics *) let print_constr8 t = - Ppconstrnew.pr_constr (Constrextern.extern_constr false (Global.env()) t) + pr_constr (Constrextern.extern_constr false (Global.env()) t) let print_lconstr8 t = - Ppconstrnew.pr_lconstr (Constrextern.extern_constr false (Global.env()) t) + pr_lconstr (Constrextern.extern_constr false (Global.env()) t) let pr_prim_rule = function | Intro id -> diff --git a/parsing/tactic_printer.ml b/parsing/tactic_printer.ml index 32a673cdc..827cfcd0e 100644 --- a/parsing/tactic_printer.ml +++ b/parsing/tactic_printer.ml @@ -21,9 +21,9 @@ open Printer let pr_tactic = function | TacArg (Tacexp t) -> (*top tactic from tacinterp*) - Pptacticnew.pr_glob_tactic (Global.env()) t + Pptactic.pr_glob_tactic (Global.env()) t | t -> - Pptacticnew.pr_tactic (Global.env()) t + Pptactic.pr_tactic (Global.env()) t let pr_rule = function | Prim r -> hov 0 (pr_prim_rule r) |