From cb1ae314411d78952062e5092804b85d981ad6e1 Mon Sep 17 00:00:00 2001 From: barras Date: Wed, 12 Mar 2003 17:49:21 +0000 Subject: *** empty log message *** git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@3761 85f007b7-540e-0410-9357-904b9bb8a0f7 --- contrib/field/field.ml4 | 13 + contrib/interface/centaur.ml4 | 6 +- contrib/interface/name_to_ast.ml | 2 +- contrib/interface/pbp.ml | 4 +- contrib/interface/showproof.ml | 4 +- contrib/interface/xlate.ml | 8 +- contrib/omega/coq_omega.ml | 8 +- contrib/ring/Ring_theory.v | 15 +- contrib/ring/Setoid_ring_theory.v | 14 +- contrib/romega/ReflOmegaCore.v | 35 +- contrib/romega/const_omega.ml | 8 +- contrib/xml/cic2acic.ml | 4 +- contrib/xml/xmlcommand.ml4 | 12 +- contrib/xml/xmlentries.ml4 | 4 +- interp/constrextern.ml | 172 +++-- interp/constrextern.mli | 6 +- interp/constrintern.ml | 32 +- interp/topconstr.ml | 4 +- kernel/names.ml | 9 +- lib/options.ml | 3 +- lib/options.mli | 1 + lib/pp_control.ml | 4 +- library/declare.ml | 2 +- library/goptions.ml | 2 +- library/library.mli | 2 +- library/nametab.ml | 37 +- library/nametab.mli | 13 +- parsing/argextend.ml4 | 27 +- parsing/ast.ml | 20 +- parsing/egrammar.ml | 52 +- parsing/egrammar.mli | 6 + parsing/extend.ml | 3 +- parsing/g_basevernac.ml4 | 63 +- parsing/g_cases.ml4 | 6 +- parsing/g_constr.ml4 | 68 +- parsing/g_constrnew.ml4 | 256 +++++++ parsing/g_ltac.ml4 | 8 +- parsing/g_ltacnew.ml4 | 209 ++++++ parsing/g_natsyntaxnew.ml | 195 +++++ parsing/g_natsyntaxnew.mli | 11 + parsing/g_prim.ml4 | 2 + parsing/g_primnew.ml4 | 143 ++++ parsing/g_proofs.ml4 | 2 + parsing/g_proofsnew.ml4 | 126 ++++ parsing/g_rsyntax.ml | 8 +- parsing/g_rsyntaxnew.ml | 113 +++ parsing/g_tactic.ml4 | 10 +- parsing/g_tacticnew.ml4 | 355 +++++++++ parsing/g_vernac.ml4 | 17 +- parsing/g_vernacnew.ml4 | 817 +++++++++++++++++++++ parsing/g_zsyntax.ml | 10 +- parsing/g_zsyntaxnew.ml | 168 +++++ parsing/g_zsyntaxnew.mli | 11 + parsing/lexer.ml4 | 58 +- parsing/pcoq.ml4 | 93 ++- parsing/pcoq.mli | 6 +- parsing/ppconstr.ml | 85 +-- parsing/ppconstr.mli | 6 +- parsing/pptactic.ml | 116 +-- parsing/pptactic.mli | 23 +- parsing/prettyp.ml | 6 - parsing/printer.ml | 154 ++-- parsing/printer.mli | 10 - parsing/q_coqast.ml4 | 5 +- parsing/search.ml | 2 +- parsing/tacextend.ml4 | 56 +- pretyping/classops.ml | 12 +- pretyping/detyping.ml | 2 +- pretyping/indrec.ml | 2 +- pretyping/tacred.ml | 2 +- pretyping/termops.ml | 24 +- pretyping/termops.mli | 5 +- proofs/tacexpr.ml | 2 +- scripts/coqmktop.ml | 14 +- states/MakeInitialNew.v | 10 + tactics/auto.ml | 8 +- tactics/eauto.ml4 | 31 +- tactics/extratactics.ml4 | 10 + tactics/hiddentac.ml | 2 +- tactics/setoid_replace.ml | 2 +- tactics/tacinterp.ml | 16 +- theories/Arith/Between.v | 1 + theories/Arith/Bool_nat.v | 1 + theories/Arith/Compare.v | 1 + theories/Arith/Compare_dec.v | 1 + theories/Arith/Div.v | 1 + theories/Arith/Div2.v | 1 + theories/Arith/EqNat.v | 1 + theories/Arith/Euclid.v | 1 + theories/Arith/Even.v | 1 + theories/Arith/Gt.v | 1 + theories/Arith/Le.v | 2 + theories/Arith/Lt.v | 1 + theories/Arith/Max.v | 1 + theories/Arith/Min.v | 1 + theories/Arith/Minus.v | 1 + theories/Arith/Mult.v | 1 + theories/Arith/Peano_dec.v | 1 + theories/Arith/Wf_nat.v | 2 + theories/Init/DatatypesSyntax.v | 11 +- theories/Init/LogicSyntax.v | 39 +- theories/Init/Logic_TypeSyntax.v | 28 +- theories/Init/PeanoSyntax.v | 24 +- theories/Init/SpecifSyntax.v | 55 +- theories/Init/Wf.v | 2 +- theories/Lists/PolyList.v | 3 +- theories/Lists/PolyListSyntax.v | 3 +- theories/Reals/Alembert.v | 3 +- theories/Reals/AltSeries.v | 1 + theories/Reals/ArithProp.v | 3 +- theories/Reals/Binomial.v | 1 + theories/Reals/Cauchy_prod.v | 3 +- theories/Reals/Cos_plus.v | 2 + theories/Reals/Cos_rel.v | 1 + theories/Reals/DiscrR.v | 1 + theories/Reals/Exp_prop.v | 1 + theories/Reals/MVT.v | 1 + theories/Reals/NewtonInt.v | 1 + theories/Reals/PSeries_reg.v | 1 + theories/Reals/PartSum.v | 1 + theories/Reals/RIneq.v | 5 + theories/Reals/RList.v | 1 + theories/Reals/R_sqr.v | 1 + theories/Reals/R_sqrt.v | 1 + theories/Reals/Ranalysis.v | 3 +- theories/Reals/Ranalysis1.v | 3 +- theories/Reals/Ranalysis2.v | 1 + theories/Reals/Ranalysis3.v | 2 +- theories/Reals/Ranalysis4.v | 1 + theories/Reals/Raxioms.v | 4 + theories/Reals/Rbasic_fun.v | 1 + theories/Reals/Rcomplete.v | 1 + theories/Reals/Rderiv.v | 1 + theories/Reals/Rfunctions.v | 2 +- theories/Reals/Rgeom.v | 2 +- theories/Reals/RiemannInt.v | 83 +-- theories/Reals/RiemannInt_SF.v | 1 + theories/Reals/Rlimit.v | 1 + theories/Reals/Rpower.v | 1 + theories/Reals/Rprod.v | 1 + theories/Reals/Rsigma.v | 1 + theories/Reals/Rsqrt_def.v | 1 + theories/Reals/Rsyntax.v | 67 +- theories/Reals/Rtopology.v | 4 +- theories/Reals/Rtrigo.v | 1 + theories/Reals/Rtrigo_alt.v | 1 + theories/Reals/Rtrigo_calc.v | 1 + theories/Reals/Rtrigo_def.v | 1 + theories/Reals/Rtrigo_reg.v | 1 + theories/Reals/SeqProp.v | 1 + theories/Reals/SeqSeries.v | 1 + theories/Reals/Sqrt_reg.v | 1 + .../Wellfounded/Lexicographic_Exponentiation.v | 8 +- theories/ZArith/Wf_Z.v | 1 + theories/ZArith/ZArith_dec.v | 2 + theories/ZArith/Zcomplements.v | 1 + theories/ZArith/Zdiv.v | 23 +- theories/ZArith/Zlogarithm.v | 1 + theories/ZArith/Zmisc.v | 26 +- theories/ZArith/Zpower.v | 5 +- theories/ZArith/Zsqrt.v | 1 + theories/ZArith/Zsyntax.v | 66 +- theories/ZArith/Zwf.v | 1 + theories/ZArith/fast_integer.v | 8 +- toplevel/class.ml | 2 +- toplevel/coqtop.ml | 6 +- toplevel/himsg.ml | 2 +- toplevel/metasyntax.ml | 133 ++-- toplevel/metasyntax.mli | 8 +- toplevel/mltop.ml4 | 7 +- toplevel/recordobj.ml | 2 +- toplevel/toplevel.ml | 5 +- toplevel/vernac.ml | 39 +- toplevel/vernacentries.ml | 13 +- toplevel/vernacexpr.ml | 12 +- translate/ppconstrnew.ml | 381 +++++----- translate/ppconstrnew.mli | 21 +- translate/pptacticnew.ml | 612 +++++++-------- translate/pptacticnew.mli | 22 +- translate/ppvernacnew.ml | 542 ++++++++++---- 180 files changed, 4695 insertions(+), 1549 deletions(-) create mode 100644 parsing/g_constrnew.ml4 create mode 100644 parsing/g_ltacnew.ml4 create mode 100644 parsing/g_natsyntaxnew.ml create mode 100644 parsing/g_natsyntaxnew.mli create mode 100644 parsing/g_primnew.ml4 create mode 100644 parsing/g_proofsnew.ml4 create mode 100644 parsing/g_rsyntaxnew.ml create mode 100644 parsing/g_tacticnew.ml4 create mode 100644 parsing/g_vernacnew.ml4 create mode 100644 parsing/g_zsyntaxnew.ml create mode 100644 parsing/g_zsyntaxnew.mli create mode 100644 states/MakeInitialNew.v diff --git a/contrib/field/field.ml4 b/contrib/field/field.ml4 index 49a187caa..c67dea9aa 100644 --- a/contrib/field/field.ml4 +++ b/contrib/field/field.ml4 @@ -107,6 +107,19 @@ let _ = Tacinterp.add_genarg_interp "minus_div_arg" (in_gen (wit_pair (wit_opt rawwit_constr) (wit_opt rawwit_constr)) (out_gen rawwit_minus_div_arg x)))))) +open Ppconstrnew +let pp_minus_div_arg (omin,odiv) = str "still no printer for minus_div_arg" +let pp_raw_minus_div_arg (omin,odiv) = + if omin=None && odiv=None then mt() else + spc() ++ str "with" ++ + pr_opt (fun c -> str "minus := " ++ pr_constr c) omin ++ + pr_opt (fun c -> str "div := " ++ pr_constr c) odiv + +let () = + Pptactic.declare_extra_genarg_pprule true + (rawwit_minus_div_arg,pp_raw_minus_div_arg) + (wit_minus_div_arg,pp_minus_div_arg) + open Pcoq.Constr GEXTEND Gram GLOBAL: minus_div_arg; diff --git a/contrib/interface/centaur.ml4 b/contrib/interface/centaur.ml4 index c9197ab34..87f8c8af1 100644 --- a/contrib/interface/centaur.ml4 +++ b/contrib/interface/centaur.ml4 @@ -259,7 +259,7 @@ let filter_by_module_from_varg_list l = let add_search (global_reference:global_reference) assumptions cstr = try let id_string = - string_of_qualid (Nametab.shortest_qualid_of_global None + string_of_qualid (Nametab.shortest_qualid_of_global Idset.empty global_reference) in let ast = try @@ -323,10 +323,10 @@ let globcv x = match x with | Node(_,"MUTIND", (Path(_,sp))::(Num(_,tyi))::_) -> convert_qualid - (Nametab.shortest_qualid_of_global None (IndRef(sp,tyi))) + (Nametab.shortest_qualid_of_global Idset.empty (IndRef(sp,tyi))) | Node(_,"MUTCONSTRUCT",(Path(_,sp))::(Num(_,tyi))::(Num(_,i))::_) -> convert_qualid - (Nametab.shortest_qualid_of_global None + (Nametab.shortest_qualid_of_global Idset.empty (ConstructRef ((sp, tyi), i))) | _ -> failwith "globcv : unexpected value";; diff --git a/contrib/interface/name_to_ast.ml b/contrib/interface/name_to_ast.ml index 4150a0948..5bd871bc9 100644 --- a/contrib/interface/name_to_ast.ml +++ b/contrib/interface/name_to_ast.ml @@ -114,7 +114,7 @@ let convert_one_inductive sp tyi = let (ref, params, arity, cstrnames, cstrtypes) = build_inductive sp tyi in let env = Global.env () in let envpar = push_rel_context params env in - let sp = sp_of_global None (IndRef (sp, tyi)) in + let sp = sp_of_global (IndRef (sp, tyi)) in (basename sp, convert_env(List.rev params), (extern_constr true envpar arity), diff --git a/contrib/interface/pbp.ml b/contrib/interface/pbp.ml index 4fb3fbe38..b5d6601c6 100644 --- a/contrib/interface/pbp.ml +++ b/contrib/interface/pbp.ml @@ -158,7 +158,7 @@ let make_pbp_atomic_tactic = function | PbpTryAssumption (Some a) -> TacTry (TacAtom (zz, TacExact (make_var a))) | PbpExists x -> - TacAtom (zz, TacSplit (ImplicitBindings [make_pbp_pattern x])) + TacAtom (zz, TacSplit (true,ImplicitBindings [make_pbp_pattern x])) | PbpGeneralize (h,args) -> let l = List.map make_pbp_pattern args in TacAtom (zz, TacGeneralize [make_app (make_var h) l]) @@ -176,7 +176,7 @@ let make_pbp_atomic_tactic = function (zz, TacElim ((make_var hyp_name,ExplicitBindings bind),None)) | PbpTryClear l -> TacTry (TacAtom (zz, TacClear (List.map (fun s -> AN s) l))) - | PbpSplit -> TacAtom (zz, TacSplit NoBindings);; + | PbpSplit -> TacAtom (zz, TacSplit (false,NoBindings));; let rec make_pbp_tactic = function | PbpThen tl -> make_then (List.map make_pbp_atomic_tactic tl) diff --git a/contrib/interface/showproof.ml b/contrib/interface/showproof.ml index 4ae1f280d..9ff567588 100644 --- a/contrib/interface/showproof.ml +++ b/contrib/interface/showproof.ml @@ -1271,8 +1271,8 @@ let rec natural_ntree ig ntree = TacIntroPattern _ -> natural_intros ig lh g gs ltree | TacIntroMove _ -> natural_intros ig lh g gs ltree | TacFix (_,n) -> natural_fix ig lh g gs n ltree - | TacSplit NoBindings -> natural_split ig lh g gs ge [] ltree - | TacSplit(ImplicitBindings l) -> natural_split ig lh g gs ge l ltree + | TacSplit (_,NoBindings) -> natural_split ig lh g gs ge [] ltree + | TacSplit(_,ImplicitBindings l) -> natural_split ig lh g gs ge l ltree | TacGeneralize l -> natural_generalize ig lh g gs ge l ltree | TacRight _ -> natural_right ig lh g gs ltree | TacLeft _ -> natural_left ig lh g gs ltree diff --git a/contrib/interface/xlate.ml b/contrib/interface/xlate.ml index 494cf593f..3941d10c6 100644 --- a/contrib/interface/xlate.ml +++ b/contrib/interface/xlate.ml @@ -822,7 +822,8 @@ and xlate_tac = | TacIntroMove (None, None) -> CT_intro (CT_coerce_NONE_to_ID_OPT CT_none) | TacLeft bindl -> CT_left (xlate_bindings bindl) | TacRight bindl -> CT_right (xlate_bindings bindl) - | TacSplit bindl -> CT_split (xlate_bindings bindl) + | TacSplit (false,bindl) -> CT_split (xlate_bindings bindl) + | TacSplit (true,bindl) -> CT_exists (xlate_bindings bindl) | TacExtend (_,"Replace", [c1; c2]) -> let c1 = xlate_formula (out_gen rawwit_constr c1) in let c2 = xlate_formula (out_gen rawwit_constr c2) in @@ -1581,13 +1582,14 @@ let xlate_vernac = CT_coerce_NONE_to_STRING_OPT CT_none) | VernacRequire (_,_,([]|_::_::_)) -> xlate_error "TODO: general form of future Require" - | VernacRequireFrom (impexp, spec, filename) -> + | VernacRequireFrom (Some impexp, spec, filename) -> let ct_impexp, ct_spec = get_require_flags impexp spec and id = id_of_string (Filename.basename filename) in CT_require (ct_impexp, ct_spec, xlate_ident id, CT_coerce_STRING_to_STRING_OPT (CT_string filename)) + | VernacRequireFrom (None, _, _) -> xlate_error "TODO: Read Module" | VernacSyntax (phylum, l) -> xlate_error "SYNTAX not implemented" (*Two versions of the syntax node with and without the binder list. *) @@ -1607,7 +1609,7 @@ let xlate_vernac = | VernacSyntaxExtension _ -> xlate_error "Syntax Extension not implemented" - | VernacInfix (str_assoc, n, str, id, false, None) -> + | VernacInfix (str_assoc, n, str, id, false, _, None) -> CT_infix ( (match str_assoc with | Some Gramext.LeftA -> CT_lefta diff --git a/contrib/omega/coq_omega.ml b/contrib/omega/coq_omega.ml index 4373da462..651b808c0 100644 --- a/contrib/omega/coq_omega.ml +++ b/contrib/omega/coq_omega.ml @@ -151,7 +151,7 @@ let dest_const_apply t = | Ind isp -> IndRef isp | _ -> raise Destruct in - id_of_global None ref, args + id_of_global ref, args type result = | Kvar of string @@ -164,11 +164,11 @@ let destructurate t = (* let env = Global.env() in*) match kind_of_term c, args with | Const sp, args -> - Kapp (string_of_id (id_of_global None (ConstRef sp)),args) + Kapp (string_of_id (id_of_global (ConstRef sp)),args) | Construct csp , args -> - Kapp (string_of_id (id_of_global None (ConstructRef csp)), args) + Kapp (string_of_id (id_of_global (ConstructRef csp)), args) | Ind isp, args -> - Kapp (string_of_id (id_of_global None (IndRef isp)),args) + Kapp (string_of_id (id_of_global (IndRef isp)),args) | Var id,[] -> Kvar(string_of_id id) | Prod (Anonymous,typ,body), [] -> Kimp(typ,body) | Prod (Name _,_,_),[] -> error "Omega: Not a quantifier-free goal" diff --git a/contrib/ring/Ring_theory.v b/contrib/ring/Ring_theory.v index 1c48c4b45..08bb53955 100644 --- a/contrib/ring/Ring_theory.v +++ b/contrib/ring/Ring_theory.v @@ -25,8 +25,8 @@ Variable Azero : A. is a good choice. The proof of A_eq_prop is in this case easy. *) Variable Aeq : A -> A -> bool. -Infix 4 "+" Aplus. -Infix 4 "*" Amult. +Infix 4 "+" Aplus V8only 40 (left associativity). +Infix 4 "*" Amult V8only 30 (left associativity). Notation "0" := Azero. Notation "1" := Aone. @@ -146,11 +146,16 @@ Variable Azero : A. Variable Aopp : A -> A. Variable Aeq : A -> A -> bool. -Infix 4 "+" Aplus. -Infix 4 "*" Amult. +Infix 4 "+" Aplus V8only 40 (left associativity). +Infix 4 "*" Amult V8only 30 (left associativity). Notation "0" := Azero. Notation "1" := Aone. -Notation "- x" := (Aopp x) (at level 3). +Notation "- 0" := (Aopp Azero) (at level 3) + V8only (at level 40, left associativity). +Notation "- 1" := (Aopp Aone) (at level 3) + V8only (at level 40, left associativity). +Notation "- x" := (Aopp x) (at level 3) + V8only (at level 40, left associativity). Record Ring_Theory : Prop := { Th_plus_sym : (n,m:A) n + m == m + n; diff --git a/contrib/ring/Setoid_ring_theory.v b/contrib/ring/Setoid_ring_theory.v index 3456b1051..4fc7ea104 100644 --- a/contrib/ring/Setoid_ring_theory.v +++ b/contrib/ring/Setoid_ring_theory.v @@ -18,7 +18,8 @@ Section Setoid_rings. Variable A : Type. Variable Aequiv : A -> A -> Prop. -Infix "==" Aequiv (at level 5). +Infix "==" Aequiv (at level 5) + V8only (at level 50). Variable S : (Setoid_Theory A Aequiv). @@ -31,11 +32,16 @@ Variable Azero : A. Variable Aopp : A -> A. Variable Aeq : A -> A -> bool. -Infix "+" Aplus (at level 4). -Infix "*" Amult (at level 3). +Infix 4 "+" Aplus V8only 40 (left associativity). +Infix 4 "*" Amult V8only 30 (left associativity). Notation "0" := Azero. Notation "1" := Aone. -Notation "- x" := (Aopp x) (at level 3). +Notation "- 0" := (Aopp Azero) (at level 3) + V8only (at level 40, left associativity). +Notation "- 1" := (Aopp Aone) (at level 3) + V8only (at level 40, left associativity). +Notation "- x" := (Aopp x) (at level 3) + V8only (at level 40, left associativity). Variable plus_morph : (a,a0,a1,a2:A) a == a0 -> a1 == a2 -> a+a1 == a0+a2. Variable mult_morph : (a,a0,a1,a2:A) a == a0 -> a1 == a2 -> a*a1 == a0*a2. diff --git a/contrib/romega/ReflOmegaCore.v b/contrib/romega/ReflOmegaCore.v index 95e4a5f5a..ef35d3245 100644 --- a/contrib/romega/ReflOmegaCore.v +++ b/contrib/romega/ReflOmegaCore.v @@ -642,25 +642,46 @@ Recursive Tactic Definition loop t := ( | [(Topp ?1)] -> (loop ?1) | [(Tint ?1)] -> (loop ?1) (* Eliminations *) - | [(Case ?1 of ? ? ? ? ? ? ? ? ? end)] -> + | [(Cases ?1 of + | (EqTerm _ _) => ? + | (LeqTerm _ _) => ? + | TrueTerm => ? + | FalseTerm => ? + | (Tnot _) => ? + | (GeqTerm _ _) => ? + | (GtTerm _ _) => ? + | (LtTerm _ _) => ? + | (NeqTerm _ _) => ? + end)] -> (Case ?1; [ Intro; Intro | Intro; Intro | Idtac | Idtac | Intro | Intro; Intro | Intro; Intro | Intro; Intro | Intro; Intro ]); Auto; Simplify - | [(Case ?1 of ? ? ? ? ? ? end)] -> + | [(Cases ?1 of + (Tint _) => ? + | (Tplus _ _) => ? + | (Tmult _ _) => ? + | (Tminus _ _) => ? + | (Topp _) => ? + | (Tvar _) => ? + end)] -> (Case ?1; [ Intro | Intro; Intro | Intro; Intro | Intro; Intro | Intro | Intro ]); Auto; Simplify - | [(Case (Zcompare ?1 ?2) of ? ? ? end)] -> + | [(Cases (Zcompare ?1 ?2) of + EGAL => ? + | INFERIEUR => ? + | SUPERIEUR => ? + end)] -> (Elim_Zcompare ?1 ?2) ; Intro ; Auto; Simplify - | [(Case ?1 of ? ? ? end)] -> + | [(Cases ?1 of ZERO => ? | (POS _) => ? | (NEG _) => ? end)] -> (Case ?1; [ Idtac | Intro | Intro ]); Auto; Simplify - | [(Case (eq_Z ?1 ?2) of ? ? end)] -> + | [(if (eq_Z ?1 ?2) then ? else ?)] -> ((Elim_eq_Z ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]); Simpl; Auto; Simplify - | [(Case (eq_term ?1 ?2) of ? ? end)] -> + | [(if (eq_term ?1 ?2) then ? else ?)] -> ((Elim_eq_term ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]); Simpl; Auto; Simplify - | [(Case (eq_pos ?1 ?2) of ? ? end)] -> + | [(if (eq_pos ?1 ?2) then ? else ?)] -> ((Elim_eq_pos ?1 ?2); Intro H; [Rewrite H; Clear H | Clear H]); Simpl; Auto; Simplify | _ -> Fail) diff --git a/contrib/romega/const_omega.ml b/contrib/romega/const_omega.ml index 42b61a150..8b055e296 100644 --- a/contrib/romega/const_omega.ml +++ b/contrib/romega/const_omega.ml @@ -21,15 +21,15 @@ let destructurate t = match Term.kind_of_term c, args with | Term.Const sp, args -> Kapp (Names.string_of_id - (Nametab.id_of_global None (Libnames.ConstRef sp)), + (Nametab.id_of_global (Libnames.ConstRef sp)), args) | Term.Construct csp , args -> Kapp (Names.string_of_id - (Nametab.id_of_global None (Libnames.ConstructRef csp)), + (Nametab.id_of_global (Libnames.ConstructRef csp)), args) | Term.Ind isp, args -> Kapp (Names.string_of_id - (Nametab.id_of_global None (Libnames.IndRef isp)),args) + (Nametab.id_of_global (Libnames.IndRef isp)),args) | Term.Var id,[] -> Kvar(Names.string_of_id id) | Term.Prod (Names.Anonymous,typ,body), [] -> Kimp(typ,body) | Term.Prod (Names.Name _,_,_),[] -> @@ -47,7 +47,7 @@ let dest_const_apply t = | Term.Ind isp -> Libnames.IndRef isp | _ -> raise Destruct in - Nametab.id_of_global None ref, args + Nametab.id_of_global ref, args let recognize_number t = let rec loop t = diff --git a/contrib/xml/cic2acic.ml b/contrib/xml/cic2acic.ml index c0bdf07a3..b3d164122 100644 --- a/contrib/xml/cic2acic.ml +++ b/contrib/xml/cic2acic.ml @@ -254,11 +254,11 @@ print_endline "PASSATO" ; flush stdout ; match g with Libnames.ConstructRef ((induri,_),_) | Libnames.IndRef (induri,_) -> - Nametab.sp_of_global None (Libnames.IndRef (induri,0)) + Nametab.sp_of_global (Libnames.IndRef (induri,0)) | Libnames.VarRef id -> (* Invariant: variables are never cooked in Coq *) raise Not_found - | _ -> Nametab.sp_of_global None g + | _ -> Nametab.sp_of_global g in Dischargedhypsmap.get_discharged_hyps sp, get_module_path_of_section_path sp diff --git a/contrib/xml/xmlcommand.ml4 b/contrib/xml/xmlcommand.ml4 index e2b49a391..1ee1ef4ea 100644 --- a/contrib/xml/xmlcommand.ml4 +++ b/contrib/xml/xmlcommand.ml4 @@ -420,18 +420,18 @@ let print_term inner_types l env csr = >] ) | T.Const kn -> - let sp = Nt.sp_of_global None (L.ConstRef kn) in + let sp = Nt.sp_of_global (L.ConstRef kn) in X.xml_empty "CONST" (add_sort_attribute false ["uri",(uri_of_path sp Constant) ; "id", next_id]) | T.Ind (kn,i) -> - let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in + let sp = Nt.sp_of_global (L.IndRef(kn,0)) in X.xml_empty "MUTIND" ["uri",(uri_of_path sp Inductive) ; "noType",(string_of_int i) ; "id", next_id] | T.Construct ((kn,i),j) -> - let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in + let sp = Nt.sp_of_global (L.IndRef(kn,0)) in X.xml_empty "MUTCONSTRUCT" (add_sort_attribute false ["uri",(uri_of_path sp Inductive) ; @@ -439,7 +439,7 @@ let print_term inner_types l env csr = "noConstr",(string_of_int j) ; "id", next_id]) | T.Case ({T.ci_ind=(kn,i)},ty,term,a) -> - let sp = Nt.sp_of_global None (L.IndRef(kn,0)) in + let sp = Nt.sp_of_global (L.IndRef(kn,0)) in let (uri, typeno) = (uri_of_path sp Inductive),i in X.xml_nempty "MUTCASE" (add_sort_attribute true @@ -728,7 +728,7 @@ let print (_,qid as locqid) fn = let (_,body,typ) = G.lookup_named id in sp,Variable,print_variable id body (T.body_of_type typ) env inner_types | Ln.ConstRef kn -> - let sp = Nt.sp_of_global None glob_ref in + let sp = Nt.sp_of_global glob_ref in let {D.const_body=val0 ; D.const_type = typ ; D.const_hyps = hyps} = G.lookup_constant kn in let hyps = string_list_of_named_context_list hyps in @@ -742,7 +742,7 @@ let print (_,qid as locqid) fn = print_definition id c typ [] hyps env inner_types end | Ln.IndRef (kn,_) -> - let sp = Nt.sp_of_global None (Ln.IndRef(kn,0)) in + let sp = Nt.sp_of_global (Ln.IndRef(kn,0)) in let {D.mind_packets=packs ; D.mind_hyps=hyps; D.mind_finite=finite} = G.lookup_mind kn in diff --git a/contrib/xml/xmlentries.ml4 b/contrib/xml/xmlentries.ml4 index bcfcbd2ff..2807a3d6e 100644 --- a/contrib/xml/xmlentries.ml4 +++ b/contrib/xml/xmlentries.ml4 @@ -50,7 +50,7 @@ END let pr_filename = function Some fn -> str " File" ++ str fn | None -> mt () let _ = - Pptactic.declare_extra_genarg_pprule + Pptactic.declare_extra_genarg_pprule true (rawwit_filename, pr_filename) (wit_filename, pr_filename) @@ -76,7 +76,7 @@ open Pp let pr_diskname = function Some fn -> str " Disk" ++ str fn | None -> mt () let _ = - Pptactic.declare_extra_genarg_pprule + Pptactic.declare_extra_genarg_pprule true (rawwit_diskname, pr_diskname) (wit_diskname, pr_diskname) diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 7035bfdfa..8cbddd87c 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -117,8 +117,8 @@ let raw_string_of_ref = function | VarRef id -> "SECVAR("^(string_of_id id)^")" -let extern_reference loc r = - try Qualid (loc,shortest_qualid_of_global None r) +let extern_reference loc vars r = + try Qualid (loc,shortest_qualid_of_global vars r) with Not_found -> (* happens in debugger *) Ident (loc,id_of_string (raw_string_of_ref r)) @@ -126,7 +126,7 @@ let extern_reference loc r = (**********************************************************************) (* conversion of terms and cases patterns *) -let rec extern_cases_pattern_in_scope scopes pat = +let rec extern_cases_pattern_in_scope scopes vars pat = try if !print_no_symbol then raise No_match; let (sc,n) = Symbols.uninterp_cases_numeral pat in @@ -138,8 +138,9 @@ let rec extern_cases_pattern_in_scope scopes pat = | PatVar (_,Name id) -> CPatAtom (loc,Some (Ident (loc,id))) | PatVar (_,Anonymous) -> CPatAtom (loc, None) | PatCstr(_,cstrsp,args,na) -> - let args = List.map (extern_cases_pattern_in_scope scopes) args in - let p = CPatCstr (loc,extern_reference loc (ConstructRef cstrsp),args) in + let args = List.map (extern_cases_pattern_in_scope scopes vars) args in + let p = CPatCstr + (loc,extern_reference loc vars (ConstructRef cstrsp),args) in (match na with | Name id -> CPatAlias (loc,p,id) | Anonymous -> p) @@ -187,16 +188,16 @@ let extern_app impl f args = else explicitize impl (CRef f) args -let rec extern_args extern scopes args subscopes = +let rec extern_args extern scopes env args subscopes = match args with | [] -> [] | a::args -> let argscopes, subscopes = match subscopes with | [] -> scopes, [] | scopt::subscopes -> option_cons scopt scopes, subscopes in - extern argscopes a :: extern_args extern scopes args subscopes + extern argscopes env a :: extern_args extern scopes env args subscopes -let rec extern scopes r = +let rec extern scopes vars r = try if !print_no_symbol then raise No_match; extern_numeral scopes r (Symbols.uninterp_numeral r) @@ -204,10 +205,10 @@ let rec extern scopes r = try if !print_no_symbol then raise No_match; - extern_symbol scopes r (Symbols.uninterp_notations r) + extern_symbol scopes vars r (Symbols.uninterp_notations r) with No_match -> match r with - | RRef (_,ref) -> CRef (extern_reference loc ref) + | RRef (_,ref) -> CRef (extern_reference loc vars ref) | RVar (_,id) -> CRef (Ident (loc,id)) @@ -222,52 +223,59 @@ let rec extern scopes r = | REvar (loc,ev) -> extern_evar loc ev (* we drop args *) | RRef (loc,ref) -> let subscopes = Symbols.find_arguments_scope ref in - let args = extern_args extern scopes args subscopes in - extern_app (implicits_of_global ref) (extern_reference loc ref) + let args = extern_args extern scopes vars args subscopes in + extern_app (implicits_of_global ref) + (extern_reference loc vars ref) args | _ -> - explicitize [] (extern scopes f) (List.map (extern scopes) args)) + explicitize [] (extern scopes vars f) + (List.map (extern scopes vars) args)) | RProd (_,Anonymous,t,c) -> (* Anonymous product are never factorized *) - CArrow (loc,extern scopes t,extern scopes c) + CArrow (loc,extern scopes vars t, extern scopes vars c) | RLetIn (_,na,t,c) -> - CLetIn (loc,(loc,na),extern scopes t,extern scopes c) + CLetIn (loc,(loc,na),extern scopes vars t, + extern scopes (add_vname vars na) c) | RProd (_,na,t,c) -> - let t = extern scopes t in - let (idl,c) = factorize_prod scopes t c in + let t = extern scopes vars t in + let (idl,c) = factorize_prod scopes (add_vname vars na) t c in CProdN (loc,[(loc,na)::idl,t],c) | RLambda (_,na,t,c) -> - let t = extern scopes t in - let (idl,c) = factorize_lambda scopes t c in + let t = extern scopes vars t in + let (idl,c) = factorize_lambda scopes (add_vname vars na) t c in CLambdaN (loc,[(loc,na)::idl,t],c) | RCases (_,typopt,tml,eqns) -> - let pred = option_app (extern scopes) typopt in - let tml = List.map (extern scopes) tml in - let eqns = List.map (extern_eqn scopes) eqns in + let pred = option_app (extern scopes vars) typopt in + let tml = List.map (extern scopes vars) tml in + let eqns = List.map (extern_eqn scopes vars) eqns in CCases (loc,pred,tml,eqns) | ROrderedCase (_,cs,typopt,tm,bv) -> - let bv = Array.to_list (Array.map (extern scopes) bv) in + let bv = Array.to_list (Array.map (extern scopes vars) bv) in COrderedCase - (loc,cs,option_app (extern scopes) typopt,extern scopes tm,bv) + (loc,cs,option_app (extern scopes vars) typopt, + extern scopes vars tm,bv) | RRec (_,fk,idv,tyv,bv) -> + let vars' = Array.fold_right Idset.add idv vars in (match fk with | RFix (nv,n) -> let listdecl = Array.mapi (fun i fi -> - (fi,nv.(i),extern scopes tyv.(i),extern scopes bv.(i))) idv + (fi,nv.(i),extern scopes vars tyv.(i), + extern scopes vars' bv.(i))) idv in CFix (loc,(loc,idv.(n)),Array.to_list listdecl) | RCoFix n -> let listdecl = Array.mapi (fun i fi -> - (fi,extern scopes tyv.(i),extern scopes bv.(i))) idv + (fi,extern scopes vars tyv.(i), + extern scopes vars' bv.(i))) idv in CCoFix (loc,(loc,idv.(n)),Array.to_list listdecl)) @@ -280,33 +288,36 @@ let rec extern scopes r = | RHole (_,e) -> CHole loc - | RCast (_,c,t) -> CCast (loc,extern scopes c,extern scopes t) + | RCast (_,c,t) -> CCast (loc,extern scopes vars c,extern scopes vars t) | RDynamic (_,d) -> CDynamic (loc,d) -and factorize_prod scopes aty = function +and factorize_prod scopes vars aty = function | RProd (_,Name id,ty,c) - when aty = extern scopes ty + when aty = extern scopes vars ty & not (occur_var_constr_expr id aty) (* avoid na in ty escapes scope *) - -> let (nal,c) = factorize_prod scopes aty c in ((loc,Name id)::nal,c) - | c -> ([],extern scopes c) + -> let (nal,c) = factorize_prod scopes (Idset.add id vars) aty c in + ((loc,Name id)::nal,c) + | c -> ([],extern scopes vars c) -and factorize_lambda scopes aty = function +and factorize_lambda scopes vars aty = function | RLambda (_,na,ty,c) - when aty = extern scopes ty + when aty = extern scopes vars ty & not (occur_name na aty) (* To avoid na in ty' escapes scope *) - -> let (nal,c) = factorize_lambda scopes aty c in ((loc,na)::nal,c) - | c -> ([],extern scopes c) + -> let (nal,c) = factorize_lambda scopes (add_vname vars na) aty c in + ((loc,na)::nal,c) + | c -> ([],extern scopes vars c) -and extern_eqn scopes (loc,ids,pl,c) = - (loc,List.map (extern_cases_pattern_in_scope scopes) pl,extern scopes c) +and extern_eqn scopes vars (loc,ids,pl,c) = + (loc,List.map (extern_cases_pattern_in_scope scopes vars) pl, + extern scopes vars c) and extern_numeral scopes t (sc,n) = match Symbols.availability_of_numeral sc scopes with | None -> raise No_match | Some key -> insert_delimiters (CNumeral (loc,n)) key -and extern_symbol scopes t = function +and extern_symbol scopes vars t = function | [] -> raise No_match | (keyrule,pat,n as rule)::rules -> try @@ -329,34 +340,36 @@ and extern_symbol scopes t = function | Some (scopt,key) -> let scopes = option_cons scopt scopes in let l = - List.map (fun (c,scl) -> extern (scl@scopes) c) subst in + List.map (fun (c,scl) -> extern (scl@scopes) vars c) + subst in insert_delimiters (CNotation (loc,ntn,l)) key) | SynDefRule kn -> CRef (Qualid (loc, shortest_qualid_of_syndef kn)) in if args = [] then e - else explicitize [] e (List.map (extern scopes) args) + else explicitize [] e (List.map (extern scopes vars) args) with - No_match -> extern_symbol scopes t rules + No_match -> extern_symbol scopes vars t rules -let extern_rawconstr = - extern (Symbols.current_scopes()) +let extern_rawconstr vars c = + extern (Symbols.current_scopes()) vars c -let extern_cases_pattern = - extern_cases_pattern_in_scope (Symbols.current_scopes()) +let extern_cases_pattern vars p = + extern_cases_pattern_in_scope (Symbols.current_scopes()) vars p (******************************************************************) (* Main translation function from constr -> constr_expr *) let extern_constr at_top env t = + let vars = vars_of_env env in let avoid = if at_top then ids_of_context env else [] in - extern (Symbols.current_scopes()) + extern (Symbols.current_scopes()) vars (Detyping.detype env avoid (names_of_rel_context env) t) (******************************************************************) (* Main translation function from pattern -> constr_expr *) -let rec extern_pattern tenv env = function - | PRef ref -> CRef (extern_reference loc ref) +let rec extern_pattern tenv vars env = function + | PRef ref -> CRef (extern_reference loc vars ref) | PVar id -> CRef (Ident (loc,id)) @@ -379,47 +392,53 @@ let rec extern_pattern tenv env = function let (f,args) = skip_coercion (function PRef r -> Some r | _ -> None) (f,Array.to_list args) in - let args = List.map (extern_pattern tenv env) args in + let args = List.map (extern_pattern tenv vars env) args in (match f with | PRef ref -> - extern_app (implicits_of_global ref) (extern_reference loc ref) - args - | _ -> explicitize [] (extern_pattern tenv env f) args) + extern_app (implicits_of_global ref) + (extern_reference loc vars ref) + args + | _ -> explicitize [] (extern_pattern tenv vars env f) args) | PSoApp (n,args) -> - let args = List.map (extern_pattern tenv env) args in + let args = List.map (extern_pattern tenv vars env) args in (* [-n] is the trick to embed a so patten into a regular application *) (* see constrintern.ml and g_constr.ml4 *) explicitize [] (CMeta (loc,-n)) args | PProd (Anonymous,t,c) -> (* Anonymous product are never factorized *) - CArrow (loc,extern_pattern tenv env t,extern_pattern tenv env c) + CArrow (loc,extern_pattern tenv vars env t, + extern_pattern tenv vars env c) | PLetIn (na,t,c) -> - CLetIn (loc,(loc,na),extern_pattern tenv env t,extern_pattern tenv env c) + CLetIn (loc,(loc,na),extern_pattern tenv vars env t, + extern_pattern tenv (add_vname vars na) (na::env) c) | PProd (na,t,c) -> - let t = extern_pattern tenv env t in - let (idl,c) = factorize_prod_pattern tenv (add_name na env) t c in + let t = extern_pattern tenv vars env t in + let (idl,c) = + factorize_prod_pattern tenv (add_vname vars na) (na::env) t c in CProdN (loc,[(loc,na)::idl,t],c) | PLambda (na,t,c) -> - let t = extern_pattern tenv env t in - let (idl,c) = factorize_lambda_pattern tenv (add_name na env) t c in + let t = extern_pattern tenv vars env t in + let (idl,c) = + factorize_lambda_pattern tenv (add_vname vars na) (na::env) t c in CLambdaN (loc,[(loc,na)::idl,t],c) | PCase (cs,typopt,tm,bv) -> - let bv = Array.to_list (Array.map (extern_pattern tenv env) bv) in + let bv = Array.to_list (Array.map (extern_pattern tenv vars env) bv) in COrderedCase - (loc,cs,option_app (extern_pattern tenv env) typopt, - extern_pattern tenv env tm,bv) + (loc,cs,option_app (extern_pattern tenv vars env) typopt, + extern_pattern tenv vars env tm,bv) | PFix f -> - extern (Symbols.current_scopes()) (Detyping.detype tenv [] env (mkFix f)) + extern (Symbols.current_scopes()) vars + (Detyping.detype tenv [] env (mkFix f)) | PCoFix c -> - extern (Symbols.current_scopes()) + extern (Symbols.current_scopes()) vars (Detyping.detype tenv [] env (mkCoFix c)) | PSort s -> @@ -429,18 +448,23 @@ let rec extern_pattern tenv env = function | RType _ -> RType None in CSort (loc,s) -and factorize_prod_pattern tenv env aty = function +and factorize_prod_pattern tenv vars env aty = function | PProd (Name id as na,ty,c) - when aty = extern_pattern tenv env ty - & not (occur_var_constr_expr id aty) (*To avoid na in ty escapes scope*) - -> let (nal,c) = factorize_prod_pattern tenv (na::env) aty c in + when aty = extern_pattern tenv vars env ty + & not (occur_var_constr_expr id aty) (*To avoid na in ty escapes scope*) + -> let (nal,c) = + factorize_prod_pattern tenv (add_vname vars na) (na::env) aty c in ((loc,Name id)::nal,c) - | c -> ([],extern_pattern tenv env c) + | c -> ([],extern_pattern tenv vars env c) -and factorize_lambda_pattern tenv env aty = function +and factorize_lambda_pattern tenv vars env aty = function | PLambda (na,ty,c) - when aty = extern_pattern tenv env ty + when aty = extern_pattern tenv vars env ty & not (occur_name na aty) (* To avoid na in ty' escapes scope *) - -> let (nal,c) = factorize_lambda_pattern tenv (add_name na env) aty c - in ((loc,na)::nal,c) - | c -> ([],extern_pattern tenv env c) + -> let (nal,c) = + factorize_lambda_pattern tenv (add_vname vars na) (na::env) aty c + in ((loc,na)::nal,c) + | c -> ([],extern_pattern tenv vars env c) + +let extern_pattern tenv env pat = + extern_pattern tenv (vars_of_env tenv) env pat diff --git a/interp/constrextern.mli b/interp/constrextern.mli index 2aa16c8c3..b2dd833aa 100644 --- a/interp/constrextern.mli +++ b/interp/constrextern.mli @@ -25,15 +25,15 @@ open Symbols (* Translation of pattern, cases pattern, rawterm and term into syntax trees for printing *) -val extern_cases_pattern : cases_pattern -> cases_pattern_expr -val extern_rawconstr : rawconstr -> constr_expr +val extern_cases_pattern : Idset.t -> cases_pattern -> cases_pattern_expr +val extern_rawconstr : Idset.t -> rawconstr -> constr_expr val extern_pattern : env -> names_context -> constr_pattern -> constr_expr (* If [b=true] in [extern_constr b env c] then the variables in the first level of quantification clashing with the variables in [env] are renamed *) val extern_constr : bool -> env -> constr -> constr_expr -val extern_reference : loc -> global_reference -> reference +val extern_reference : loc -> Idset.t -> global_reference -> reference (* For debugging *) val print_implicits : bool ref diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 74740abce..e7db4afe4 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -105,7 +105,7 @@ let add_glob loc ref = in let s = string_of_dirpath (find_module dir) in i*) - let sp = Nametab.sp_of_global None ref in + let sp = Nametab.sp_of_global ref in let id = let _,id = repr_path sp in string_of_id id in let dp = string_of_dirpath (Declare.library_part ref) in dump_string (Printf.sprintf "R%d %s.%s\n" (fst loc) dp id) @@ -418,6 +418,8 @@ let internalise isarity sigma env allow_soapp lvar c = | CLetIn (loc,(_,na),c1,c2) -> RLetIn (loc, na, intern false env c1, intern false (name_fold Idset.add na ids,impls,scopes) c2) + | CNotation (loc,"- _",[CNumeral(_,Bignat.POS p)]) -> + Symbols.interp_numeral loc (Bignat.NEG p) scopes | CNotation (loc,ntn,args) -> let scopes = if isarity then Symbols.type_scope::scopes else scopes in let (ids,c) = Symbols.interp_notation ntn scopes in @@ -665,10 +667,38 @@ let rec pat_of_raw metas vars lvar = function PCase (st,option_app (pat_of_raw metas vars lvar) po, pat_of_raw metas vars lvar c, Array.map (pat_of_raw metas vars lvar) br) + | RCases (loc,po,[c],br) -> + PCase (Term.RegularStyle,option_app (pat_of_raw metas vars lvar) po, + pat_of_raw metas vars lvar c, + Array.init (List.length br) + (pat_of_raw_branch loc metas vars lvar br)) | r -> let loc = loc_of_rawconstr r in user_err_loc (loc,"pattern_of_rawconstr", str "Not supported pattern") +and pat_of_raw_branch loc metas vars lvar brs i = + let bri = List.filter + (function + (_,_,[PatCstr(_,c,lv,_)],_) -> snd c = i+1 + | (loc,_,_,_) -> + user_err_loc (loc,"pattern_of_rawconstr", + str "Not supported pattern")) brs in + match bri with + [(_,_,[PatCstr(_,_,lv,_)],br)] -> + let lna = + List.map + (function PatVar(_,na) -> na + | PatCstr(loc,_,_,_) -> + user_err_loc (loc,"pattern_of_rawconstr", + str "Not supported pattern")) lv in + let vars' = List.rev lna @ vars in + List.fold_right (fun na b -> PLambda(na,PMeta None,b)) lna + (pat_of_raw metas vars' lvar br) + | _ -> user_err_loc (loc,"pattern_of_rawconstr", + str "No unique branch for " ++ int (i+1) ++ + str"-th constructor") + + let pattern_of_rawconstr lvar c = let metas = ref [] in let p = pat_of_raw metas [] lvar c in diff --git a/interp/topconstr.ml b/interp/topconstr.ml index aa78ec5e0..17b324937 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -217,9 +217,9 @@ let rec match_ alp metas sigma a1 a2 = match (a1,a2) with | RCast(_,c1,t1), ACast(c2,t2) -> match_ alp metas (match_ alp metas sigma c1 c2) t1 t2 | RSort (_,s1), ASort s2 when s1 = s2 -> sigma - | RHole _, a -> sigma | RMeta _, AHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match - | a, AHole _ -> sigma + | a, AHole _ when not(Options.do_translate()) -> sigma + | RHole _, AHole _ -> sigma | (RDynamic _ | RRec _ | REvar _), _ | _,_ -> raise No_match diff --git a/kernel/names.ml b/kernel/names.ml index ddc592c09..09971cc56 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -17,9 +17,16 @@ type identifier = string let id_ord = Pervasives.compare -let string_of_id id = String.copy id let id_of_string s = String.copy s +let map_ident id = + if Options.do_translate() then + match id with + "fix" -> "Fix" + | _ -> id + else id +let string_of_id id = String.copy (map_ident id) + (* Hash-consing of identifier *) module Hident = Hashcons.Make( struct diff --git a/lib/options.ml b/lib/options.ml index f2cc7a09d..90517561f 100644 --- a/lib/options.ml +++ b/lib/options.ml @@ -25,10 +25,11 @@ let term_quality = ref false let xml_export = ref false let v7 = ref true +let v7_only = ref false (* Translate *) let translate = ref false -let make_translate f = translate := f; () +let make_translate f = translate := f; v7 := f; () let do_translate () = !translate let translate_file = ref false diff --git a/lib/options.mli b/lib/options.mli index af08f689e..b6b895178 100644 --- a/lib/options.mli +++ b/lib/options.mli @@ -24,6 +24,7 @@ val term_quality : bool ref val xml_export : bool ref val v7 : bool ref +val v7_only : bool ref val translate : bool ref val make_translate : bool -> unit diff --git a/lib/pp_control.ml b/lib/pp_control.ml index 3b1266c2c..44f351458 100644 --- a/lib/pp_control.ml +++ b/lib/pp_control.ml @@ -22,7 +22,7 @@ let dflt_gp = { margin = 72; max_indent = 62; max_depth = 50; - ellipsis = "." } + ellipsis = ".." } (* A deeper pretty-printer to print proof scripts *) @@ -30,7 +30,7 @@ let deep_gp = { margin = 72; max_indent = 62; max_depth = 10000; - ellipsis = "." } + ellipsis = "..." } (* set_gp : Format.formatter -> pp_global_params -> unit * set the parameters of a formatter *) diff --git a/library/declare.ml b/library/declare.ml index 6999ee669..d5dc01cf5 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -431,7 +431,7 @@ let strength_of_global = function | IndRef _ | ConstructRef _ | ConstRef _ -> Global let library_part ref = - let sp = Nametab.sp_of_global None ref in + let sp = Nametab.sp_of_global ref in let dir,_ = repr_path sp in match strength_of_global ref with | Local -> diff --git a/library/goptions.ml b/library/goptions.ml index 4c2d15206..b1868fdd7 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -309,7 +309,7 @@ let msg_option_value (name,v) = | IntValue (Some n) -> int n | IntValue None -> str "undefined" | StringValue s -> str s - | IdentValue r -> pr_global_env None r + | IdentValue r -> pr_global_env Idset.empty r let print_option_value key = let (name,(_,read,_)) = get_option key in diff --git a/library/library.mli b/library/library.mli index d1d345aab..d69ecef9e 100644 --- a/library/library.mli +++ b/library/library.mli @@ -46,7 +46,7 @@ val require_library : bool option -> qualid located list -> bool -> unit val require_library_from_file : - bool option -> identifier option -> string -> bool -> unit + bool option -> identifier option -> System.physical_path -> bool -> unit (*s [save_library_to s f] saves the current environment as a library [s] in the file [f]. *) diff --git a/library/nametab.ml b/library/nametab.ml index cbb3b23e9..372590911 100755 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -70,7 +70,7 @@ module type NAMETREE = sig val find : user_name -> 'a t -> 'a val exists : user_name -> 'a t -> bool val user_name : qualid -> 'a t -> user_name - val shortest_qualid : user_name -> 'a t -> qualid + val shortest_qualid : Idset.t -> user_name -> 'a t -> qualid end module Make(U:UserName) : NAMETREE with type user_name = U.t @@ -207,15 +207,17 @@ let exists uname tab = with Not_found -> false -let shortest_qualid uname tab = +let shortest_qualid ctx uname tab = + let id,dir = U.repr uname in + let hidden = Idset.mem id ctx in let rec find_uname pos dir (path,tab) = match path with - | Absolute (u,_) | Relative (u,_) when u=uname -> List.rev pos + | Absolute (u,_) | Relative (u,_) + when u=uname && not(pos=[] && hidden) -> List.rev pos | _ -> match dir with [] -> raise Not_found | id::dir -> find_uname (id::pos) dir (ModIdmap.find id tab) in - let id,dir = U.repr uname in let ptab = Idmap.find id tab in let found_dir = find_uname [] dir ptab in make_qualid (make_dirpath found_dir) id @@ -413,16 +415,14 @@ let exists_modtype sp = SpTab.exists sp !the_modtypetab (* Reverse locate functions ***********************************************) -let sp_of_global ctx_opt ref = - match (ctx_opt,ref) with - | Some ctx, VarRef id -> - let _ = Sign.lookup_named id ctx in - make_path empty_dirpath id +let sp_of_global ref = + match ref with + | VarRef id -> make_path empty_dirpath id | _ -> Globrevtab.find (TrueGlobal ref) !the_globrevtab -let id_of_global ctx_opt ref = - let (_,id) = repr_path (sp_of_global ctx_opt ref) in +let id_of_global ref = + let (_,id) = repr_path (sp_of_global ref) in id let sp_of_syntactic_definition kn = @@ -434,21 +434,24 @@ let dir_of_mp mp = (* Shortest qualid functions **********************************************) -let shortest_qualid_of_global ctx_opt ref = - let sp = sp_of_global ctx_opt ref in - SpTab.shortest_qualid sp !the_ccitab +let shortest_qualid_of_global ctx ref = + match ref with + | VarRef id -> make_qualid empty_dirpath id + | _ -> + let sp = Globrevtab.find (TrueGlobal ref) !the_globrevtab in + SpTab.shortest_qualid ctx sp !the_ccitab let shortest_qualid_of_syndef kn = let sp = sp_of_syntactic_definition kn in - SpTab.shortest_qualid sp !the_ccitab + SpTab.shortest_qualid Idset.empty sp !the_ccitab let shortest_qualid_of_module mp = let dir = MPmap.find mp !the_modrevtab in - DirTab.shortest_qualid dir !the_dirtab + DirTab.shortest_qualid Idset.empty dir !the_dirtab let shortest_qualid_of_modtype kn = let sp = KNmap.find kn !the_modtyperevtab in - SpTab.shortest_qualid sp !the_modtypetab + SpTab.shortest_qualid Idset.empty sp !the_modtypetab let pr_global_env env ref = (* Il est important de laisser le let-in, car les streams s'évaluent diff --git a/library/nametab.mli b/library/nametab.mli index f14b1a123..9ee45a9ad 100755 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -133,23 +133,20 @@ val full_name_module : qualid -> dir_path (*s Reverse lookup -- finding user names corresponding to the given internal name *) - val sp_of_syntactic_definition : kernel_name -> section_path - -val sp_of_global : - Sign.named_context option -> global_reference -> section_path val shortest_qualid_of_global : - Sign.named_context option -> global_reference -> qualid + Idset.t -> global_reference -> qualid val shortest_qualid_of_syndef : kernel_name -> qualid -val id_of_global : - Sign.named_context option -> global_reference -> identifier val dir_of_mp : module_path -> dir_path +val sp_of_global : global_reference -> section_path +val id_of_global : global_reference -> identifier + (* Printing of global references using names as short as possible *) -val pr_global_env : Sign.named_context option -> global_reference -> std_ppcmds +val pr_global_env : Idset.t -> global_reference -> std_ppcmds (* The [shortest_qualid] functions given an object with user_name diff --git a/parsing/argextend.ml4 b/parsing/argextend.ml4 index 576c57a53..4d6edda0e 100644 --- a/parsing/argextend.ml4 +++ b/parsing/argextend.ml4 @@ -78,7 +78,7 @@ let make_rule loc (prods,act) = let (symbs,pil) = List.split prods in <:expr< ($mlexpr_of_list (fun x -> x) symbs$,$make_act loc act pil$) >> -let declare_tactic_argument loc s typ pr f rawtyppr cl = +let declare_tactic_argument for_v8 loc s typ pr f rawtyppr cl = let interp = match f with | None -> <:expr< Tacinterp.genarg_interp >> | Some f -> <:expr< $lid:f$>> in @@ -103,12 +103,13 @@ let declare_tactic_argument loc s typ pr f rawtyppr cl = Pcoq.Gram.extend ($lid:s$ : Pcoq.Gram.Entry.e 'a) None [(None, None, $rules$)]; Pptactic.declare_extra_genarg_pprule + $mlexpr_of_bool for_v8$ ($rawwit$, $lid:rawpr$) ($wit$, $lid:pr$); end >> -let declare_vernac_argument loc s cl = +let declare_vernac_argument for_v8 loc s cl = let se = mlexpr_of_string s in let typ = ExtraArgType s in let wit = <:expr< $lid:"wit_"^s$ >> in @@ -173,13 +174,31 @@ EXTEND "END" -> if String.capitalize s = s then failwith "Argument entry names must be lowercase"; - declare_tactic_argument loc s typ pr f rawtyppr l + declare_tactic_argument true loc s typ pr f rawtyppr l | "VERNAC"; "ARGUMENT"; "EXTEND"; s = [ UIDENT | LIDENT ]; OPT "|"; l = LIST1 argrule SEP "|"; "END" -> if String.capitalize s = s then failwith "Argument entry names must be lowercase"; - declare_vernac_argument loc s l ] ] + declare_vernac_argument true loc s l + | "V7"; "ARGUMENT"; "EXTEND"; s = [ UIDENT | LIDENT ]; + "TYPED"; "AS"; typ = argtype; + "PRINTED"; "BY"; pr = LIDENT; + f = OPT [ "INTERPRETED"; "BY"; f = LIDENT -> f ]; + rawtyppr = + OPT [ "RAW"; "TYPED"; "AS"; t = argtype; + "RAW"; "PRINTED"; "BY"; pr = LIDENT -> (t,pr) ]; + OPT "|"; l = LIST1 argrule SEP "|"; + "END" -> + if String.capitalize s = s then + failwith "Argument entry names must be lowercase"; + declare_tactic_argument false loc s typ pr f rawtyppr l + | "V7"; "VERNAC"; "ARGUMENT"; "EXTEND"; s = [ UIDENT | LIDENT ]; + OPT "|"; l = LIST1 argrule SEP "|"; + "END" -> + if String.capitalize s = s then + failwith "Argument entry names must be lowercase"; + declare_vernac_argument false loc s l ] ] ; argtype: [ [ e = LIDENT -> fst (interp_entry_name loc e) ] ] diff --git a/parsing/ast.ml b/parsing/ast.ml index 13989bcbb..1b68f69bf 100755 --- a/parsing/ast.ml +++ b/parsing/ast.ml @@ -122,21 +122,35 @@ type grammar_action = type env = (string * typed_ast) list +let string_of_dirpath = function + | [] -> "" + | sl -> + String.concat "." (List.map string_of_id (List.rev sl)) + +let pr_id id = str (string_of_id id) + +let print_kn kn = + let (mp,dp,l) = repr_kn kn in + let dpl = repr_dirpath dp in + str (string_of_mp mp) ++ str "." ++ + prlist_with_sep (fun _ -> str".") pr_id dpl ++ + str (string_of_label l) + (* Pretty-printing *) let rec print_ast ast = match ast with | Num(_,n) -> int n | Str(_,s) -> qs s - | Path(_,sl) -> str (string_of_kn sl) + | Path(_,sl) -> print_kn sl | Id (_,s) -> str "{" ++ str s ++ str "}" - | Nvar(_,s) -> str (string_of_id s) + | Nvar(_,s) -> pr_id s | Nmeta(_,s) -> str s | Node(_,op,l) -> hov 3 (str "(" ++ str op ++ spc () ++ print_astl l ++ str ")") | Slam(_,None,ast) -> hov 1 (str "[<>]" ++ print_ast ast) | Slam(_,Some x,ast) -> hov 1 - (str "[" ++ str (string_of_id x) ++ str "]" ++ cut () ++ + (str "[" ++ pr_id x ++ str "]" ++ cut () ++ print_ast ast) | Smetalam(_,id,ast) -> hov 1 (str id ++ print_ast ast) | Dynamic(_,d) -> diff --git a/parsing/egrammar.ml b/parsing/egrammar.ml index 192f84d6b..77ac447c9 100644 --- a/parsing/egrammar.ml +++ b/parsing/egrammar.ml @@ -42,11 +42,26 @@ let assoc_level = function | _ -> "" let constr_level = function - | 8,assoc -> assert (assoc <> Gramext.LeftA); "top" | n,assoc -> (string_of_int n)^(assoc_level assoc) -let default_levels = [8,Gramext.RightA; 1,Gramext.RightA; 0,Gramext.RightA] -let level_stack = ref [default_levels] +let default_levels_v7 = + [10,Gramext.RightA; + 9,Gramext.RightA; + 8,Gramext.RightA; + 1,Gramext.RightA; + 0,Gramext.RightA] + +let default_levels_v8 = + [200,Gramext.RightA; + 100,Gramext.RightA; + 80,Gramext.RightA; + 40,Gramext.LeftA; + 10,Gramext.LeftA; + 9,Gramext.RightA; + 1,Gramext.LeftA; + 0,Gramext.RightA] + +let level_stack = ref [default_levels_v7] (* At a same level, LeftA takes precedence over RightA and NoneA *) (* In case, several associativity exists for a level, we make two levels, *) @@ -65,15 +80,17 @@ let create_assoc = function | Some a -> a let find_position other assoc lev = + let default = + if !Options.v7 then (10,Gramext.RightA) else (200,Gramext.RightA) in let current = List.hd !level_stack in match lev with | None -> level_stack := current :: !level_stack; None, (if other then assoc else None), None | Some n -> - if n = 8 & assoc = Some Gramext.LeftA then + if !Options.v7 & n = 8 & assoc = Some Gramext.LeftA then error "Left associativity not allowed at level 8"; - let after = ref (8,Gramext.RightA) in + let after = ref default in let rec add_level q = function | (p,_ as pa)::l when p > n -> pa :: add_level pa l | (p,a as pa)::l as l' when p = n -> @@ -92,7 +109,7 @@ let find_position other assoc lev = try (* Create the entry *) let current = List.hd !level_stack in - level_stack := add_level (8,Gramext.RightA) current :: !level_stack; + level_stack := add_level default current :: !level_stack; let assoc = create_assoc assoc in Some (Gramext.After (constr_level !after)), Some assoc, Some (constr_level (n,assoc)) @@ -347,7 +364,7 @@ let extend_constr_notation (n,assoc,ntn,rule) = let extend_constr_delimiters (sc,rule,pat_rule) = let mkact loc env = CDelimiters (loc,sc,snd (List.hd env)) in - extend_constr Constr.constr (ETConstr(0,()),Some 0,Some Gramext.NonA,false) + extend_constr Constr.operconstr (ETConstr(0,()),Some 0,Some Gramext.NonA,false) (make_act mkact) rule; let mkact loc env = CPatDelimiters (loc,sc,snd (List.hd env)) in extend_constr Constr.pattern (ETPattern,None,None,false) @@ -361,13 +378,21 @@ let make_rule univ f g (s,pt) = let act = make_gen_act f ntl in (symbs, act) +let tac_exts = ref [] +let get_extend_tactic_grammars () = !tac_exts + let extend_tactic_grammar s gl = + tac_exts := (s,gl) :: !tac_exts; let univ = get_univ "tactic" in let make_act loc l = Tacexpr.TacExtend (loc,s,List.map snd l) in let rules = List.map (make_rule univ make_act make_prod_item) gl in Gram.extend Tactic.simple_tactic None [(None, None, List.rev rules)] +let vernac_exts = ref [] +let get_extend_vernac_grammars () = !vernac_exts + let extend_vernac_command_grammar s gl = + vernac_exts := (s,gl) :: !vernac_exts; let univ = get_univ "vernac" in let make_act loc l = Vernacexpr.VernacExtend (s,List.map snd l) in let rules = List.map (make_rule univ make_act make_prod_item) gl in @@ -418,6 +443,15 @@ let extend_grammar gram = | TacticGrammar l -> add_tactic_entries l); grammar_state := gram :: !grammar_state +let reset_extend_grammars_v8 () = + let te = List.rev !tac_exts in + let tv = List.rev !vernac_exts in + tac_exts := []; + vernac_exts := []; + List.iter (fun (s,gl) -> extend_tactic_grammar s gl) te; + List.iter (fun (s,gl) -> extend_vernac_command_grammar s gl) tv; + level_stack := [default_levels_v8] + (* Summary functions: the state of the lexer is included in that of the parser. Because the grammar affects the set of keywords when adding or removing @@ -448,12 +482,10 @@ let unfreeze (grams, lex) = grammar_state := common; Lexer.unfreeze lex; List.iter extend_grammar (List.rev redo) - + let init_grammar () = remove_grammars (number_of_entries !grammar_state); grammar_state := [] -let _ = Lexer.init () - let init () = init_grammar () diff --git a/parsing/egrammar.mli b/parsing/egrammar.mli index d24264abc..091f98424 100644 --- a/parsing/egrammar.mli +++ b/parsing/egrammar.mli @@ -42,5 +42,11 @@ val extend_tactic_grammar : val extend_vernac_command_grammar : string -> (string * grammar_tactic_production list) list -> unit +val get_extend_tactic_grammars : + unit -> (string * (string * grammar_tactic_production list) list) list +val get_extend_vernac_grammars : + unit -> (string * (string * grammar_tactic_production list) list) list +val reset_extend_grammars_v8 : unit -> unit + val subst_all_grammar_command : Names.substitution -> all_grammar_command -> all_grammar_command diff --git a/parsing/extend.ml b/parsing/extend.ml index 14a33a656..076f7f026 100644 --- a/parsing/extend.ml +++ b/parsing/extend.ml @@ -175,7 +175,8 @@ let explicitize_prod_entry pos univ nt = | "constr5" -> ETConstr (5,pos) | "constr6" -> ETConstr (6,pos) | "constr7" -> ETConstr (7,pos) - | "constr8" | "constr" -> ETConstr (8,pos) + | "constr8" -> ETConstr (8,pos) + | "constr" when !Options.v7 -> ETConstr (8,pos) | "constr9" -> ETConstr (9,pos) | "constr10" | "lconstr" -> ETConstr (10,pos) | "pattern" -> ETPattern diff --git a/parsing/g_basevernac.ml4 b/parsing/g_basevernac.ml4 index 683b2ae54..a1208f48c 100644 --- a/parsing/g_basevernac.ml4 +++ b/parsing/g_basevernac.ml4 @@ -18,6 +18,15 @@ open Goptions open Constr open Prim +let vernac_kw = + [ "Quit"; "Load"; "Compile"; "Fixpoint"; "CoFixpoint"; + "Definition"; "Inductive"; "CoInductive"; + "Theorem"; "Variable"; "Axiom"; "Parameter"; "Hypothesis"; + "."; ">->" ] +let _ = List.iter (fun s -> Lexer.add_token ("",s)) vernac_kw + +let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" + GEXTEND Gram GLOBAL: class_rawexpr; @@ -155,7 +164,7 @@ GEXTEND Gram [ [ IDENT "Term"; qid = global -> PrintName qid | IDENT "All" -> PrintFullContext | IDENT "Section"; s = global -> PrintSectionContext s - | "Grammar"; uni = IDENT; ent = IDENT -> + | IDENT "Grammar"; uni = IDENT; ent = IDENT -> (* This should be in "syntax" section but is here for factorization*) PrintGrammar (uni, ent) | IDENT "LoadPath" -> PrintLoadPath @@ -213,19 +222,19 @@ GEXTEND Gram [ [ IDENT "Token"; s = STRING -> Pp.warning "Token declarations are now useless"; VernacNop - | "Grammar"; IDENT "tactic"; IDENT "simple_tactic"; + | IDENT "Grammar"; IDENT "tactic"; IDENT "simple_tactic"; OPT [ ":"; IDENT "tactic" ]; ":="; OPT "|"; tl = LIST0 grammar_tactic_rule SEP "|" -> VernacTacticGrammar tl - | "Grammar"; u = univ; + | IDENT "Grammar"; u = univ; tl = LIST1 grammar_entry SEP "with" -> VernacGrammar (rename_command_entry u,tl) - | "Syntax"; u = univ; el = LIST1 syntax_entry SEP ";" -> + | IDENT "Syntax"; u = univ; el = LIST1 syntax_entry SEP ";" -> VernacSyntax (u,el) - | "Uninterpreted"; IDENT "Notation"; s = STRING; + | IDENT "Uninterpreted"; IDENT "Notation"; s = STRING; l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] -> VernacSyntaxExtension (s,l) @@ -237,19 +246,46 @@ GEXTEND Gram | IDENT "Arguments"; IDENT "Scope"; qid = global; "["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl) - | IDENT "Infix"; a = entry_prec; n = OPT natural; op = STRING; + | IDENT "Infix"; a = entry_prec; n = OPT natural; + op = STRING; p = global; modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ] -> - let (a,n,b) = Metasyntax.interp_infix_modifiers a n modl in - VernacInfix (a,n,op,p,b,sc) + sc = OPT [ ":"; sc = IDENT -> sc]; + v8 = OPT[IDENT "V8only"; + op8=OPT STRING; + n8=OPT natural; + mv8=OPT["("; l = LIST1 syntax_modifier SEP ","; ")" -> l] -> + (op8,n8,mv8) ] -> + let (ai,ni,b) = Metasyntax.interp_infix_modifiers a n modl in + let v8 = + let (op8,nv8,mv8) = + match v8 with + Some (op8,nv8,mv8) -> + let op8 = match op8 with Some s -> s | _ -> op in + let nv8 = match nv8 with Some n -> Some n + | _ -> Util.option_app (( * ) 10) n in + let mv8 = match mv8 with Some m -> m | _ -> modl in + (op8,nv8,mv8) + | None -> (op,Util.option_app(( * ) 10) n, modl) in + let (a8,n8,_) = + Metasyntax.interp_infix_modifiers a nv8 mv8 in + Some(a8,n8,op8) in + VernacInfix (ai,ni,op,p,b,v8,sc) | IDENT "Distfix"; a = entry_prec; n = natural; s = STRING; p = global; sc = OPT [ ":"; sc = IDENT -> sc ] -> VernacDistfix (a,n,s,p,sc) | IDENT "Notation"; s = IDENT; ":="; c = constr -> - VernacNotation ("'"^s^"'",c,[],None) + VernacNotation ("'"^s^"'",c,[],None,None) | IDENT "Notation"; s = STRING; ":="; c = constr; modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; - sc = OPT [ ":"; sc = IDENT -> sc ] -> VernacNotation (s,c,modl,sc) + sc = OPT [ ":"; sc = IDENT -> sc ]; + v8 = OPT[IDENT "V8only"; + s8=OPT STRING; + mv8=["(";mv8=LIST1 syntax_modifier SEP ","; ")" -> mv8 + | ->[] ] -> (s8,mv8) ] -> + let v8 = Util.option_app (fun (s8,mv8) -> + let s8 = match s8 with Some s -> s | _ -> s in + (s8,mv8)) v8 in + VernacNotation (s,c,modl,v8,sc) (* "Print" "Grammar" should be here but is in "command" entry in order to factorize with other "Print"-based vernac entries *) @@ -269,9 +305,8 @@ GEXTEND Gram ; syntax_extension_type: [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference - | IDENT "annot" -> - if !Options.v7 <> true then Util.error "annot not allowed in new syntax"; - ETOther ("constr","annot") + | IDENT "bigint" -> ETBigint + | i=IDENT -> ETOther ("constr",i) ] ] ; opt_scope: diff --git a/parsing/g_cases.ml4 b/parsing/g_cases.ml4 index 62ee20cad..679d174c2 100644 --- a/parsing/g_cases.ml4 +++ b/parsing/g_cases.ml4 @@ -20,7 +20,7 @@ let pair loc = Qualid (loc, Libnames.qualid_of_string "Coq.Init.Datatypes.pair") GEXTEND Gram - GLOBAL: constr pattern; + GLOBAL: operconstr pattern; pattern: [ [ r = Prim.reference -> CPatAtom (loc,Some r) @@ -47,12 +47,12 @@ GEXTEND Gram | p = pattern -> p ] ] ; equation: - [ [ lhs = LIST1 pattern; "=>"; rhs = constr9 -> (loc,lhs,rhs) ] ] + [ [ lhs = LIST1 pattern; "=>"; rhs = operconstr LEVEL "9" -> (loc,lhs,rhs) ] ] ; ne_eqn_list: [ [ leqn = LIST1 equation SEP "|" -> leqn ] ] ; - constr: LEVEL "1" + operconstr: LEVEL "1" [ [ "<"; p = annot; ">"; "Cases"; lc = LIST1 constr; "of"; OPT "|"; eqs = ne_eqn_list; "end" -> CCases (loc, Some p, lc, eqs) diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index eac6ce12f..cea564953 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -17,6 +17,19 @@ open Libnames open Prim open Topconstr +(* Initialize the lexer *) +let constr_kw = + [ "Cases"; "of"; "with"; "end"; "as"; "in"; "Prop"; "Set"; "Type"; + ":"; "("; ")"; "["; "]"; "{"; "}"; ","; ";"; "->"; "="; ":="; "!"; + "::"; "<:"; ":<"; "=>"; "<"; ">"; "|"; "?"; "/"; + "<->"; "\\/"; "/\\"; "`"; "``"; "&"; "*"; "+"; "@"; "^"; "#"; "-"; + "~"; "'"; "<<"; ">>"; "<>"; + ] +let _ = 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 ? *) + + let coerce_to_var = function | CRef (Ident (_,id)) -> id | ast -> Util.user_err_loc @@ -78,8 +91,9 @@ let test_ident_colon = end | _ -> raise Stream.Failure) + GEXTEND Gram - GLOBAL: constr9 lconstr constr sort global constr_pattern Constr.ident annot + GLOBAL: operconstr lconstr constr sort global constr_pattern Constr.ident annot (*ne_name_comma_list*); Constr.ident: [ [ id = Prim.ident -> id @@ -105,8 +119,24 @@ GEXTEND Gram | "Type" -> RType None ] ] ; constr: - [ "top" RIGHTA - [ c1 = constr; "->"; c2 = constr -> CArrow (loc, c1, c2) ] + [ [ c = operconstr LEVEL "8" -> c ] ] + ; + lconstr: + [ [ c = operconstr LEVEL "10" -> c ] ] + ; + operconstr: + [ "10" RIGHTA + [ "!"; f = global; args = LIST0 (operconstr LEVEL "9") -> + CAppExpl (loc, f, args) +(* + | "!"; f = global; "with"; b = binding_list -> + <:ast< (APPLISTWITH $f $b) >> +*) + | f = operconstr; args = LIST1 constr91 -> CApp (loc, f, args) ] + | "9" RIGHTA + [ c1 = operconstr; "::"; c2 = operconstr -> CCast (loc, c1, c2) ] + | "8" RIGHTA + [ c1 = operconstr; "->"; c2 = operconstr -> CArrow (loc, c1, c2) ] | "1" RIGHTA [ "<"; p = annot; ">"; IDENT "Match"; c = constr; "with"; cl = LIST0 constr; "end" -> @@ -119,32 +149,32 @@ GEXTEND Gram | 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 LEVEL "top"-> + 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 LEVEL "top" -> + "in"; c1 = constr -> CLetIn (loc, na, c, c1) | IDENT "if"; c1 = constr; IDENT "then"; c2 = constr; - IDENT "else"; c3 = constr LEVEL "top" -> + 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 LEVEL "top" -> + "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 LEVEL "top" -> + IDENT "else"; c3 = constr -> COrderedCase (loc, IfStyle, Some p, c1, [c2; c3]) ] | "0" RIGHTA [ "?" -> CHole loc | "?"; n = Prim.natural -> CMeta (loc, n) - | bll = binders; c = constr LEVEL "top" -> abstract_constr loc c bll + | bll = binders; c = constr -> abstract_constr loc c bll (* Hack to parse syntax "(n)" as a natural number *) | "("; test_int_rparen; n = bigint; ")" -> CDelimiters (loc,"N",CNumeral (loc,n)) @@ -182,24 +212,10 @@ GEXTEND Gram | "'"; test_ident_colon; key = IDENT; ":"; c = constr; "'" -> CDelimiters (loc,key,c) ] ] ; - lconstr: - [ "10" RIGHTA - [ "!"; f = global; args = LIST0 constr9 -> CAppExpl (loc, f, args) -(* - | "!"; f = global; "with"; b = binding_list -> - <:ast< (APPLISTWITH $f $b) >> -*) - | f = constr9; args = LIST1 constr91 -> CApp (loc, f, args) - | f = constr9 -> f ] ] - ; constr91: - [ [ test_int_bang; n = INT; "!"; c = constr9 -> + [ [ test_int_bang; n = INT; "!"; c = operconstr LEVEL "9" -> (c, Some (int_of_string n)) - | c = constr9 -> (c, None) ] ] - ; - constr9: - [ RIGHTA [ c1 = constr -> c1 - | c1 = constr; "::"; c2 = constr -> CCast (loc, c1, c2) ] ] + | c = operconstr LEVEL "9" -> (c, None) ] ] ; (* annot and product_annot_tail are hacks to forbid concrete syntax *) (* ">" (e.g. for gt, Zgt, ...) in annotations *) @@ -232,7 +248,7 @@ GEXTEND Gram [ c1 = SELF; "=="; c2 = NEXT -> CNotation (loc, "_ == _", [c1;c2]) ] | RIGHTA [ c1 = SELF; "="; c2 = NEXT -> CNotation (loc, "_ = _", [c1;c2]) ] - | [ c = constr LEVEL "4L" -> c ] ] + | [ c = operconstr LEVEL "4L" -> c ] ] ; product_annot_tail: [ [ ";"; idl = ne_name_comma_list; ":"; c = constr; diff --git a/parsing/g_constrnew.ml4 b/parsing/g_constrnew.ml4 new file mode 100644 index 000000000..ba8be907c --- /dev/null +++ b/parsing/g_constrnew.ml4 @@ -0,0 +1,256 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Lexer.add_token("",s)) constr_kw +let _ = Options.v7 := false + +let pair loc = + Qualid (loc, Libnames.qualid_of_string "Coq.Init.Datatypes.pair") + +let mk_cast = function + (c,(_,None)) -> c + | (c,(_,Some ty)) -> CCast(join_loc (constr_loc c) (constr_loc ty), c, ty) + +let mk_lam = function + ([],c) -> c + | (bl,c) -> CLambdaN(constr_loc c, bl,c) + +let coerce_to_id c = + match c with + CRef(Ident(loc,id)) -> (loc,Name id) + | CHole loc -> (loc,Anonymous) + | _ -> error "ill-formed match type annotation" + +let match_bind_of_pat loc (oid,ty) = + let l2 = + match oid with + None -> [] + | Some x -> [([x],CHole loc)] in + let l1 = + match ty with + None -> [] (* No: should lookup inductive arity! *) + | Some (CApp(_,_,args)) -> (* parameters do not appear *) + List.map (fun (c,_) -> ([coerce_to_id c],CHole loc)) args + | _ -> error "ill-formed match type annotation" in + l1@l2 + +let mk_match (loc,cil,rty,br) = + let (lc,pargs) = List.split cil in + let pr = + match rty with + None -> None + | Some ty -> + let idargs = (* TODO: not forget the list lengths for PP! *) + List.flatten (List.map (match_bind_of_pat loc) pargs) in + Some (CLambdaN(loc,idargs,ty)) in + CCases(loc,pr,lc,br) + +let mk_fixb (loc,id,bl,ann,body,(tloc,tyc)) = + let n = + match bl,ann with + [([_],_)], None -> 0 + | _, Some x -> + let ids = List.map snd (List.flatten (List.map fst bl)) in + (try list_index (snd x) ids - 1 + with Not_found -> error "no such fix variable") + | _ -> error "cannot find decreasing arg of fix" in + let ty = match tyc with + None -> CHole tloc + | Some t -> CProdN(loc,bl,t) in + (snd id,n,ty,CLambdaN(loc,bl,body)) + +let mk_cofixb (loc,id,bl,ann,body,(tloc,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 + None -> CHole tloc + | Some t -> CProdN(loc,bl,t) in + (snd id,ty,CLambdaN(loc,bl,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 binder_constr = + create_constr_entry (get_univ "constr") "binder_constr" + +GEXTEND Gram + GLOBAL: binder_constr lconstr constr operconstr sort global + constr_pattern Constr.ident; + Constr.ident: + [ [ id = Prim.ident -> id + + (* This is used in quotations and Syntax *) + | id = METAIDENT -> id_of_string id ] ] + ; + global: + [ [ r = Prim.reference -> r + + (* This is used in quotations *) + | id = METAIDENT -> Ident (loc,id_of_string id) ] ] + ; + constr_pattern: + [ [ c = constr -> c ] ] + ; + sort: + [ [ "Set" -> RProp Pos + | "Prop" -> RProp Null + | "Type" -> RType None ] ] + ; + lconstr: + [ [ c = operconstr -> c ] ] + ; + constr: + [ [ c = operconstr LEVEL "9" -> c ] ] + ; + operconstr: + [ "200" RIGHTA + [ c = binder_constr -> c ] + | "100" RIGHTA + [ c1 = operconstr; ":"; c2 = binder_constr -> CCast(loc,c1,c2) + | c1 = operconstr; ":"; c2 = operconstr -> CCast(loc,c1,c2) ] + | "80" RIGHTA + [ c1 = operconstr; "->"; c2 = binder_constr -> CArrow(loc,c1,c2) + | c1 = operconstr; "->"; c2 = operconstr -> CArrow(loc,c1,c2) ] + | "40L" LEFTA + [ "-"; c = operconstr -> CNotation(loc,"- _",[c]) ] + | "10L" LEFTA + [ f=operconstr; args=LIST1 appl_arg -> CApp(loc,f,args) + | "@"; f=global; args=LIST0 NEXT -> CAppExpl(loc,f,args) ] + | "9" [ ] + | "1L" LEFTA + [ c=operconstr; "%"; key=IDENT -> CDelimiters (loc,key,c) ] + | "0" + [ c=atomic_constr -> c + | c=match_constr -> c + | "("; c=operconstr; ")" -> c ] ] + ; + binder_constr: + [ [ "!"; bl = LIST1 binder; "."; c = operconstr LEVEL "200" -> + CProdN(loc,bl,c) + | "fun"; bl = LIST1 binder; ty = type_cstr; "=>"; + c = operconstr LEVEL "200" -> + CLambdaN(loc,bl,mk_cast(c,ty)) + | "let"; id=name; bl = LIST0 binder; ty = type_cstr; ":="; + c1 = operconstr; "in"; c2 = operconstr LEVEL "200" -> + CLetIn(loc,id,mk_lam(bl,mk_cast(c1,ty)),c2) + | "let"; "("; lb = LIST1 name SEP ","; ")"; ":="; + c1 = operconstr; "in"; c2 = operconstr LEVEL "200" -> + COrderedCase (loc, LetStyle, None, c1, + [CLambdaN (loc, [lb, CHole loc], c2)]) + | "if"; c1=operconstr; "then"; c2=operconstr LEVEL "200"; + "else"; c3=operconstr LEVEL "200" -> + COrderedCase (loc, IfStyle, None, c1, [c2; c3]) + | c=fix_constr -> c ] ] + ; + appl_arg: + [ [ "@"; n=INT; ":="; c=operconstr LEVEL "9" -> (c,Some(int_of_string n)) + | c=operconstr LEVEL "9" -> (c,None) ] ] + ; + atomic_constr: + [ [ g=global -> CRef g + | s=sort -> CSort(loc,s) + | n=INT -> CNumeral (loc,Bignat.POS (Bignat.of_string n)) + | IDENT "_" -> CHole loc + | "?"; n=Prim.natural -> CMeta(loc,n) ] ] + ; + fix_constr: + [ [ kw=fix_kw; dcl=fix_decl -> + let (_,n,_,_,_,_) = dcl in mk_fix(loc,kw,n,[dcl]) + | kw=fix_kw; dcl1=fix_decl; "with"; dcls=LIST1 fix_decl SEP "with"; + "for"; id=identref -> + mk_fix(loc,kw,id,dcl1::dcls) + ] ] + ; + fix_kw: + [ [ "fix" -> true + | "cofix" -> false ] ] + ; + fix_decl: + [ [ id=identref; bl=LIST0 binder; ann=fixannot; ty=type_cstr; ":="; + c=operconstr LEVEL "200" -> (loc,id,bl,ann,c,ty) ] ] + ; + fixannot: + [ [ "{"; "struct"; id=name; "}" -> Some id + | -> None ] ] + ; + match_constr: + [ [ "match"; ci=LIST1 case_item SEP ","; ty=case_type; "with"; + br=branches; "end" -> mk_match (loc,ci,ty,br) ] ] + ; + case_item: + [ [ c=operconstr LEVEL "80"; p=pred_pattern -> (c,p) ] ] + ; + pred_pattern: + [ [ oid = OPT ["as"; id=name -> id]; + (_,ty) = type_cstr -> (oid,ty) ] ] + ; + case_type: + [ [ ty = OPT [ "=>"; c = lconstr -> c ] -> ty ] ] + ; + branches: + [ [ OPT"|"; br=LIST0 eqn SEP "|" -> br ] ] + ; + eqn: + [ [ pl = LIST1 pattern SEP ","; "=>"; rhs = lconstr -> (loc,pl,rhs) ] ] + ; + simple_pattern: + [ [ r = Prim.reference -> CPatAtom (loc,Some r) + | IDENT "_" -> CPatAtom (loc,None) + | "("; p = lpattern; ")" -> p + | n = bigint -> CPatNumeral (loc,n) + ] ] + ; + pattern: + [ [ p = pattern ; lp = LIST1 simple_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) + | c = pattern; "%"; key=IDENT -> + CPatDelimiters (loc,key,c) + | p = simple_pattern -> p ] ] + ; + lpattern: + [ [ c = pattern -> c + | p1=pattern; ","; p2=lpattern -> CPatCstr (loc, pair loc, [p1;p2]) ] ] + ; + binder: + [ [ id=name -> ([id],CHole loc) + | "("; id=name; ")" -> ([id],CHole loc) + | "("; id=name; ":"; c=lconstr; ")" -> ([id],c) + | id=name; ":"; c=constr -> ([id],c) ] ] (* tolerance *) + ; + type_cstr: + [ [ c=OPT [":"; c=lconstr -> c] -> (loc,c) ] ] + ; +END;; diff --git a/parsing/g_ltac.ml4 b/parsing/g_ltac.ml4 index 7d82fa715..1cc0c0297 100644 --- a/parsing/g_ltac.ml4 +++ b/parsing/g_ltac.ml4 @@ -60,6 +60,10 @@ end open Prelude +let arg_of_expr = function + TacArg a -> a + | e -> Tacexp e + (* Tactics grammar rules *) GEXTEND Gram @@ -220,8 +224,8 @@ GEXTEND Gram | ta = tactic_arg0 -> ta ] ] ; tactic_arg0: - [ [ "("; a = tactic_expr; ")" -> Tacexp a - | "()" -> TacVoid + [ [ "("; a = tactic_expr; ")" -> arg_of_expr a + | "()" -> Integer 0 | r = reference -> Reference r | n = integer -> Integer n | id = METAIDENT -> MetaIdArg (loc,id) diff --git a/parsing/g_ltacnew.ml4 b/parsing/g_ltacnew.ml4 new file mode 100644 index 000000000..fe7c17ad1 --- /dev/null +++ b/parsing/g_ltacnew.ml4 @@ -0,0 +1,209 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* user_err_loc (loc, "", (str "Syntax Error")) + | Qast.Node ("LETCLAUSE", [id;c;d]) -> + Qast.Tuple [id;c;d] + | _ -> anomaly "out_letin_clause" + +let make_letin_clause _ = function + | Qast.List l -> Qast.List (List.map (out_letin_clause dummy_loc) l) + | _ -> anomaly "make_letin_clause" +end +else +module Prelude = struct +let fail_default_value = 0 + +let out_letin_clause loc = function + | LETTOPCLAUSE _ -> user_err_loc (loc, "", (str "Syntax Error")) + | LETCLAUSE (id,c,d) -> (id,c,d) + +let make_letin_clause loc = List.map (out_letin_clause loc) +end + +let arg_of_expr = function + TacArg a -> a + | e -> Tacexp e + +open Prelude + +(* Tactics grammar rules *) + +let tactic_expr = Gram.Entry.create "tactic:tactic_expr" + +GEXTEND Gram + GLOBAL: tactic Vernac_.command tactic_expr tactic_arg; + + tactic_expr: + [ "5" LEFTA + [ ta0 = tactic_expr; "&"; ta1 = tactic_expr -> TacThen (ta0, ta1) + | ta = tactic_expr; "&"; "["; lta = LIST0 tactic_expr SEP "|"; "]" -> + TacThens (ta, lta) ] + | "4" + [ ] + | "3" RIGHTA + [ IDENT "try"; ta = tactic_expr -> TacTry ta + | IDENT "do"; n = natural; 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 ] + | "2" RIGHTA + [ ta0 = tactic_expr; "||"; ta1 = tactic_expr -> TacOrelse (ta0,ta1) ] + | "1" RIGHTA + [ "fun"; it = LIST1 input_fun ; "=>"; body = tactic_expr -> + TacFun (it,body) + | 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) + | "match"; IDENT "context"; "with"; mrl = match_context_list; "end" -> + TacMatchContext (false,mrl) + | "match"; IDENT "reverse"; IDENT "context"; "with"; + mrl = match_context_list; "end" -> + TacMatchContext (true,mrl) + | "match"; c = constrarg; "with"; mrl = match_list; "end" -> + TacMatch (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" -> TacId + | IDENT "fail" -> TacFail fail_default_value + | IDENT "fail"; n = natural -> TacFail n + | st = simple_tactic -> TacAtom (loc,st) + | IDENT "eval"; rtc = red_expr; "in"; c = Constr.lconstr -> + TacArg(ConstrMayEval (ConstrEval (rtc,c))) + | IDENT "inst"; id = identref; "["; c = Constr.lconstr; "]" -> + TacArg(ConstrMayEval (ConstrContext (id,c))) + | IDENT "check"; c = Constr.lconstr -> + TacArg(ConstrMayEval (ConstrTypeOf c)) + | "'"; c = Constr.constr -> TacArg(ConstrMayEval(ConstrTerm c)) + | 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: + [ [ "'"; a = tactic_expr LEVEL "0" -> arg_of_expr a + | a = tactic_atom -> a + | c = Constr.constr -> ConstrMayEval (ConstrTerm c) ] ] + ; + tactic_atom: + [ [ id = METAIDENT -> MetaIdArg (loc,id) + | r = reference -> Reference r + | "?"; n = natural -> MetaNumArg (loc,n) + | "()" -> TacVoid ] ] + ; + input_fun: + [ [ IDENT "_" -> None + | l = base_ident -> Some l ] ] + ; + let_clause: + [ [ id = identref; ":="; te = tactic_expr -> + LETCLAUSE (id, None, arg_of_expr te) + | (_,id) = identref; ":"; c = Constr.constr; ":="; IDENT "proof" -> + LETTOPCLAUSE (id, c) + | id = identref; ":"; c = constrarg; ":="; te = tactic_expr -> + LETCLAUSE (id, Some c, arg_of_expr te) + | (_,id) = identref; ":"; c = Constr.lconstr -> + LETTOPCLAUSE (id, c) ] ] + ; + rec_clause: + [ [ 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 ] ] + ; + match_hyps: + [ [ id = identref; ":"; mp = match_pattern -> Hyp (id, mp) + | IDENT "_"; ":"; mp = match_pattern -> NoHypId mp ] ] + ; + match_context_rule: + [ [ "["; largs = LIST0 match_hyps SEP ","; "|-"; mp = match_pattern; "]"; + "=>"; te = tactic_expr -> Pat (largs, mp, te) + | IDENT "_"; "=>"; te = tactic_expr -> All te ] ] + ; + match_context_list: + [ [ mrl = LIST1 match_context_rule SEP "|" -> mrl + | "|"; mrl = LIST1 match_context_rule SEP "|" -> mrl ] ] + ; + match_rule: + [ [ "["; mp = match_pattern; "]"; "=>"; te = tactic_expr -> Pat ([],mp,te) + | IDENT "_"; "=>"; 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: + [ [ deftok; IDENT "Definition"; b = tacdef_body -> + VernacDeclareTacticDefinition (false, [b]) + | IDENT "Recursive"; deftok; IDENT "Definition"; + l = LIST1 tacdef_body SEP "with" -> + VernacDeclareTacticDefinition (true, l) ] ] + ; + END diff --git a/parsing/g_natsyntaxnew.ml b/parsing/g_natsyntaxnew.ml new file mode 100644 index 000000000..b124c4ec0 --- /dev/null +++ b/parsing/g_natsyntaxnew.ml @@ -0,0 +1,195 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* > *) +let nat_of_int n dloc = + let ast_O = set_loc dloc ast_O in + let ast_S = set_loc dloc ast_S in + let rec mk_nat n = + if n <= 0 then + ast_O + else + Node(dloc,"APPLIST", [ast_S; mk_nat (n-1)]) + in + mk_nat n + +let pat_nat_of_int n dloc = + let ast_O = set_loc dloc ast_O in + let ast_S = set_loc dloc ast_S in + let rec mk_nat n = + if n <= 0 then + ast_O + else + Node(dloc,"PATTCONSTRUCT", [ast_S; mk_nat (n-1)]) + in + mk_nat n + +let nat_of_string s dloc = + nat_of_int (int_of_string s) dloc + +let pat_nat_of_string s dloc = + pat_nat_of_int (int_of_string s) dloc + +exception Non_closed_number + +let rec int_of_nat_rec astS astO p = + match p with + | Node (_,"APPLIST", [b; a]) when alpha_eq(b,astS) -> + (int_of_nat_rec astS astO a)+1 + | a when alpha_eq(a,astO) -> 1 + (***** YES, 1, non 0 ... to print the successor of p *) + | _ -> raise Non_closed_number + +let int_of_nat p = + try + Some (int_of_nat_rec ast_S ast_O p) + with + Non_closed_number -> None + +let pr_S a = hov 0 (str "S" ++ brk (1,1) ++ a) + +let rec pr_external_S std_pr = function + | Node (l,"APPLIST", [b; a]) when alpha_eq (b,ast_S) -> + str"(" ++ pr_S (pr_external_S std_pr a) ++ str")" + | p -> std_pr p + +(* Declare the primitive printer *) + +(* Prints not p, but the SUCCESSOR of p !!!!! *) +let nat_printer std_pr p = + match (int_of_nat p) with + | Some i -> str "(" ++ str (string_of_int i) ++ str ")" + | None -> str "(" ++ pr_S (pr_external_S std_pr p) ++ str ")" + +let _ = Esyntax.Ppprim.add ("nat_printer", nat_printer) +*) +(* +(* Declare the primitive parser *) + +let unat = create_univ_if_new "nat" + +let number = create_constr_entry unat "number" +let pat_number = create_constr_entry unat "pat_number" + +let _ = + Gram.extend number None + [None, None, + [[Gramext.Stoken ("INT", "")], + Gramext.action nat_of_string]] + +let _ = + Gram.extend pat_number None + [None, None, + [[Gramext.Stoken ("INT", "")], + Gramext.action pat_nat_of_string]] +*) + +(*i*) +open Rawterm +open Libnames +open Bignat +open Coqlib +open Symbols +open Pp +open Util +open Names +(*i*) + +(**********************************************************************) +(* Parsing via scopes *) +(* For example, (nat_of_string "3") is <<(S (S (S O)))>> *) + +let nat_of_int dloc n = + match n with + | POS n -> + if less_than (of_string "5000") n & Options.is_verbose () then begin + warning ("You may experiment stack overflow and segmentation fault\ + \nwhile parsing numbers in nat greater than 5000"); + flush_all () + end; + let ref_O = RRef (dloc, glob_O) in + let ref_S = RRef (dloc, glob_S) in + let rec mk_nat acc n = + if is_nonzero n then + mk_nat (RApp (dloc,ref_S, [acc])) (sub_1 n) + else + acc + in + mk_nat ref_O n + | NEG n -> + user_err_loc (dloc, "nat_of_int", + str "Cannot interpret a negative number as a natural number") + +let pat_nat_of_int dloc n name = + match n with + | POS n -> + let rec mk_nat n name = + if is_nonzero n then + PatCstr (dloc,path_of_S,[mk_nat (sub_1 n) Anonymous],name) + else + PatCstr (dloc,path_of_O,[],name) + in + mk_nat n name + | NEG n -> + user_err_loc (dloc, "pat_nat_of_int", + str "Cannot interpret a negative number as a natural number") + +(***********************************************************************) +(* Printing via scopes *) + +exception Non_closed_number + +let rec int_of_nat = function + | RApp (_,RRef (_,s),[a]) when s = glob_S -> add_1 (int_of_nat a) + | RRef (_,z) when z = glob_O -> zero + | _ -> raise Non_closed_number + +let uninterp_nat p = + try + Some (POS (int_of_nat p)) + with + Non_closed_number -> None + +let rec int_of_nat_pattern = function + | PatCstr (_,s,[a],_) when ConstructRef s = glob_S -> + add_1 (int_of_nat_pattern a) + | PatCstr (_,z,[],_) when ConstructRef z = glob_O -> zero + | _ -> raise Non_closed_number + +let uninterp_nat_pattern p = + try + Some (POS (int_of_nat_pattern p)) + with + Non_closed_number -> None + +(***********************************************************************) +(* Declare the primitive parsers and printers *) + +let _ = + Symbols.declare_numeral_interpreter "nat_scope" + ["Coq";"Init";"Peano"] + (nat_of_int,Some pat_nat_of_int) + ([RRef (dummy_loc,glob_S); RRef (dummy_loc,glob_O)], uninterp_nat, None) diff --git a/parsing/g_natsyntaxnew.mli b/parsing/g_natsyntaxnew.mli new file mode 100644 index 000000000..91596ef5c --- /dev/null +++ b/parsing/g_natsyntaxnew.mli @@ -0,0 +1,11 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* "; "<<"; ">>"; "'";"."] +let _ = List.iter (fun s -> Lexer.add_token("",s)) prim_kw + +ifdef Quotify then +open Qast + +open Prim + +ifdef Quotify then +module Prelude = struct +let local_id_of_string s = Apply ("Names","id_of_string", [s]) +let local_make_dirpath l = Qast.Apply ("Names","make_dirpath",[l]) +let local_make_posint s = s +let local_make_negint s = Apply ("Pervasives","~-", [s]) +let local_make_qualid l id' = + Qast.Apply ("Nametab","make_qualid", [local_make_dirpath l;id']) +let local_make_short_qualid id = + Qast.Apply ("Nametab","make_short_qualid",[id]) +let local_make_path l id' = + Qast.Apply ("Libnames","encode_kn", [local_make_dirpath l;id']) +let local_make_binding loc a b = + match a with + | Qast.Node ("Nvar", [_;id]) -> + Qast.Node ("Slam", [Qast.Loc; Qast.Option (Some id); b]) + | Qast.Node ("Nmeta", [_;s]) -> + Qast.Node ("Smetalam", [Qast.Loc;s;b]) + | _ -> + Qast.Apply ("Pervasives", "failwith", [Qast.Str "Slam expects a var or a metavar"]) +let local_append l id = Qast.Apply ("List","append", [l; Qast.List [id]]) +end + +else + +module Prelude = struct +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] +end + +open Prelude + +ifdef Quotify then +open Q + +GEXTEND Gram + GLOBAL: ast bigint qualid reference; + metaident: + [ [ s = METAIDENT -> Nmeta (loc,s) ] ] + ; + field: + [ [ "."; s = IDENT -> local_id_of_string s ] ] + ; + fields: + [ [ id = field; (l,id') = fields -> (local_append l id,id') + | id = field -> ([],id) + ] ] + ; + astpath: + [ [ id = base_ident; (l,a) = fields -> + Path(loc, local_make_path (local_append l id) a) + | id = base_ident -> Nvar(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 + ] ] + ; + reference: + [ [ id = base_ident; (l,id') = fields -> + Qualid (loc, local_make_qualid (local_append l id) id') + | id = base_ident -> Ident (loc,id) + ] ] + ; + qualid: + [ [ qid = basequalid -> loc, qid ] ] + ; + 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 + | "'"; a = ast -> Node(loc,"$QUOTE",[a]) ] ] + ; + bigint: (* Negative numbers are dealt with specially *) + [ [ i = INT -> Bignat.POS (Bignat.of_string i) ] ] + ; +END + +let constr_to_ast a = + Termast.ast_of_rawconstr + (Constrintern.interp_rawconstr Evd.empty (Global.env()) a) + +let constr_parser_with_glob = Pcoq.map_entry constr_to_ast Constr.constr + +let _ = define_ast_quotation true "constr" constr_parser_with_glob diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 879ce98ae..f6801bf31 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -18,6 +18,8 @@ open Vernacexpr open Prim open Constr +let thm_token = Gram.Entry.create "vernac:thm_token" + (* Proof commands *) GEXTEND Gram GLOBAL: command; diff --git a/parsing/g_proofsnew.ml4 b/parsing/g_proofsnew.ml4 new file mode 100644 index 000000000..db8da7101 --- /dev/null +++ b/parsing/g_proofsnew.ml4 @@ -0,0 +1,126 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Tacexpr.ConclLocation () + | discard = [ IDENT "Discardable" -> true | -> false ]; IDENT "Hypothesis" + -> Tacexpr.HypLocation discard ] ] + ; + opt_hintbases: + [ [ -> [] + | ":"; l = LIST1 IDENT -> l ] ] + ; + command: + [ [ IDENT "Goal"; c = Constr.lconstr -> VernacGoal c + | IDENT "Proof" -> VernacNop + | 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 "Qed" -> VernacEndProof (true,None) + | IDENT "Save" -> VernacEndProof (true,None) + | IDENT "Save"; tok = thm_token; id = base_ident -> + VernacEndProof (true,Some (id,Some tok)) + | IDENT "Save"; id = base_ident -> VernacEndProof (true,Some (id,None)) + | IDENT "Defined" -> VernacEndProof (false,None) + | IDENT "Defined"; id=base_ident -> VernacEndProof (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 "Implicits"; n = natural -> + VernacShow (ShowGoalImplicitly (Some n)) + | IDENT "Show"; IDENT "Implicits" -> VernacShow (ShowGoalImplicitly None) + | IDENT "Show"; IDENT "Node" -> VernacShow ShowNode + | IDENT "Show"; IDENT "Script" -> VernacShow ShowScript + | IDENT "Show"; IDENT "Existentials" -> VernacShow ShowExistentials + | IDENT "Show"; IDENT "Tree" -> VernacShow ShowTree + | IDENT "Show"; IDENT "Conjectures" -> VernacShow ShowProofNames + | IDENT "Show"; IDENT "Proof" -> VernacShow ShowProof + | IDENT "Show"; IDENT "Intro" -> VernacShow (ShowIntros false) + | IDENT "Show"; IDENT "Intros" -> VernacShow (ShowIntros true) + | 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 "HintDestruct"; + dloc = destruct_location; + id = base_ident; + hyptyp = Constr.constr_pattern; + pri = natural; + "["; tac = tactic; "]" -> + VernacHintDestruct (id,dloc,hyptyp,pri,tac) + + | IDENT "Hint"; hintname = base_ident; dbnames = opt_hintbases; ":="; h = hint + -> VernacHints (dbnames, h hintname) + + | IDENT "Hints"; (dbnames,h) = hints -> VernacHints (dbnames, h) + + +(*This entry is not commented, only for debug*) + | IDENT "PrintConstr"; c = Constr.constr -> + VernacExtend ("PrintConstr", + [Genarg.in_gen Genarg.rawwit_constr c]) + ] ]; + + hint: + [ [ IDENT "Resolve"; c = Constr.constr -> fun name -> HintsResolve [Some name, c] + | IDENT "Immediate"; c = Constr.constr -> fun name -> HintsImmediate [Some name, c] + | IDENT "Unfold"; qid = global -> fun name -> HintsUnfold [Some name,qid] + | IDENT "Constructors"; c = global -> fun n -> HintsConstructors (n,c) + | IDENT "Extern"; n = natural; c = Constr.constr ; tac = tactic -> + fun name -> HintsExtern (name,n,c,tac) ] ] + ; + hints: + [ [ IDENT "Resolve"; l = LIST1 global; dbnames = opt_hintbases -> + (dbnames, HintsResolve (List.map (fun qid -> (None, CRef qid)) l)) + | IDENT "Immediate"; l = LIST1 global; dbnames = opt_hintbases -> + (dbnames, HintsImmediate (List.map (fun qid -> (None, CRef qid)) l)) + | IDENT "Unfold"; l = LIST1 global; dbnames = opt_hintbases -> + (dbnames, HintsUnfold (List.map (fun qid -> (None,qid)) l)) ] ] + ; + constr_body: + [ [ ":="; c = lconstr -> c + | ":"; t = lconstr; ":="; c = lconstr -> CCast(loc,c,t) ] ] + ; +END diff --git a/parsing/g_rsyntax.ml b/parsing/g_rsyntax.ml index 6edbf7c55..13eb76334 100644 --- a/parsing/g_rsyntax.ml +++ b/parsing/g_rsyntax.ml @@ -66,7 +66,13 @@ let r_of_int n dloc = let r_of_string s dloc = r_of_int (int_of_string s) dloc -let rnumber = create_constr_entry (get_univ "rnatural") "rnumber" +let rsyntax_create name = + let e = + Pcoq.create_constr_entry (Pcoq.get_univ "rnatural") name in + Pcoq.Gram.Unsafe.clear_entry e; + e + +let rnumber = rsyntax_create "rnumber" let _ = Gram.extend rnumber None diff --git a/parsing/g_rsyntaxnew.ml b/parsing/g_rsyntaxnew.ml new file mode 100644 index 000000000..081762ae3 --- /dev/null +++ b/parsing/g_rsyntaxnew.ml @@ -0,0 +1,113 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* RApp (dloc, RRef(dloc,glob_Ropp), [r_of_posint dloc n]) + | POS n -> r_of_posint dloc n + +(**********************************************************************) +(* Printing R via scopes *) +(**********************************************************************) + +exception Non_closed_number + +(* for numbers > 1 *) +let rec bignat_of_pos = function + (* 1+1 *) + | RApp (_,RRef (_,p), [RRef (_,o1); RRef (_,o2)]) + when p = glob_Rplus & o1 = glob_R1 & o2 = glob_R1 -> two + (* 1+1+1 *) + | RApp (_,RRef (_,p1), [RRef (_,o1); + RApp(_,RRef (_,p2),[RRef(_,o2);RRef(_,o3)])]) + when p1 = glob_Rplus & p2 = glob_Rplus & + o1 = glob_R1 & o2 = glob_R1 & o3 = glob_R1 -> three + (* (1+1)*b *) + | RApp (_,RRef (_,p), [a; b]) when p = glob_Rmult -> + if bignat_of_pos a <> two then raise Non_closed_number; + mult_2 (bignat_of_pos b) + (* 1+(1+1)*b *) + | RApp (_,RRef (_,p1), [RRef (_,o); RApp (_,RRef (_,p2),[a;b])]) + when p1 = glob_Rplus & p2 = glob_Rmult & o = glob_R1 -> + if bignat_of_pos a <> two then raise Non_closed_number; + add_1 (mult_2 (bignat_of_pos b)) + | _ -> raise Non_closed_number + +let bignat_of_r = function + | RRef (_,a) when a = glob_R0 -> zero + | RRef (_,a) when a = glob_R1 -> one + | r -> bignat_of_pos r + +let bigint_of_r = function + | RApp (_,RRef (_,o), [a]) when o = glob_Ropp -> NEG (bignat_of_r a) + | a -> POS (bignat_of_r a) + +let uninterp_r p = + try + Some (bigint_of_r p) + with Non_closed_number -> + None + +let _ = Symbols.declare_numeral_interpreter "R_scope" + ["Coq";"Reals";"Rsyntax"] + (r_of_int,None) + ([RRef(dummy_loc,glob_Ropp);RRef(dummy_loc,glob_R0); + RRef(dummy_loc,glob_Rplus);RRef(dummy_loc,glob_Rmult);RRef(dummy_loc,glob_R1)], + uninterp_r, + None) diff --git a/parsing/g_tactic.ml4 b/parsing/g_tactic.ml4 index c7f8217d1..fe77a3095 100644 --- a/parsing/g_tactic.ml4 +++ b/parsing/g_tactic.ml4 @@ -22,6 +22,10 @@ open Constr open Prim open Tactic +let tactic_kw = + [ "using"; "Orelse"; "Proof"; "Qed"; "And"; "()"; "|-" ] +let _ = List.iter (fun s -> Lexer.add_token ("",s)) tactic_kw + (* Functions overloaded by quotifier *) let induction_arg_of_constr c = @@ -319,9 +323,9 @@ GEXTEND Gram (* Constructors *) | IDENT "Left"; bl = with_binding_list -> TacLeft bl | IDENT "Right"; bl = with_binding_list -> TacRight bl - | IDENT "Split"; bl = with_binding_list -> TacSplit bl - | IDENT "Exists"; bl = binding_list -> TacSplit bl - | IDENT "Exists" -> TacSplit NoBindings + | IDENT "Split"; bl = with_binding_list -> TacSplit (false,bl) + | IDENT "Exists"; bl = binding_list -> TacSplit (true,bl) + | IDENT "Exists" -> TacSplit (true,NoBindings) | IDENT "Constructor"; n = num_or_meta; l = with_binding_list -> TacConstructor (n,l) | IDENT "Constructor"; t = OPT tactic -> TacAnyConstructor t diff --git a/parsing/g_tacticnew.ml4 b/parsing/g_tacticnew.ml4 new file mode 100644 index 000000000..4083fe82c --- /dev/null +++ b/parsing/g_tacticnew.ml4 @@ -0,0 +1,355 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* "; "<-" ] +let _ = List.iter (fun s -> Lexer.add_token("",s)) tactic_kw + +(* Hack to parse "with n := c ..." as a binder without conflicts with the *) +(* admissible notation "with c1 c2..." *) +let test_coloneq2 = + Gram.Entry.of_parser "test_int_coloneq" + (fun strm -> + match Stream.npeek 2 strm with + | [_; ("", ":=")] -> () + | _ -> raise Stream.Failure) + +(* open grammar entries, possibly in quotified form *) +ifdef Quotify then open Qast + +open Constr +open Prim +open Tactic + +(* Functions overloaded by quotifier *) + +let induction_arg_of_constr c = + try ElimOnIdent (Topconstr.constr_loc c,coerce_to_id c) with _ -> ElimOnConstr c + +let local_compute = [FBeta;FIota;FDeltaBut [];FZeta] + +let error_oldelim _ = error "OldElim no longer supported" + +ifdef Quotify then + let induction_arg_of_constr = function + | Qast.Node ("Nvar", [_;id]) -> Qast.Node ("ElimOnIdent", [id]) + | c -> Qast.Node ("ElimOnConstr", [c]) + +ifdef Quotify then +let make_red_flag s = Qast.Apply ("make_red_flag", [s]) + +ifdef Quotify then +let local_compute = + Qast.List [ + Qast.Node ("FBeta", []); + Qast.Node ("FDeltaBut", [Qast.List []]); + Qast.Node ("FIota", []); + Qast.Node ("FZeta", [])] + +ifdef Quotify then open Q + +let join_to_constr loc c2 = (fst loc), snd (Topconstr.constr_loc c2) + +(* Auxiliary grammar rules *) + +GEXTEND Gram + GLOBAL: simple_tactic constrarg constr_with_bindings quantified_hypothesis + red_expr int_or_var castedopenconstr; + + 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 metavariable) *) + id_or_ltac_ref: + [ [ id = base_ident -> AN id + | "?"; n = natural -> MetaNum (loc,n) ] ] + ; + (* Either a global ref or a ltac ref (variable or pattern metavariable) *) + global_or_ltac_ref: + [ [ qid = global -> AN qid + | "?"; n = natural -> MetaNum (loc,n) ] ] + ; + (* 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) + ] ] + ; + constrarg: + [ [ IDENT "inst"; id = identref; "["; c = Constr.lconstr; "]" -> + ConstrContext (id, c) + | IDENT "eval"; rtc = Tactic.red_expr; "in"; c = Constr.lconstr -> + ConstrEval (rtc,c) + | IDENT "check"; c = Constr.lconstr -> ConstrTypeOf c + | c = Constr.lconstr -> ConstrTerm c ] ] + ; + castedopenconstr: + [ [ c = lconstr -> c ] ] + ; + induction_arg: + [ [ n = natural -> ElimOnAnonHyp n + | c = lconstr -> induction_arg_of_constr c + ] ] + ; + quantified_hypothesis: + [ [ id = base_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) ] ] + ; + pattern_occ: + [ [ nl = LIST0 integer; c = [c=constr->c | g=global->Topconstr.CRef g]-> (nl,c) ] ] + ; + pattern_occ_hyp_list: + [ [ nl = LIST1 natural; IDENT "Goal" -> (Some nl,[]) + | nl = LIST1 natural; id = id_or_meta; (g,l) = pattern_occ_hyp_list -> + (g,(id,nl)::l) + | IDENT "Goal" -> (Some [],[]) + | id = id_or_meta; (g,l) = pattern_occ_hyp_list -> (g,(id,[])::l) ] ] + ; + clause_pattern: + [ [ "in"; p = pattern_occ_hyp_list -> p | -> None, [] ] ] + ; + intropatterns: + [ [ l = LIST0 simple_intropattern -> l ]] + ; + simple_intropattern: + [ [ "["; tc = LIST1 intropatterns SEP "|" ; "]" -> IntroOrAndPattern tc + | "("; tc = LIST1 simple_intropattern SEP "," ; ")" -> IntroOrAndPattern [tc] + | IDENT "_" -> IntroWildcard + | id = base_ident -> IntroIdentifier id + ] ] + ; + simple_binding: + [ [ id = base_ident; ":="; c = constr -> (loc, NamedHyp id, c) + | n = natural; ":="; c = constr -> (loc, AnonHyp n, c) ] ] + ; + binding_list: + [ [ test_coloneq2; n = natural; ":="; c = constr; + bl = LIST0 simple_binding -> + ExplicitBindings ((join_to_constr loc c,AnonHyp n, c) :: bl) + | test_coloneq2; id = base_ident; ":="; c2 = constr; + bl = LIST0 simple_binding -> + ExplicitBindings + ((join_to_constr loc c2,NamedHyp id, c2) :: bl) + | c1 = constr; bl = LIST0 constr -> + ImplicitBindings (c1 :: bl) ] ] + ; + constr_with_bindings: + [ [ c = constr; l = with_binding_list -> (c, l) ] ] + ; + with_binding_list: + [ [ "with"; bl = binding_list -> 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 + ] ] + ; + 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 "unfold"; ul = LIST1 unfold_occ -> Unfold ul + | IDENT "fold"; cl = LIST1 constr -> Fold cl + | IDENT "pattern"; pl = LIST1 pattern_occ -> 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 + | s = IDENT; c = constr -> ExtraRedExpr (s,c) ] ] + ; + hypident: + [ [ id = id_or_meta -> InHyp id + | "("; "type"; "of"; id = id_or_meta; ")" -> InHypType id ] ] + ; + clause: + [ [ "in"; idl = LIST1 hypident -> idl + | -> [] ] ] + ; + fixdecl: + [ [ id = base_ident; "/"; n = natural; ":"; c = constr -> (id,n,c) ] ] + ; + cofixdecl: + [ [ id = base_ident; ":"; c = constr -> (id,c) ] ] + ; + hintbases: + [ [ "with"; "*" -> None + | "with"; l = LIST1 IDENT -> Some l + | -> Some [] ] ] + ; + eliminator: + [ [ "using"; el = constr_with_bindings -> el ] ] + ; + with_names: + [ [ "as"; "["; ids = LIST1 (LIST0 base_ident) SEP "|"; "]" -> ids + | -> [] ] ] + ; + simple_tactic: + [ [ + (* Basic tactics *) + IDENT "intros"; IDENT "until"; id = quantified_hypothesis -> + TacIntrosUntil id + | IDENT "intros"; pl = intropatterns -> TacIntroPattern pl + | IDENT "intro"; id = base_ident; IDENT "after"; id2 = identref -> + TacIntroMove (Some id, Some id2) + | 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 "assumption" -> TacAssumption + | IDENT "exact"; c = lconstr -> TacExact c + + | IDENT "apply"; cl = constr_with_bindings -> TacApply cl + | IDENT "elim"; cl = constr_with_bindings; el = OPT eliminator -> + TacElim (cl,el) + | IDENT "elimtype"; c = lconstr -> TacElimType c + | IDENT "case"; cl = constr_with_bindings -> TacCase cl + | IDENT "casetype"; c = lconstr -> TacCaseType c + | "fix"; n = natural -> TacFix (None,n) + | "fix"; id = base_ident; n = natural -> TacFix (Some id,n) + | "fix"; id = base_ident; n = natural; "with"; fd = LIST0 fixdecl -> + TacMutualFix (id,n,fd) + | "cofix" -> TacCofix None + | "cofix"; id = base_ident -> TacCofix (Some id) + | "cofix"; id = base_ident; "with"; fd = LIST0 cofixdecl -> + TacMutualCofix (id,fd) + + | IDENT "cut"; c = lconstr -> TacCut c + | IDENT "assert"; c = lconstr -> + (match c with + Topconstr.CCast(_,c,t) -> TacTrueCut (Some (coerce_to_id c),t) + | _ -> TacTrueCut (None,c)) + | IDENT "assert"; c = lconstr; ":="; b = lconstr -> + TacForward (false,Names.Name (coerce_to_id c),b) + | IDENT "pose"; c = lconstr; ":="; b = lconstr -> + TacForward (true,Names.Name (coerce_to_id c),b) + | IDENT "pose"; b = lconstr -> TacForward (true,Names.Anonymous,b) + | IDENT "generalize"; lc = LIST1 constr -> TacGeneralize lc + | IDENT "generalize"; IDENT "dependent"; c = lconstr -> + TacGeneralizeDep c + | IDENT "lettac"; id = base_ident; ":="; c = lconstr; p = clause_pattern + -> TacLetTac (id,c,p) + | IDENT "instantiate"; n = natural; c = lconstr -> TacInstantiate (n,c) + + | IDENT "specialize"; n = OPT natural; lcb = constr_with_bindings -> + TacSpecialize (n,lcb) + | IDENT "lapply"; c = lconstr -> TacLApply c + + (* Derived basic tactics *) + | IDENT "induction"; h = quantified_hypothesis -> TacOldInduction h + | IDENT "newinduction"; c = induction_arg; el = OPT eliminator; + ids = with_names -> TacNewInduction (c,el,ids) + | IDENT "double"; IDENT "induction"; h1 = quantified_hypothesis; + h2 = quantified_hypothesis -> TacDoubleInduction (h1,h2) + | IDENT "destruct"; h = quantified_hypothesis -> TacOldDestruct h + | IDENT "newdestruct"; c = induction_arg; el = OPT eliminator; + ids = with_names -> TacNewDestruct (c,el,ids) + | IDENT "decompose"; IDENT "record" ; c = lconstr -> TacDecomposeAnd c + | IDENT "decompose"; IDENT "sum"; c = lconstr -> TacDecomposeOr c + | IDENT "decompose"; "["; l = LIST1 global_or_ltac_ref; "]"; c = lconstr + -> TacDecompose (l,c) + + (* Automation tactic *) + | IDENT "trivial"; db = hintbases -> TacTrivial db + | IDENT "auto"; n = OPT natural; db = hintbases -> TacAuto (n, db) + + | IDENT "autotdb"; n = OPT natural -> TacAutoTDB n + | IDENT "cdhyp"; id = identref -> TacDestructHyp (true,id) + | IDENT "dhyp"; id = identref -> TacDestructHyp (false,id) + | IDENT "dconcl" -> TacDestructConcl + | IDENT "superauto"; l = autoargs -> TacSuperAuto l + | IDENT "auto"; n = OPT natural; IDENT "decomp"; p = OPT natural -> + TacDAuto (n, p) + + (* Context management *) + | IDENT "clear"; l = LIST1 id_or_ltac_ref -> TacClear l + | IDENT "clearbody"; l = LIST1 id_or_ltac_ref -> TacClearBody l + | IDENT "move"; id1 = identref; IDENT "after"; id2 = identref -> + TacMove (true,id1,id2) + | IDENT "rename"; id1 = identref; IDENT "into"; id2 = identref -> + TacRename (id1,id2) + + (* Constructors *) + | IDENT "left"; bl = with_binding_list -> TacLeft bl + | IDENT "right"; bl = with_binding_list -> TacRight bl + | IDENT "split"; bl = with_binding_list -> TacSplit (false,bl) + | IDENT "exists"; bl = binding_list -> TacSplit (true,bl) + | IDENT "exists" -> TacSplit (true,NoBindings) + | IDENT "constructor"; n = num_or_meta; l = with_binding_list -> + TacConstructor (n,l) + | IDENT "constructor"; t = OPT tactic -> TacAnyConstructor t + + (* Equivalence relations *) + | IDENT "reflexivity" -> TacReflexivity + | IDENT "symmetry" -> TacSymmetry + | IDENT "transitivity"; c = lconstr -> TacTransitivity c + + (* 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 68e6fe858..a082772ac 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -27,6 +27,9 @@ open Genarg let evar_constr loc = CHole loc +let class_rawexpr = G_basevernac.class_rawexpr +let thm_token = G_proofs.thm_token + (* Rem: do not join the different GEXTEND into one, it breaks native *) (* compilation on PowerPC and Sun architectures *) @@ -325,15 +328,13 @@ GEXTEND Gram | IDENT "Chapter"; id = base_ident -> VernacBeginSection id ] ] ; - module_vardecls: (* This is interpreted by Pcoq.abstract_binder *) - [ [ id = base_ident; idl = ident_comma_list_tail; ":"; mty = Module.module_type + module_vardecls: + [ [ "("; id = base_ident; idl = ident_comma_list_tail; + ":"; mty = Module.module_type; ")" -> (id::idl,mty) ] ] ; - module_binders: - [ [ "["; bl = LIST1 module_vardecls SEP ";"; "]" -> bl ] ] - ; module_binders_list: - [ [ bls = LIST0 module_binders -> List.flatten bls ] ] + [ [ bl = LIST0 module_vardecls -> bl ] ] ; of_module_type: [ [ ":"; mty = Module.module_type -> (mty, true) @@ -417,7 +418,7 @@ GEXTEND Gram let l = list_tabulate (fun _ -> (CHole (loc),None)) n in CApp (loc,c,l) | None -> c in - VernacNotation ("'"^id^"'",c,[],None) + VernacNotation ("'"^id^"'",c,[],None,None) | IDENT "Implicits"; qid = global; "["; l = LIST0 natural; "]" -> VernacDeclareImplicits (qid,Some l) | IDENT "Implicits"; qid = global -> VernacDeclareImplicits (qid,None) @@ -474,7 +475,7 @@ GEXTEND Gram VernacRequireFrom (export, specif, id, filename) *) | IDENT "Require"; export = export_token; specif = specif_token; filename = STRING -> - VernacRequireFrom (export, specif, filename) + VernacRequireFrom (Some export, specif, filename) | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 STRING -> VernacDeclareMLModule l | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl) diff --git a/parsing/g_vernacnew.ml4 b/parsing/g_vernacnew.ml4 new file mode 100644 index 000000000..e88e97141 --- /dev/null +++ b/parsing/g_vernacnew.ml4 @@ -0,0 +1,817 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* ->"; ":<"; "<:" ] +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 + | d = use_default_tac; tac = Tactic.tactic -> + (fun g -> + let g = match g with Some gl -> gl | _ -> 1 in + VernacSolve(g,tac,d)) ] ] + ; + located_vernac: + [ [ v = vernac -> loc, v ] ] + ; + use_default_tac: + [ [ "!!" -> false + | -> true ] ] + ; +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 = base_ident; bl = LIST0 binder; ":"; + c = lconstr -> + VernacStartTheoremProof (thm, id, (bl, c), false, (fun _ _ -> ())) + | (f,d,e) = def_token; id = base_ident; b = def_body -> + VernacDefinition (d, id, b, f, e) + | stre = assumption_token; bl = LIST1 simple_binder_coe -> + VernacAssumption (stre, bl) + | stre = assumptions_token; bl = LIST1 simple_binder_coe -> + test_plurial_form bl; + VernacAssumption (stre, bl) + (* Gallina inductive declarations *) + | OPT[IDENT "Mutual"]; f = finite_token; + indl = LIST1 inductive_definition SEP "with" -> + VernacInductive (f,indl) + | IDENT "Fixpoint"; recs = LIST1 rec_definition SEP "with" -> + VernacFixpoint recs + | IDENT "CoFixpoint"; corecs = LIST1 corec_definition SEP "with" -> + VernacCoFixpoint corecs + | IDENT "Scheme"; l = LIST1 scheme SEP "with" -> VernacScheme l ] ] + ; + gallina_ext: + [ [ record_token; oc = opt_coercion; name = base_ident; + ps = LIST0 simple_binder; ":"; + s = lconstr; ":="; cstr = OPT base_ident; "{"; + fs = LIST0 local_decl_expr; "}" -> + VernacRecord ((oc,name),ps,s,cstr,fs) + | f = finite_token; s = csort; id = base_ident; + indpar = LIST0 simple_binder; ":="; lc = constructor_list -> + VernacInductive (f,[id,indpar,s,lc]) + ] ] + ; + thm_token: + [ [ IDENT "Theorem" -> Theorem + | IDENT "Lemma" -> Lemma + | IDENT "Fact" -> Fact + | IDENT "Remark" -> Remark ] ] + ; + def_token: + [ [ IDENT "Definition" -> (fun _ _ -> ()), Global, GDefinition + | IDENT "Local" -> (fun _ _ -> ()), Local, LDefinition + | IDENT "SubClass" -> Class.add_subclass_hook, Global, GCoercion + | IDENT "Local"; IDENT "SubClass" -> + Class.add_subclass_hook, Local, LCoercion ] ] + ; + assumption_token: + [ [ IDENT "Hypothesis" -> (Local, Logical) + | IDENT "Variable" -> (Local, Definitional) + | IDENT "Axiom" -> (Global, Logical) + | IDENT "Parameter" -> (Global, Definitional) ] ] + ; + assumptions_token: + [ [ IDENT "Hypotheses" -> (Local, Logical) + | IDENT "Variables" -> (Local, Definitional) + | IDENT "Parameters" -> (Global, Definitional) ] ] + ; + finite_token: + [ [ IDENT "Inductive" -> true + | IDENT "CoInductive" -> false ] ] + ; + record_token: + [ [ IDENT "Record" -> () | IDENT "Structure" -> () ] ] + ; + (* Simple definitions *) + def_body: + [ [ bl = LIST0 binder; ":="; red = reduce; c = lconstr -> + (match c with + CCast(_,c,t) -> DefineBody (bl, red, c, Some t) + | _ -> DefineBody (bl, red, c, None)) + | bl = LIST0 binder; ":"; t = lconstr; ":="; red = reduce; c = lconstr -> + DefineBody (bl, red, c, Some t) + | bl = LIST0 binder; ":"; t = lconstr -> + ProveBody (bl, t) ] ] + ; + reduce: + [ [ IDENT "Eval"; r = Tactic.red_expr; "in" -> Some r + | -> None ] ] + ; + (* Inductives and records *) + inductive_definition: + [ [ id = base_ident; indpar = LIST0 simple_binder; ":"; c = lconstr; ":="; + lc = constructor_list -> (id,indpar,c,lc) ] ] + ; + constructor_list: + [ [ "|"; l = LIST1 constructor_binder SEP "|" -> l + | l = LIST1 constructor_binder SEP "|" -> l + | -> [] ] ] + ; + csort: + [ [ s = sort -> CSort (loc,s) ] ] + ; + opt_coercion: + [ [ ">" -> true + | -> false ] ] + ; + (* (co)-fixpoints *) + rec_definition: + [ [ id = base_ident; bl = LIST1 binder_nodef; ":"; type_ = lconstr; + ":="; def = lconstr -> + let ni = List.length (List.flatten (List.map fst bl)) - 1 in + let loc0 = fst (List.hd (fst (List.hd bl))) in + let loc1 = join_loc loc0 (constr_loc type_) in + let loc2 = join_loc loc0 (constr_loc def) in + (id, ni, CProdN (loc1,bl,type_), CLambdaN (loc2,bl,def)) ] ] + ; + corec_definition: + [ [ id = base_ident; ":"; c = lconstr; ":="; def = lconstr -> + (id,c,def) ] ] + ; + (* Inductive schemes *) + scheme: + [ [ id = base_ident; ":="; 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 *) + simple_binder: + [ [ b = simple_binder_coe -> no_coercion loc b ] ] + ; + binder: + [ [ na = name -> LocalRawAssum([na],CHole loc) + | "("; na = name; ")" -> LocalRawAssum([na],CHole loc) + | "("; na = name; ":"; c = lconstr; ")" + -> LocalRawAssum([na],c) + | "("; na = name; ":="; c = lconstr; ")" -> + LocalRawDef (na,c) + ] ] + ; + binder_nodef: + [ [ b = binder -> + (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 *) + local_decl_expr: + [ [ id = base_ident -> (false,AssumExpr(id,CHole loc)) + | "("; id = base_ident; ")" -> (false,AssumExpr(id,CHole loc)) + | "("; id = base_ident; oc = of_type_with_opt_coercion; + t = lconstr; ")" -> + (oc,AssumExpr (id,t)) + | "("; id = base_ident; oc = of_type_with_opt_coercion; + t = lconstr; ":="; b = lconstr; ")" -> + (oc,DefExpr (id,b,Some t)) + | "("; id = base_ident; ":="; b = lconstr; ")" -> + match b with + CCast(_,b,t) -> (false,DefExpr(id,b,Some t)) + | _ -> (false,DefExpr(id,b,None)) ] ] + ; + simple_binder_coe: + [ [ id=base_ident -> (false,(id,CHole loc)) + | "("; id = base_ident; ")" -> (false,(id,CHole loc)) + | "("; id = base_ident; oc = of_type_with_opt_coercion; + t = lconstr; ")" -> + (oc,(id,t)) ] ] + ; + constructor_binder: + [ [ id = base_ident; coe = of_type_with_opt_coercion; c = lconstr -> + (coe,(id,c)) ] ] + ; + 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"; id = base_ident; + bl = LIST0 module_binder; mty_o = OPT of_module_type; + mexpr_o = OPT is_module_expr -> + VernacDefineModule (id, bl, mty_o, mexpr_o) + + | IDENT "Module"; "Type"; id = base_ident; + bl = LIST0 module_binder; mty_o = OPT is_module_type -> + VernacDeclareModuleType (id, bl, mty_o) + + | IDENT "Declare"; IDENT "Module"; id = base_ident; + bl = LIST0 module_binder; mty_o = OPT of_module_type; + mexpr_o = OPT is_module_expr -> + VernacDeclareModule (id, bl, mty_o, mexpr_o) + (* Section beginning *) + | IDENT "Section"; id = base_ident -> VernacBeginSection id + | IDENT "Chapter"; id = base_ident -> VernacBeginSection id + + (* This end a Section a Module or a Module Type *) + | IDENT "End"; id = base_ident -> 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 = 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 + | IDENT "Closed" -> None + | -> Some false ] ] + ; + 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: + [ [ "("; id = base_ident; ":"; mty = module_type; ")" -> + ([id],mty) ] ] + ; + + (* Module expressions *) + module_expr: + [ [ qid = qualid -> CMEident qid + | me1 = module_expr; me2 = module_expr -> CMEapply (me1,me2) + | "("; me = module_expr; ")" -> me +(* ... *) + ] ] + ; + with_declaration: + [ [ IDENT "Definition"; id = base_ident; ":="; c = Constr.constr -> + CWith_Definition (id,c) + | IDENT "Module"; id = base_ident; ":="; qid = qualid -> + CWith_Module (id,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 = Ast.coerce_global_to_id qid in + VernacDefinition (Global,s,d,Recordobj.add_object_hook,SCanonical) + + (* Coercions *) + | IDENT "Coercion"; qid = global; d = def_body -> + let s = Ast.coerce_global_to_id qid in + VernacDefinition (Global,s,d,Class.add_coercion_hook,GCoercion) + | IDENT "Coercion"; IDENT "Local"; qid = global; d = def_body -> + let s = Ast.coerce_global_to_id qid in + VernacDefinition (Local,s,d,Class.add_coercion_hook,LCoercion) + | IDENT "Identity"; IDENT "Coercion"; IDENT "Local"; f = base_ident; + ":"; s = class_rawexpr; ">->"; t = class_rawexpr -> + VernacIdentityCoercion (Local, f, s, t) + | IDENT "Identity"; IDENT "Coercion"; f = base_ident; ":"; + 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 "Syntactic"; IDENT "Definition"; id = IDENT; ":="; c = lconstr; + n = OPT [ "|"; n = natural -> n ] -> + let c = match n with + | Some n -> + let l = list_tabulate (fun _ -> (CHole (loc),None)) n in + CApp (loc,c,l) + | None -> c in + VernacNotation ("'"^id^"'",c,[],None,None) + | IDENT "Implicits"; qid = global; "["; l = LIST0 natural; "]" -> + VernacDeclareImplicits (qid,Some l) + | IDENT "Implicits"; qid = global -> VernacDeclareImplicits (qid,None) + + (* 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) ] ] + ; +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 = 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 = STRING -> s | s = IDENT -> s ] -> + VernacLoad (verbosely, s) + | IDENT "Declare"; IDENT "ML"; IDENT "Module"; l = LIST1 STRING -> + VernacDeclareMLModule l + + (* Dump of the universe graph - to file or to stdout *) + | IDENT "Dump"; IDENT "Universes"; fopt = OPT STRING -> + VernacPrint (PrintUniverses fopt) + + | IDENT "Locate"; l = locatable -> VernacLocate l + + (* Managing load paths *) + | IDENT "Add"; IDENT "LoadPath"; dir = STRING; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "Add"; IDENT "Rec"; IDENT "LoadPath"; dir = STRING; + alias = as_dirpath -> VernacAddLoadPath (true, dir, alias) + | IDENT "Remove"; IDENT "LoadPath"; dir = STRING -> + VernacRemoveLoadPath dir + + (* For compatibility *) + | IDENT "AddPath"; dir = STRING; "as"; alias = as_dirpath -> + VernacAddLoadPath (false, dir, alias) + | IDENT "AddRecPath"; dir = STRING; "as"; alias = as_dirpath -> + VernacAddLoadPath (true, dir, alias) + | IDENT "DelPath"; dir = STRING -> + VernacRemoveLoadPath dir + + (* 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" -> 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) + + (* 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"; qid = global; l = in_or_out_modules -> + VernacSearch (SearchAbout qid, l) + + | IDENT "Add"; IDENT "ML"; IDENT "Path"; dir = STRING -> + VernacAddMLPath (false, dir) + | IDENT "Add"; IDENT "Rec"; IDENT "ML"; IDENT "Path"; dir = STRING -> + VernacAddMLPath (true, dir) + + (* Pour intervenir sur les tables de paramčtres *) + | "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"; uni = IDENT; ent = IDENT -> + (* This should be in "syntax" section but is here for factorization*) + PrintGrammar (uni, ent) + | IDENT "LoadPath" -> PrintLoadPath + | IDENT "Modules" -> PrintModules + + | IDENT "ML"; IDENT "Path" -> PrintMLLoadPath + | IDENT "ML"; IDENT "Modules" -> PrintMLModules + | IDENT "Graph" -> PrintGraph + | IDENT "Classes" -> PrintClasses + | IDENT "Coercions" -> PrintCoercions + | IDENT "Coercion"; IDENT "Paths"; s = class_rawexpr; t = class_rawexpr + -> PrintCoercionPaths (s,t) + | IDENT "Tables" -> PrintTables + | IDENT "Proof"; qid = global -> PrintOpaqueName qid + | IDENT "Hint" -> PrintHintGoal + | IDENT "Hint"; qid = global -> PrintHint qid + | IDENT "Hint"; "*" -> PrintHintDb + | IDENT "HintDb"; s = IDENT -> PrintHintDbName s + | IDENT "Scope"; s = IDENT -> PrintScope s ] ] + ; + class_rawexpr: + [ [ IDENT "FUNCLASS" -> FunClass + | IDENT "SORTCLASS" -> SortClass + | qid = global -> RefClass qid ] ] + ; + locatable: + [ [ qid = global -> LocateTerm qid + | IDENT "File"; f = STRING -> LocateFile f + | IDENT "Library"; qid = global -> LocateLibrary qid + | s = 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 "Restore"; IDENT "State"; s = IDENT -> VernacRestoreState s + | IDENT "Restore"; IDENT "State"; s = 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 + +(* Tactic Debugger *) + | IDENT "Debug"; IDENT "On" -> VernacDebug true + | IDENT "Debug"; IDENT "Off" -> VernacDebug false + + ] ]; + END +;; + +(* Grammar extensions *) + +GEXTEND Gram + GLOBAL: syntax; + + syntax: + [ [ IDENT "Open"; IDENT "Scope"; sc = IDENT -> VernacOpenScope sc + + | IDENT "Delimits"; IDENT "Scope"; sc = IDENT; "with"; key = IDENT -> + VernacDelimiters (sc,key) + + | IDENT "Arguments"; IDENT "Scope"; qid = global; + "["; scl = LIST0 opt_scope; "]" -> VernacArgumentsScope (qid,scl) + + | IDENT "Infix"; a = entry_prec; n = OPT natural; op = STRING; + p = global; + modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; + sc = OPT [ ":"; sc = IDENT -> sc ] -> + let (a,n,b) = Metasyntax.interp_infix_modifiers a n modl in + VernacInfix (a,n,op,p,b,None,sc) + | IDENT "Notation"; s = IDENT; ":="; c = constr -> + VernacNotation ("'"^s^"'",c,[],None,None) + | IDENT "Notation"; s = STRING; ":="; c = constr; + modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ]; + sc = OPT [ ":"; sc = IDENT -> sc ] -> + VernacNotation (s,c,modl,None,sc) + + | IDENT "Grammar"; IDENT "tactic"; IDENT "simple_tactic"; + OPT [ ":"; IDENT "tactic" ]; ":="; + OPT "|"; tl = LIST0 grammar_tactic_rule SEP "|" -> + VernacTacticGrammar 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"; s = STRING; + l = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ] + -> VernacSyntaxExtension (s,l) + + (* "Print" "Grammar" should be here but is in "command" entry in order + to factorize with other "Print"-based vernac entries *) + ] ] + ; + univ: + [ [ univ = IDENT -> + set_default_action_parser (parser_type_from_name univ); univ ] ] + ; + syntax_modifier: + [ [ x = IDENT; IDENT "at"; IDENT "level"; n = natural -> + SetItemLevel ([x],n) + | x = IDENT; ","; l = LIST1 IDENT SEP ","; IDENT "at"; IDENT "level"; + n = natural -> SetItemLevel (x::l,n) + | 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 ] ] + ; + syntax_extension_type: + [ [ IDENT "ident" -> ETIdent | IDENT "global" -> ETReference + | IDENT "bigint" -> ETBigint + | e=IDENT -> ETOther("constr",e) + ] ] + ; + 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; "["; s = STRING; pil = LIST0 production_item; "]"; + "->"; "["; t = Tactic.tactic; "]" -> (name, (s,pil), t) ]] + ; + grammar_rule: + [[ name = rule_name; "["; pil = LIST0 production_item; "]"; "->"; + a = action -> (name, pil, a) ]] + ; + rule_name: + [[ name = IDENT -> name ]] + ; + production_item: + [[ s = STRING -> 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 = STRING -> 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 = STRING; 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.lconstr -> + Termast.ast_of_rawconstr + (Constrintern.interp_rawconstr 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 + +(* Reinstall tactic and vernac extensions *) +let _ = Egrammar.reset_extend_grammars_v8() diff --git a/parsing/g_zsyntax.ml b/parsing/g_zsyntax.ml index 5103aedb0..11c88d113 100644 --- a/parsing/g_zsyntax.ml +++ b/parsing/g_zsyntax.ml @@ -54,11 +54,15 @@ let z_of_string pos_or_neg s dloc = aZERO (* Declare the primitive parser with Grammar and without the scope mechanism *) -open Pcoq +let zsyntax_create name = + let e = + Pcoq.create_constr_entry (Pcoq.get_univ "znatural") name in + Pcoq.Gram.Unsafe.clear_entry e; + e -let number = create_constr_entry (get_univ "znatural") "number" +let number = zsyntax_create "number" -let negnumber = create_constr_entry (get_univ "znatural") "negnumber" +let negnumber = zsyntax_create "negnumber" let _ = Gram.extend number None diff --git a/parsing/g_zsyntaxnew.ml b/parsing/g_zsyntaxnew.ml new file mode 100644 index 000000000..24b6d4137 --- /dev/null +++ b/parsing/g_zsyntaxnew.ml @@ -0,0 +1,168 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* RApp (dloc, ref_xO,[pos_of q]) + | (q,true) when is_nonzero q -> RApp (dloc,ref_xI,[pos_of q]) + | (q,true) -> ref_xH + in + pos_of x + +let interp_positive dloc = function + | POS n when is_nonzero n -> pos_of_bignat dloc n + | _ -> + user_err_loc (dloc, "interp_positive", + str "Only strictly positive numbers in type \"positive\"!") + +let rec pat_pos_of_bignat dloc x name = + match div2_with_rest x with + | (q,false) -> + PatCstr (dloc,path_of_xO,[pat_pos_of_bignat dloc q Anonymous],name) + | (q,true) when is_nonzero q -> + PatCstr (dloc,path_of_xI,[pat_pos_of_bignat dloc q Anonymous],name) + | (q,true) -> + PatCstr (dloc,path_of_xH,[],name) + +let pat_interp_positive dloc = function + | POS n -> pat_pos_of_bignat dloc n + | NEG n -> + user_err_loc (dloc, "interp_positive", + str "No negative number in type \"positive\"!") + +(**********************************************************************) +(* Printing positive via scopes *) +(**********************************************************************) + +exception Non_closed_number + +let rec bignat_of_pos = function + | RApp (_, RRef (_,b),[a]) when b = glob_xO -> mult_2(bignat_of_pos a) + | RApp (_, RRef (_,b),[a]) when b = glob_xI -> add_1(mult_2(bignat_of_pos a)) + | RRef (_, a) when a = glob_xH -> Bignat.one + | _ -> raise Non_closed_number + +let uninterp_positive p = + try + Some (POS (bignat_of_pos p)) + with Non_closed_number -> + None + +(***********************************************************************) +(* Declaring interpreters and uninterpreters for positive *) +(***********************************************************************) + +let _ = Symbols.declare_numeral_interpreter "positive_scope" + ["Coq";"ZArith";"Zsyntax"] + (interp_positive,Some pat_interp_positive) + ([RRef (dummy_loc, glob_xI); + RRef (dummy_loc, glob_xO); + RRef (dummy_loc, glob_xH)], + uninterp_positive, + None) + +(**********************************************************************) +(* Parsing Z via scopes *) +(**********************************************************************) + +let z_path = make_path fast_integer_module (id_of_string "Z") +let glob_z = IndRef (z_path,0) +let path_of_ZERO = ((z_path,0),1) +let path_of_POS = ((z_path,0),2) +let path_of_NEG = ((z_path,0),3) +let glob_ZERO = ConstructRef path_of_ZERO +let glob_POS = ConstructRef path_of_POS +let glob_NEG = ConstructRef path_of_NEG + +let z_of_posint dloc pos_or_neg n = + if is_nonzero n then + let sgn = if pos_or_neg then glob_POS else glob_NEG in + RApp(dloc, RRef (dloc,sgn), [pos_of_bignat dloc n]) + else + RRef (dloc, glob_ZERO) + +let z_of_int dloc z = + match z with + | POS n -> z_of_posint dloc true n + | NEG n -> z_of_posint dloc false n + +let pat_z_of_posint dloc pos_or_neg n name = + if is_nonzero n then + let sgn = if pos_or_neg then path_of_POS else path_of_NEG in + PatCstr (dloc, sgn, [pat_pos_of_bignat dloc n Anonymous], name) + else + PatCstr (dloc, path_of_ZERO, [], name) + +let pat_z_of_int dloc n name = + match n with + | POS n -> pat_z_of_posint dloc true n name + | NEG n -> pat_z_of_posint dloc false n name + +(**********************************************************************) +(* Printing Z via scopes *) +(**********************************************************************) + +let bigint_of_z = function + | RApp (_, RRef (_,b),[a]) when b = glob_POS -> POS (bignat_of_pos a) + | RApp (_, RRef (_,b),[a]) when b = glob_NEG -> NEG (bignat_of_pos a) + | RRef (_, a) when a = glob_ZERO -> POS (Bignat.zero) + | _ -> raise Non_closed_number + +let uninterp_z p = + try + Some (bigint_of_z p) + with Non_closed_number -> None + +(***********************************************************************) +(* Declaring interpreters and uninterpreters for Z *) + +let _ = Symbols.declare_numeral_interpreter "Z_scope" + ["Coq";"ZArith";"Zsyntax"] + (z_of_int,Some pat_z_of_int) + ([RRef (dummy_loc, glob_ZERO); + RRef (dummy_loc, glob_POS); + RRef (dummy_loc, glob_NEG)], + uninterp_z, + None) diff --git a/parsing/g_zsyntaxnew.mli b/parsing/g_zsyntaxnew.mli new file mode 100644 index 000000000..a8370f630 --- /dev/null +++ b/parsing/g_zsyntaxnew.mli @@ -0,0 +1,11 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* "; "\\/"; "/\\"; "|-"; ":="; "<->"; "!"; "$"; "%"; "&"; - "*"; "+"; ","; "@"; "^"; "#"; "\\"; "/"; "-"; "<"; ">"; - "|"; "?"; "="; "~"; "'"; "<<"; ">>"; "<>"; "::"; - "<:"; ":<"; "=>"; ">->" ] + unfreeze (empty_ttree, empty_ttree) let _ = init() @@ -318,20 +303,11 @@ let parse_226_tail tk = parser | [< len = ident (store 0 '\226') >] -> TokIdent (get_buff len) -(* Parse a token in a char stream *) -let rec next_token = parser bp - | [< ''\n'; s >] -> (match !current with - | BeVernac s -> current := InComment s - | InComment _ -> add_comment_pp (fnl ()) - | _ -> ()); next_token s - | [< '' ' | '\t'; s >] -> (match !current with - | BeVernac _ | InComment _ -> add_comment_pp (spc ()) - | _ -> ()); next_token s - | [< ''\r'; s >] -> next_token s - | [< ''$'; len = ident (store 0 '$') >] ep -> - (("METAIDENT", get_buff len), (bp,ep)) - | [< ''.' as c; t = parser +(* Parse what foolows a dot *) +let parse_after_dot bp c strm = + if !Options.v7 then + match strm with parser | [< ' ('_' | 'a'..'z' | 'A'..'Z' as c); len = ident (store 0 c) >] -> ("FIELD", get_buff len) (* Greek utf-8 letters [CE80-CEBF and CF80-CFBF] (CE=206; BF=191) *) @@ -349,10 +325,26 @@ let rec next_token = parser bp | [< ' ('\192'..'\214' | '\216'..'\246' | '\248'..'\255' as c); len = ident (store 0 c) >] -> ("FIELD", get_buff len) -(* - | [< >] -> ("", ".") >] ep -> (t, (bp,ep)) -*) - | [< (t,_) = process_chars bp c >] -> t >] ep -> + | [< (t,_) = process_chars bp c >] -> t + else + match strm with parser + | [< (t,_) = process_chars bp c >] -> t + + +(* Parse a token in a char stream *) + +let rec next_token = parser bp + | [< ''\n'; s >] -> (match !current with + | BeVernac s -> current := InComment s + | InComment _ -> add_comment_pp (fnl ()) + | _ -> ()); next_token s + | [< '' ' | '\t'; s >] -> (match !current with + | BeVernac _ | InComment _ -> add_comment_pp (spc ()) + | _ -> ()); next_token s + | [< ''\r'; s >] -> next_token s + | [< ''$'; len = ident (store 0 '$') >] ep -> + (("METAIDENT", get_buff len), (bp,ep)) + | [< ''.' as c; t = parse_after_dot bp c >] ep -> current:=BeVernac ""; (t, (bp,ep)) | [< ' ('_' | 'a'..'z' | 'A'..'Z' as c); len = ident (store 0 c) >] ep -> let id = get_buff len in current:=InVernac; diff --git a/parsing/pcoq.ml4 b/parsing/pcoq.ml4 index 33948d83b..6cff40c34 100644 --- a/parsing/pcoq.ml4 +++ b/parsing/pcoq.ml4 @@ -344,7 +344,7 @@ module Constr = (* Entries that can be refered via the string -> Gram.Entry.e table *) let constr = gec_constr "constr" - let constr9 = gec_constr "constr9" + let operconstr = gec_constr "operconstr" let constr_eoi = eoi_entry constr let lconstr = gec_constr "lconstr" let ident = make_gen_entry uconstr rawwit_ident "ident" @@ -398,14 +398,36 @@ module Vernac_ = let syntax = gec_vernac "syntax_command" let vernac = gec_vernac "Vernac_.vernac" - (* Various vernacular entries needed to be exported *) - let thm_token = Gram.Entry.create "vernac:thm_token" - let class_rawexpr = Gram.Entry.create "vernac:class_rawexpr" - let vernac_eoi = eoi_entry vernac end +(* Prim is not re-initialized *) +let reset_all_grammars () = + let f = Gram.Unsafe.clear_entry in + List.iter f + [Constr.constr;Constr.operconstr;Constr.lconstr;Constr.annot; + Constr.constr_pattern]; + f Constr.ident; f Constr.global; f Constr.sort; f Constr.pattern; + f Module.module_expr; f Module.module_type; + f Tactic.simple_tactic; + f Tactic.castedopenconstr; + f Tactic.constr_with_bindings; + f Tactic.constrarg; + f Tactic.quantified_hypothesis; + f Tactic.int_or_var; + f Tactic.red_expr; + f Tactic.tactic_arg; + f Tactic.tactic; + f Vernac_.gallina; + f Vernac_.gallina_ext; + f Vernac_.command; + f Vernac_.syntax; + f Vernac_.vernac; + Lexer.init() + + + let main_entry = Gram.Entry.create "vernac" GEXTEND Gram @@ -425,8 +447,6 @@ open Vernac_ let globalizer = ref (fun x -> failwith "No globalizer") let set_globalizer f = globalizer := f -let f = (ast : Coqast.t G.Entry.e) - let define_ast_quotation default s (e:Coqast.t G.Entry.e) = (if default then GEXTEND Gram @@ -462,7 +482,7 @@ GEXTEND Gram dynconstr: [ [ a = Constr.constr -> ConstrNode a (* For compatibility *) - | "<<"; a = Constr.constr; ">>" -> ConstrNode a ] ] + | "<<"; a = Constr.lconstr; ">>" -> ConstrNode a ] ] ; dyncasespattern: [ [ a = Constr.pattern -> CasesPatternNode a ] ]; END @@ -527,9 +547,7 @@ let adjust_level assoc from = function | _ -> Some (Some n) let compute_entry allow_create adjust = function - | ETConstr (10,_) -> weaken_entry Constr.lconstr, None, true - | ETConstr (9,_) -> weaken_entry Constr.constr9, None, true - | ETConstr (n,q) -> weaken_entry Constr.constr, adjust (n,q), false + | ETConstr (n,q) -> weaken_entry Constr.operconstr, adjust (n,q), false | ETIdent -> weaken_entry Constr.ident, None, false | ETBigint -> weaken_entry Prim.bigint, None, false | ETReference -> weaken_entry Constr.global, None, false @@ -543,19 +561,50 @@ let compute_entry allow_create adjust = function with e when allow_create -> create_entry u n ConstrArgType in object_of_typed_entry e, None, true -let get_constr_entry = compute_entry true (fun (n,()) -> Some n) - -let get_constr_production_entry ass from = - compute_entry false (adjust_level ass from) +let get_constr_entry en = + match en with + ETConstr(200,_) when not !Options.v7 -> + snd (get_entry (get_univ "constr") "binder_constr"), + None, + false + | _ -> compute_entry true (fun (n,()) -> Some n) en + +let get_constr_production_entry ass from en = + (* first 2 cases to help factorisation *) + match en with + | ETConstr (10,q) when !Options.v7 -> + weaken_entry Constr.lconstr, None, false + | ETConstr (8,q) when !Options.v7 -> + weaken_entry Constr.constr, None, false + | _ -> compute_entry false (adjust_level ass from) en let constr_prod_level = function - | 8 -> "top" - | 4 -> "4L" + | 4 when !Options.v7 -> "4L" | n -> string_of_int n +let is_self from e = + match from, e with + ETConstr(n,_), ETConstr(n', + BorderProd(false,Some(Gramext.NonA|Gramext.LeftA))) -> false + | ETConstr(n,_), ETConstr(n',_) -> n=n' + | (ETIdent,ETIdent | ETReference, ETReference | ETBigint,ETBigint + | ETPattern, ETPattern) -> true + | ETOther(s1,s2), ETOther(s1',s2') -> s1=s1' & s2=s2' + | _ -> false + +let is_binder_level from e = + match from, e with + ETConstr(200,_), ETConstr(200,_) -> not !Options.v7 + | _ -> false + let symbol_of_production assoc from typ = - match get_constr_production_entry assoc from typ with - | (eobj,None,_) -> Gramext.Snterm (Gram.Entry.obj eobj) - | (eobj,Some None,_) -> Gramext.Snext - | (eobj,Some (Some lev),_) -> - Gramext.Snterml (Gram.Entry.obj eobj,constr_prod_level lev) + if is_binder_level from typ then + let eobj = snd (get_entry (get_univ "constr") "operconstr") in + Gramext.Snterml (Gram.Entry.obj eobj,"200") + else if is_self from typ then Gramext.Sself + else + match get_constr_production_entry assoc from typ with + | (eobj,None,_) -> Gramext.Snterm (Gram.Entry.obj eobj) + | (eobj,Some None,_) -> Gramext.Snext + | (eobj,Some (Some lev),_) -> + Gramext.Snterml (Gram.Entry.obj eobj,constr_prod_level lev) diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index b15046c29..325d4f494 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -129,9 +129,9 @@ module Prim : module Constr : sig val constr : constr_expr Gram.Entry.e - val constr9 : constr_expr Gram.Entry.e val constr_eoi : constr_expr Gram.Entry.e val lconstr : constr_expr Gram.Entry.e + val operconstr : constr_expr Gram.Entry.e val ident : identifier Gram.Entry.e val global : reference Gram.Entry.e val sort : rawsort Gram.Entry.e @@ -164,8 +164,6 @@ module Tactic : module Vernac_ : sig open Decl_kinds - val thm_token : theorem_kind Gram.Entry.e - val class_rawexpr : class_rawexpr Gram.Entry.e val gallina : vernac_expr Gram.Entry.e val gallina_ext : vernac_expr Gram.Entry.e val command : vernac_expr Gram.Entry.e @@ -173,3 +171,5 @@ module Vernac_ : val vernac : vernac_expr Gram.Entry.e val vernac_eoi : vernac_expr Gram.Entry.e end + +val reset_all_grammars : unit -> unit diff --git a/parsing/ppconstr.ml b/parsing/ppconstr.ml index 57e813632..97c7d637b 100644 --- a/parsing/ppconstr.ml +++ b/parsing/ppconstr.ml @@ -22,55 +22,6 @@ open Topconstr open Term (*i*) -let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; - -let pr_global ref = pr_global_env None ref - -let global_const_name kn = - try pr_global (ConstRef kn) - with Not_found -> (* May happen in debug *) - (str ("CONST("^(string_of_kn kn)^")")) - -let global_var_name id = - try pr_global (VarRef id) - with Not_found -> (* May happen in debug *) - (str ("SECVAR("^(string_of_id id)^")")) - -let global_ind_name (kn,tyi) = - try pr_global (IndRef (kn,tyi)) - with Not_found -> (* May happen in debug *) - (str ("IND("^(string_of_kn kn)^","^(string_of_int tyi)^")")) - -let global_constr_name ((kn,tyi),i) = - try pr_global (ConstructRef ((kn,tyi),i)) - with Not_found -> (* May happen in debug *) - (str ("CONSTRUCT("^(string_of_kn kn)^","^(string_of_int tyi) - ^","^(string_of_int i)^")")) - -let globpr gt = match gt with - | Nvar(_,s) -> (pr_id s) - | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) - | Node(_,"CONST",[Path(_,sl)]) -> - global_const_name (section_path sl) - | Node(_,"SECVAR",[Nvar(_,s)]) -> - global_var_name s - | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> - global_ind_name (section_path sl, tyi) - | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> - global_constr_name ((section_path sl, tyi), i) - | Dynamic(_,d) -> - if (Dyn.tag d) = "constr" then (str"") - else dfltpr gt - | gt -> dfltpr gt - -let wrap_exception = function - Anomaly (s1,s2) -> - warning ("Anomaly ("^s1^")"); pp s2; - str"" - | Failure _ | UserError _ | Not_found -> - str"" - | s -> raise s - let latom = 0 let lannot = 1 let lprod = 8 (* not 1 because the scope extends to 8 on the right *) @@ -157,6 +108,8 @@ let pr_binder pr (nal,t) = let pr_binders pr bl = hv 0 (prlist_with_sep pr_semicolon (pr_binder pr) bl) +let pr_global vars ref = pr_global_env vars ref + let rec pr_lambda_tail pr bll = function | CLambdaN (_,bl,a) -> pr_lambda_tail pr (bll ++ pr_semicolon() ++ pr_binders pr bl) a @@ -230,21 +183,22 @@ let pr_annotation pr = function | None -> mt () | Some t -> str "<" ++ pr ltop t ++ str ">" ++ brk (0,2) -let rec pr_pat = function - | CPatAlias (_,p,x) -> pr_pat p ++ spc () ++ str "as" ++ spc () ++ pr_id x +let rec pr_cases_pattern = function + | CPatAlias (_,p,x) -> + pr_cases_pattern p ++ spc () ++ str "as" ++ spc () ++ pr_id x | CPatCstr (_,c,[]) -> pr_reference c | CPatCstr (_,c,pl) -> hov 0 ( str "(" ++ pr_reference c ++ spc () ++ - prlist_with_sep spc pr_pat pl ++ str ")") + prlist_with_sep spc pr_cases_pattern pl ++ str ")") | CPatAtom (_,Some c) -> pr_reference c | CPatAtom (_,None) -> str "_" | CPatNumeral (_,n) -> Bignat.pr_bigint n - | CPatDelimiters (_,key,p) -> pr_delimiters key (pr_pat p) + | CPatDelimiters (_,key,p) -> pr_delimiters key (pr_cases_pattern p) let pr_eqn pr (_,patl,rhs) = hov 0 ( - prlist_with_sep spc pr_pat patl ++ spc () ++ + prlist_with_sep spc pr_cases_pattern patl ++ spc () ++ str "=>" ++ brk (1,1) ++ pr ltop rhs) ++ spc () @@ -394,26 +348,3 @@ let rec pr_may_eval pr = function str "[" ++ pr c ++ str "]") | ConstrTypeOf c -> hov 0 (str "Check " ++ brk (1,1) ++ pr c) | ConstrTerm c -> pr c - -(**********************************************************************) -let constr_syntax_universe = "constr" -(* This is starting precedence for printing constructions or tactics *) -(* Level 9 means no parentheses except for applicative terms (at level 10) *) -let constr_initial_prec = Some (9,Ppextend.L) - -let gentermpr_fail gt = - Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt - -let gentermpr gt = - try gentermpr_fail gt - with s -> wrap_exception s - -(* [at_top] means ids of env must be avoided in bound variables *) - -let gentermpr_core at_top env t = - gentermpr (Termast.ast_of_constr at_top env t) -(* -let gentermpr_core at_top env t = - pr_constr (Constrextern.extern_constr at_top env t) -*) - diff --git a/parsing/ppconstr.mli b/parsing/ppconstr.mli index 1a2848f00..6dd3d842c 100644 --- a/parsing/ppconstr.mli +++ b/parsing/ppconstr.mli @@ -23,10 +23,7 @@ open Util val split_fix : int -> constr_expr -> constr_expr -> (name located list * constr_expr) list * constr_expr * constr_expr -val pr_global : global_reference -> std_ppcmds - -val gentermpr : Coqast.t -> std_ppcmds -val gentermpr_core : bool -> env -> constr -> std_ppcmds +val pr_global : Idset.t -> global_reference -> std_ppcmds val pr_opt : ('a -> std_ppcmds) -> 'a option -> std_ppcmds val pr_name : name -> std_ppcmds @@ -39,4 +36,5 @@ val pr_occurrences : ('a -> std_ppcmds) -> 'a occurrences -> 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_cases_pattern : cases_pattern_expr -> std_ppcmds val pr_may_eval : ('a -> std_ppcmds) -> 'a may_eval -> std_ppcmds diff --git a/parsing/pptactic.ml b/parsing/pptactic.ml index 2f18076b7..fbeb697eb 100644 --- a/parsing/pptactic.ml +++ b/parsing/pptactic.ml @@ -13,22 +13,32 @@ open Names open Nameops open Util open Extend -open Ppconstr open Tacexpr open Rawterm open Topconstr open Genarg open Libnames +let pr_red_expr = Ppconstr.pr_red_expr +let pr_may_eval = Ppconstr.pr_may_eval +let pr_sort = Ppconstr.pr_sort +let pr_global = Ppconstr.pr_global Idset.empty +let pr_name = Ppconstr.pr_name +let pr_opt = Ppconstr.pr_opt +let pr_occurrences = Ppconstr.pr_occurrences + (* Extensions *) +let prtac_tab_v7 = Hashtbl.create 17 let prtac_tab = Hashtbl.create 17 -let declare_extra_tactic_pprule s f g = - Hashtbl.add prtac_tab s (f,g) +let declare_extra_tactic_pprule for_v8 s f g = + Hashtbl.add prtac_tab_v7 s (f,g); + if for_v8 then Hashtbl.add prtac_tab s (f,g) +let genarg_pprule_v7 = ref Stringmap.empty let genarg_pprule = ref Stringmap.empty -let declare_extra_genarg_pprule (rawwit, f) (wit, g) = +let declare_extra_genarg_pprule for_v8 (rawwit, f) (wit, g) = let s = match unquote wit with | ExtraArgType s -> s | _ -> error @@ -36,7 +46,9 @@ let declare_extra_genarg_pprule (rawwit, f) (wit, g) = in let f x = f (out_gen rawwit x) in let g x = g (out_gen wit x) in - genarg_pprule := Stringmap.add s (f,g) !genarg_pprule + genarg_pprule_v7 := Stringmap.add s (f,g) !genarg_pprule_v7; + if for_v8 then + genarg_pprule := Stringmap.add s (f,g) !genarg_pprule (* [pr_rawtac] is here to cheat with ML typing system, gen_tactic_expr is polymorphic but with some occurrences of its @@ -66,8 +78,6 @@ let pr_or_meta pr = function | AI x -> pr x | _ -> failwith "pr_hyp_location: unexpected quotation meta-variable" -let pr_casted_open_constr = pr_constr - let pr_quantified_hypothesis = function | AnonHyp n -> int n | NamedHyp id -> pr_id id @@ -133,9 +143,9 @@ let pr_induction_arg prc = function | ElimOnAnonHyp n -> int n let pr_match_pattern = function - | Term a -> pr_pattern a - | Subterm (None,a) -> str "[" ++ pr_pattern a ++ str "]" - | Subterm (Some id,a) -> pr_id id ++ str "[" ++ pr_pattern a ++ str "]" + | Term a -> Ppconstr.pr_pattern a + | Subterm (None,a) -> str "[" ++ Ppconstr.pr_pattern a ++ str "]" + | Subterm (Some id,a) -> pr_id id ++ str "[" ++ Ppconstr.pr_pattern a ++ str "]" let pr_match_hyps = function | NoHypId mp -> str "_:" ++ pr_match_pattern mp @@ -146,7 +156,8 @@ let pr_match_rule m pr = function str "[" ++ pr_match_pattern mp ++ str "]" ++ spc () ++ str "->" ++ brk (1,2) ++ pr t | Pat (rl,mp,t) -> - str "[" ++ prlist_with_sep pr_semicolon pr_match_hyps rl ++ spc () ++ + str "[" ++ prlist_with_sep pr_semicolon + pr_match_hyps rl ++ spc () ++ str "|-" ++ spc () ++ pr_match_pattern mp ++ spc () ++ str "]" ++ spc () ++ str "->" ++ brk (1,2) ++ pr t | All t -> str "_" ++ spc () ++ str "->" ++ brk (1,2) ++ pr t @@ -192,7 +203,16 @@ let pr_autoarg_usingTDB = function | true -> spc () ++ str "Using TDB" | false -> mt () -let rec pr_rawgen prtac x = +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 rec pr_rawgen prc prtac x = match Genarg.genarg_tag x with | BoolArgType -> pr_arg str (if out_gen rawwit_bool x then "true" else "false") | IntArgType -> pr_arg int (out_gen rawwit_int x) @@ -202,45 +222,44 @@ let rec pr_rawgen prtac x = | IdentArgType -> pr_arg pr_id (out_gen rawwit_ident x) | RefArgType -> pr_arg pr_reference (out_gen rawwit_ref x) | SortArgType -> pr_arg pr_sort (out_gen rawwit_sort x) - | ConstrArgType -> pr_arg pr_constr (out_gen rawwit_constr x) + | ConstrArgType -> pr_arg prc (out_gen rawwit_constr x) | ConstrMayEvalArgType -> - pr_arg (pr_may_eval pr_constr) (out_gen rawwit_constr_may_eval x) + pr_arg (pr_may_eval prc) (out_gen rawwit_constr_may_eval x) | QuantHypArgType -> pr_arg pr_quantified_hypothesis (out_gen rawwit_quant_hyp x) | RedExprArgType -> - pr_arg (pr_red_expr (pr_constr,pr_metanum pr_reference)) (out_gen rawwit_red_expr x) + pr_arg (pr_red_expr (prc,pr_metanum pr_reference)) (out_gen rawwit_red_expr x) | TacticArgType -> pr_arg prtac (out_gen rawwit_tactic x) | CastedOpenConstrArgType -> - pr_arg pr_casted_open_constr (out_gen rawwit_casted_open_constr x) + pr_arg prc (out_gen rawwit_casted_open_constr x) | ConstrWithBindingsArgType -> - pr_arg (pr_with_bindings pr_constr) + pr_arg (pr_with_bindings prc) (out_gen rawwit_constr_with_bindings x) | List0ArgType _ -> - hov 0 (fold_list0 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) + hov 0 (fold_list0 (fun x a -> pr_rawgen prc prtac x ++ a) x (mt())) | List1ArgType _ -> - hov 0 (fold_list1 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) - | OptArgType _ -> hov 0 (fold_opt (pr_rawgen prtac) (mt()) x) + hov 0 (fold_list1 (fun x a -> pr_rawgen prc prtac x ++ a) x (mt())) + | OptArgType _ -> hov 0 (fold_opt (pr_rawgen prc prtac) (mt()) x) | PairArgType _ -> hov 0 (fold_pair - (fun a b -> pr_rawgen prtac a ++ spc () ++ pr_rawgen prtac b) + (fun a b -> pr_rawgen prc prtac a ++ spc () ++ pr_rawgen prc + prtac b) x) | ExtraArgType s -> - try fst (Stringmap.find s !genarg_pprule) x + let tab = if Options.do_translate() then !genarg_pprule + else !genarg_pprule_v7 in + try fst (Stringmap.find s tab) x with Not_found -> str " [no printer for " ++ str s ++ str "] " -let rec pr_raw_tacarg_using_rule pr_gen = function - | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_raw_tacarg_using_rule pr_gen (l,al) - | Egrammar.TacNonTerm _ :: l, a :: al -> pr_gen a ++ pr_raw_tacarg_using_rule pr_gen (l,al) - | [], [] -> mt () - | _ -> failwith "Inconsistent arguments of extended tactic" - -let pr_raw_extend prt s l = +let pr_raw_extend prc prt s l = + let prg = pr_rawgen prc prt in + let tab = if Options.do_translate() then prtac_tab else prtac_tab_v7 in try - let (s,pl) = fst (Hashtbl.find prtac_tab s) l in - str s ++ pr_raw_tacarg_using_rule (pr_rawgen prt) (pl,l) + let (s,pl) = fst (Hashtbl.find tab s) l in + str s ++ pr_tacarg_using_rule prg (pl,l) with Not_found -> - str "TODO(" ++ str s ++ prlist (pr_rawgen prt) l ++ str ")" + str "TODO(" ++ str s ++ prlist prg l ++ str ")" open Closure @@ -284,23 +303,21 @@ let rec pr_generic prtac x = (fun a b -> pr_generic prtac a ++ spc () ++ pr_generic prtac b) x) | ExtraArgType s -> - try snd (Stringmap.find s !genarg_pprule) x + let tab = if Options.do_translate() then !genarg_pprule + else !genarg_pprule_v7 in + try snd (Stringmap.find s tab) x with Not_found -> str "[no printer for " ++ str s ++ str "]" -let rec pr_tacarg_using_rule prt = function - | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_tacarg_using_rule prt (l,al) - | Egrammar.TacNonTerm _ :: l, a :: al -> pr_generic prt a ++ pr_tacarg_using_rule prt (l,al) - | [], [] -> mt () - | _ -> failwith "Inconsistent arguments of extended tactic" - let pr_extend prt s l = + let prg = pr_generic prt in + let tab = if Options.do_translate() then prtac_tab else prtac_tab_v7 in try - let (s,pl) = snd (Hashtbl.find prtac_tab s) l in - str s ++ pr_tacarg_using_rule prt (pl,l) + let (s,pl) = snd (Hashtbl.find tab s) l in + str s ++ pr_tacarg_using_rule prg (pl,l) with Not_found -> - str s ++ prlist (pr_generic prt) l + str s ++ prlist prg l -let make_pr_tac (pr_constr,pr_cst,pr_ind,pr_ident,pr_extend) = +let make_pr_tac (pr_constr,pr_pattern,pr_cst,pr_ind,pr_ident,pr_extend) = let pr_bindings = pr_bindings pr_constr in let pr_with_bindings = pr_with_bindings pr_constr in @@ -445,7 +462,7 @@ and pr_atom1 = function (* Constructors *) | TacLeft l -> hov 1 (str "Left" ++ pr_bindings l) | TacRight l -> hov 1 (str "Right" ++ pr_bindings l) - | TacSplit l -> hov 1 (str "Split" ++ pr_bindings l) + | TacSplit (_,l) -> hov 1 (str "Split" ++ pr_bindings l) | TacAnyConstructor (Some t) -> hov 1 (str "Constructor" ++ pr_arg !pr_rawtac0 t) | TacAnyConstructor None as t -> pr_atom0 t @@ -546,13 +563,15 @@ and pr6 = function hov 0 (str "Match" ++ spc () ++ pr_may_eval Ppconstr.pr_constr t ++ spc() ++ str "With" ++ prlist - (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule true prtac r) + (fun r -> fnl () ++ str "|" ++ spc () ++ + pr_match_rule true prtac r) lrul) | TacMatchContext (lr,lrul) -> hov 0 ( str (if lr then "Match Reverse Context With" else "Match Context With") ++ prlist - (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule false prtac r) + (fun r -> fnl () ++ str "|" ++ spc () ++ + pr_match_rule false prtac r) lrul) | TacFun (lvar,body) -> hov 0 (str "Fun" ++ @@ -586,10 +605,11 @@ in (prtac,pr0,pr_match_rule) let (pr_raw_tactic,pr_raw_tactic0,pr_match_rule) = make_pr_tac (Ppconstr.pr_constr, + Ppconstr.pr_constr, pr_metanum pr_reference, pr_reference, pr_or_meta (fun (loc,id) -> pr_id id), - pr_raw_extend) + pr_raw_extend Ppconstr.pr_constr) let _ = pr_rawtac := pr_raw_tactic let _ = pr_rawtac0 := pr_raw_tactic0 @@ -597,7 +617,9 @@ let _ = pr_rawtac0 := pr_raw_tactic0 let (pr_tactic,_,_) = make_pr_tac (Printer.prterm, + Ppconstr.pr_constr, pr_evaluable_reference, pr_inductive, pr_id, pr_extend) + diff --git a/parsing/pptactic.mli b/parsing/pptactic.mli index ca16c21e5..186a726f6 100644 --- a/parsing/pptactic.mli +++ b/parsing/pptactic.mli @@ -13,20 +13,24 @@ open Genarg open Tacexpr open Pretyping open Proof_type +open Topconstr +(* if the boolean is false then the extension applies only to old syntax *) val declare_extra_genarg_pprule : + bool -> ('a raw_abstract_argument_type * ('a -> std_ppcmds)) -> ('b closed_abstract_argument_type * ('b -> std_ppcmds)) -> unit +(* idem *) val declare_extra_tactic_pprule : - string -> + bool -> string -> (raw_generic_argument list -> string * Egrammar.grammar_tactic_production list) -> (closed_generic_argument list -> string * Egrammar.grammar_tactic_production list) -> unit -val pr_match_pattern : Tacexpr.pattern_expr match_pattern -> std_ppcmds +val pr_match_pattern : pattern_expr match_pattern -> std_ppcmds val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) -> (pattern_expr,raw_tactic_expr) match_rule -> std_ppcmds @@ -34,3 +38,18 @@ val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) -> val pr_raw_tactic : raw_tactic_expr -> std_ppcmds val pr_tactic : Proof_type.tactic_expr -> std_ppcmds + +val pr_rawgen: + (constr_expr -> std_ppcmds) -> + (raw_tactic_expr -> std_ppcmds) -> + (constr_expr, raw_tactic_expr) generic_argument -> + std_ppcmds +val pr_raw_extend: + (constr_expr -> std_ppcmds) -> + (raw_tactic_expr -> std_ppcmds) -> string -> + (constr_expr, raw_tactic_expr) generic_argument list -> + std_ppcmds +val pr_extend : + (raw_tactic_expr -> std_ppcmds) -> string -> + (Term.constr, raw_tactic_expr) generic_argument list -> + std_ppcmds diff --git a/parsing/prettyp.ml b/parsing/prettyp.ml index 4404c1929..648d4690d 100644 --- a/parsing/prettyp.ml +++ b/parsing/prettyp.ml @@ -411,12 +411,6 @@ let print_local_context () = in (print_var_rec env ++ print_last_const env) -let fprint_var name typ = - (str ("*** [" ^ name ^ " :") ++ fprtype typ ++ str "]" ++ fnl ()) - -let fprint_judge {uj_val=trm;uj_type=typ} = - (fprterm trm ++ str" : " ++ fprterm typ) - 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/printer.ml b/parsing/printer.ml index 72fe499e6..df2aabdf8 100644 --- a/parsing/printer.ml +++ b/parsing/printer.ml @@ -28,15 +28,91 @@ open Ppconstr let emacs_str s = if !Options.print_emacs then s else "" -(* These are the names of the universes where the pp rules for constr and - tactic are put (must be consistent with files PPConstr.v and PPTactic.v *) - -let tactic_syntax_universe = "tactic" +(**********************************************************************) +(* Old Ast printing *) +let constr_syntax_universe = "constr" (* This is starting precedence for printing constructions or tactics *) (* Level 9 means no parentheses except for applicative terms (at level 10) *) -let tactic_initial_prec = Some ((tactic_syntax_universe,(9,0,0)),Ppextend.L) - +let constr_initial_prec_v7 = Some (9,Ppextend.L) +let constr_initial_prec = Some (200,Ppextend.E) + +let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; + +let global_const_name kn = + try pr_global Idset.empty (ConstRef kn) + with Not_found -> (* May happen in debug *) + (str ("CONST("^(string_of_kn kn)^")")) + +let global_var_name id = + try pr_global Idset.empty (VarRef id) + with Not_found -> (* May happen in debug *) + (str ("SECVAR("^(string_of_id id)^")")) + +let global_ind_name (kn,tyi) = + try pr_global Idset.empty (IndRef (kn,tyi)) + with Not_found -> (* May happen in debug *) + (str ("IND("^(string_of_kn kn)^","^(string_of_int tyi)^")")) + +let global_constr_name ((kn,tyi),i) = + try pr_global Idset.empty (ConstructRef ((kn,tyi),i)) + with Not_found -> (* May happen in debug *) + (str ("CONSTRUCT("^(string_of_kn kn)^","^(string_of_int tyi) + ^","^(string_of_int i)^")")) + +let globpr gt = match gt with + | Nvar(_,s) -> (pr_id s) + | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) + | Node(_,"CONST",[Path(_,sl)]) -> + global_const_name (section_path sl) + | Node(_,"SECVAR",[Nvar(_,s)]) -> + global_var_name s + | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> + global_ind_name (section_path sl, tyi) + | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> + global_constr_name ((section_path sl, tyi), i) + | Dynamic(_,d) -> + if (Dyn.tag d) = "constr" then (str"") + else dfltpr gt + | gt -> dfltpr gt + + +let wrap_exception = function + Anomaly (s1,s2) -> + warning ("Anomaly ("^s1^")"); pp s2; + str"" + | Failure _ | UserError _ | Not_found -> + str"" + | s -> raise s + +let gentermpr_fail gt = + let prec = + if !Options.v7 then constr_initial_prec_v7 else constr_initial_prec in + Esyntax.genprint globpr constr_syntax_universe prec gt + +let gentermpr gt = + try gentermpr_fail gt + with s -> wrap_exception s + +(**********************************************************************) +(* Generic printing: choose old or new printers *) + +(* [at_top] means ids of env must be avoided in bound variables *) +let gentermpr_core at_top env t = + if !Options.v7 then gentermpr (Termast.ast_of_constr at_top env t) + else Ppconstrnew.pr_constr (Constrextern.extern_constr at_top env t) +let pr_cases_pattern t = + if !Options.v7 then gentermpr (Termast.ast_of_cases_pattern t) + else Ppconstrnew.pr_cases_pattern + (Constrextern.extern_cases_pattern Idset.empty t) +let pr_pattern_env tenv env t = + if !Options.v7 then gentermpr (Termast.ast_of_pattern tenv env t) + else Ppconstrnew.pr_constr + (Constrextern.extern_pattern tenv env t) + +(**********************************************************************) +(* Derived printers *) + let prterm_env_at_top env = gentermpr_core true env let prterm_env env = gentermpr_core false env let prtype_env env typ = prterm_env env typ @@ -48,25 +124,14 @@ let prterm t = prterm_env (Global.env()) t let prtype t = prtype_env (Global.env()) t let prjudge j = prjudge_env (Global.env()) j -let fprterm_env a = - warning "Fw printing not implemented, use CCI printing 1"; prterm_env a -let fprterm a = - warning "Fw printing not implemented, use CCI printing 2"; prterm a - -let fprtype_env env typ = - warning "Fw printing not implemented, use CCI printing 3"; prtype_env env typ -let fprtype a = - warning "Fw printing not implemented, use CCI printing 4"; prtype a +let pr_constant env cst = prterm_env env (mkConst cst) +let pr_existential env ev = prterm_env env (mkEvar ev) +let pr_inductive env ind = prterm_env env (mkInd ind) +let pr_constructor env cstr = prterm_env env (mkConstruct cstr) +let pr_global = pr_global Idset.empty -(* For compatibility *) -let fterm0 = fprterm_env - -let pr_constant env cst = gentermpr (ast_of_constant env cst) -let pr_existential env ev = gentermpr (ast_of_existential env ev) -let pr_inductive env ind = gentermpr (ast_of_inductive env ind) -let pr_constructor env cstr = gentermpr (ast_of_constructor env cstr) - -let pr_global = pr_global +let pr_rawterm t = + pr_constr (Constrextern.extern_rawconstr Idset.empty t) open Pattern let pr_ref_label = function (* On triche sur le contexte *) @@ -75,49 +140,8 @@ let pr_ref_label = function (* On triche sur le contexte *) | CstrNode sp -> pr_constructor (Global.env()) sp | VarNode id -> pr_id id -let pr_cases_pattern t = gentermpr (Termast.ast_of_cases_pattern t) -(* -let pr_rawterm t = gentermpr (Termast.ast_of_rawconstr t) -*) -let pr_rawterm t = pr_constr (Constrextern.extern_rawconstr t) -let pr_pattern_env tenv env t = gentermpr (Termast.ast_of_pattern tenv env t) let pr_pattern t = pr_pattern_env (Global.env()) empty_names_context t -(* -let gentacpr x = Pptactic.prtac x -*) - -(* -and default_tacpr = function - | Nvar(_,s) -> (pr_id s) - - (* constr's may occur inside tac expressions ! *) - | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) - | Node(_,"CONST",[Path(_,sl)]) -> - let sp = section_path sl in - pr_global (ConstRef sp) - | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> - let sp = section_path sl in - pr_global (IndRef (sp,tyi)) - | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> - let sp = section_path sl in - pr_global (ConstructRef ((sp,tyi),i)) - - (* This should be tactics *) - | Node(_,s,[]) -> (str s) -(* TODO - | Node(_,s,ta) -> - (str s ++ brk(1,2) ++ hov 0 (prlist_with_sep pr_spc gentacpr ta)) -*) - | Dynamic(_,d) as gt -> - let tg = Dyn.tag d in - if tg = "tactic" then (str"") - else if tg = "value" then (str"") - else if tg = "constr" then (str"") - else dfltpr gt - | gt -> dfltpr gt -*) - let pr_var_decl env (id,c,typ) = let pbody = match c with | None -> (mt ()) diff --git a/parsing/printer.mli b/parsing/printer.mli index 48ee955cc..b69a9955b 100644 --- a/parsing/printer.mli +++ b/parsing/printer.mli @@ -58,13 +58,3 @@ val pr_context_of : env -> std_ppcmds val emacs_str : string -> string -(*i*) -val fprterm_env : env -> constr -> std_ppcmds -val fprterm : constr -> std_ppcmds -val fprtype_env : env -> types -> std_ppcmds -val fprtype : types -> std_ppcmds - -(* For compatibility *) -val fterm0 : env -> constr -> std_ppcmds -(*i*) - diff --git a/parsing/q_coqast.ml4 b/parsing/q_coqast.ml4 index 1f352f6af..2b3967707 100644 --- a/parsing/q_coqast.ml4 +++ b/parsing/q_coqast.ml4 @@ -407,8 +407,9 @@ let rec mlexpr_of_atomic_tactic = function <:expr< Tacexpr.TacLeft $mlexpr_of_binding_kind l$>> | Tacexpr.TacRight l -> <:expr< Tacexpr.TacRight $mlexpr_of_binding_kind l$>> - | Tacexpr.TacSplit l -> - <:expr< Tacexpr.TacSplit $mlexpr_of_binding_kind l$>> + | Tacexpr.TacSplit (b,l) -> + <:expr< Tacexpr.TacSplit + ($mlexpr_of_bool b$,$mlexpr_of_binding_kind l$)>> | Tacexpr.TacAnyConstructor t -> <:expr< Tacexpr.TacAnyConstructor $mlexpr_of_option mlexpr_of_tactic t$>> | Tacexpr.TacConstructor (n,l) -> diff --git a/parsing/search.ml b/parsing/search.ml index 74cdd77dd..fc5792fa0 100644 --- a/parsing/search.ml +++ b/parsing/search.ml @@ -107,7 +107,7 @@ let plain_display ref a c = let filter_by_module (module_list:dir_path list) (accept:bool) (ref:global_reference) _ _ = try - let sp = sp_of_global None ref in + let sp = sp_of_global ref in let sl = dirpath sp in let rec filter_aux = function | m :: tl -> (not (is_dirpath_prefix_of m sl)) && (filter_aux tl) diff --git a/parsing/tacextend.ml4 b/parsing/tacextend.ml4 index 593fb0169..a3d5f496e 100644 --- a/parsing/tacextend.ml4 +++ b/parsing/tacextend.ml4 @@ -124,7 +124,16 @@ let make_printing_rule l = (<:patt< _ >>,None,<:expr< failwith "Tactic extension: cannot occur" >>) in List.fold_right add_printing_clause l [default] -let declare_tactic loc s cl = +let new_tac_ext (s,cl) = + (String.lowercase s, List.map + (fun (s,l,e) -> + (String.lowercase s, List.map + (function TacTerm s -> TacTerm (String.lowercase s) + | t -> t) l, + e)) + cl) + +let declare_tactic_v7 loc s cl = let pl = make_printing_rule cl in let gl = mlexpr_of_clause cl in let hide_tac (_,p,e) = @@ -142,13 +151,44 @@ let declare_tactic loc s cl = <:str_item< declare open Pcoq; - declare $list:List.map hide_tac cl$ end; + Egrammar.extend_tactic_grammar $se$ $gl$; + let pp = fun [ $list:pl$ ] in + Pptactic.declare_extra_tactic_pprule False $se$ pp pp; + end + >> + +let declare_tactic loc s cl = + let (s',cl') = new_tac_ext (s,cl) in + let pl' = make_printing_rule cl' in + let gl' = mlexpr_of_clause cl' in + let se' = mlexpr_of_string s' in + let pl = make_printing_rule cl in + let gl = mlexpr_of_clause cl in + let hide_tac (_,p,e) = + (* reste a definir les fonctions cachees avec des noms frais *) + let stac = "h_"^s' in + let e = + make_fun + <:expr< + Refiner.abstract_extended_tactic $mlexpr_of_string s'$ $make_args p$ $e$ + >> + p in + <:str_item< value $lid:stac$ = $e$ >> + in + let se = mlexpr_of_string s in + <:str_item< + declare + open Pcoq; + declare $list:List.map hide_tac cl'$ end; try - Refiner.add_tactic $se$ (fun [ $list:make_clauses s cl$ ]) + Refiner.add_tactic $se'$ (fun [ $list:make_clauses s' cl'$ ]) with e -> Pp.pp (Cerrors.explain_exn e); - Egrammar.extend_tactic_grammar $se$ $gl$; + Egrammar.extend_tactic_grammar $se'$ $gl'$; + let pp' = fun [ $list:pl'$ ] in + Pptactic.declare_extra_tactic_pprule True $se'$ pp' pp'; + Egrammar.extend_tactic_grammar $se'$ $gl$; let pp = fun [ $list:pl$ ] in - Pptactic.declare_extra_tactic_pprule $se$ pp pp; + Pptactic.declare_extra_tactic_pprule False $se'$ pp pp; end >> @@ -194,7 +234,11 @@ EXTEND [ [ "TACTIC"; "EXTEND"; s = [ UIDENT | LIDENT ]; OPT "|"; l = LIST1 tacrule SEP "|"; "END" -> - declare_tactic loc s l ] ] + declare_tactic loc s l + | "V7"; "TACTIC"; "EXTEND"; s = [ UIDENT | LIDENT ]; + OPT "|"; l = LIST1 tacrule SEP "|"; + "END" -> + declare_tactic_v7 loc s l ] ] ; tacrule: [ [ "["; s = STRING; l = LIST0 tacargs; "]"; "->"; "["; e = Pcaml.expr; "]" diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 2b452ecbb..aa2c6c1db 100755 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -255,7 +255,7 @@ let inductive_class_of ind = fst (class_info (CL_IND ind)) let class_args_of c = snd (decompose_app c) let strength_of_cl = function - | CL_CONST kn -> constant_strength (sp_of_global None (ConstRef kn)) + | CL_CONST kn -> constant_strength (sp_of_global (ConstRef kn)) | CL_SECVAR sp -> variable_strength sp | _ -> Global @@ -263,11 +263,11 @@ let string_of_class = function | CL_FUN -> "FUNCLASS" | CL_SORT -> "SORTCLASS" | CL_CONST sp -> - string_of_qualid (shortest_qualid_of_global None (ConstRef sp)) + string_of_qualid (shortest_qualid_of_global Idset.empty (ConstRef sp)) | CL_IND sp -> - string_of_qualid (shortest_qualid_of_global None (IndRef sp)) + string_of_qualid (shortest_qualid_of_global Idset.empty (IndRef sp)) | CL_SECVAR sp -> - string_of_qualid (shortest_qualid_of_global None (VarRef sp)) + string_of_qualid (shortest_qualid_of_global Idset.empty (VarRef sp)) (* coercion_value : coe_index -> unsafe_judgment * bool *) @@ -393,7 +393,7 @@ let coercion_of_qualid qid = let coe = coe_of_reference ref in if not (coercion_exists coe) then errorlabstrm "try_add_coercion" - (Nametab.pr_global_env None ref ++ str" is not a coercion"); + (Nametab.pr_global_env Idset.empty ref ++ str" is not a coercion"); coe module CoercionPrinting = @@ -401,7 +401,7 @@ module CoercionPrinting = type t = coe_typ let encode = coercion_of_qualid let subst = subst_coe_typ - let printer x = pr_global_env None x + let printer x = pr_global_env Idset.empty x let key = Goptions.SecondaryTable ("Printing","Coercion") let title = "Explicitly printed coercions: " let member_message x b = diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index f0f07e06e..f676717d7 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -70,7 +70,7 @@ module PrintingCasesMake = let kn' = subst_kn subst kn in if kn' == kn then obj else (kn',i), ints - let printer (ind,_) = pr_global_env None (IndRef ind) + let printer (ind,_) = pr_global_env Idset.empty (IndRef ind) let key = Goptions.SecondaryTable ("Printing",Test.field) let title = Test.title let member_message x = Test.member_message (printer x) diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index ed549a77e..1b7a515b9 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -560,7 +560,7 @@ let lookup_eliminator ind_sp s = (* inductive type *) let ref = ConstRef (make_kn mp dp (label_of_id id)) in try - let _ = sp_of_global None ref in + let _ = sp_of_global ref in constr_of_reference ref with Not_found -> (* Then try to get a user-defined eliminator in some other places *) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 7613cd6df..2a74ede72 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -33,7 +33,7 @@ let set_transparent_const sp = if cb.const_body <> None & cb.const_opaque then errorlabstrm "set_transparent_const" (str "Cannot make" ++ spc () ++ - Nametab.pr_global_env None (ConstRef sp) ++ + Nametab.pr_global_env Idset.empty (ConstRef sp) ++ spc () ++ str "transparent because it was declared opaque."); Conv_oracle.set_transparent_const sp diff --git a/pretyping/termops.ml b/pretyping/termops.ml index b233ce244..a41631bd2 100644 --- a/pretyping/termops.ml +++ b/pretyping/termops.ml @@ -537,7 +537,19 @@ let first_char id = let lowercase_first_char id = String.lowercase (first_char id) -let id_of_global env ref = Nametab.id_of_global (Some (named_context env)) ref +let vars_of_env env = + let s = + Sign.fold_named_context (fun (id,_,_) s -> Idset.add id s) + (named_context env) ~init:Idset.empty in + Sign.fold_rel_context + (fun (na,_,_) s -> match na with Name id -> Idset.add id s | _ -> s) + (rel_context env) ~init:s + +let add_vname vars = function + Name id -> Idset.add id vars + | _ -> vars + +let id_of_global = Nametab.id_of_global let sort_hdchar = function | Prop(_) -> "P" @@ -558,9 +570,9 @@ let hdchar env c = if i=0 then lowercase_first_char (id_of_label (label kn)) else - lowercase_first_char (id_of_global env (IndRef x)) + lowercase_first_char (id_of_global (IndRef x)) | Construct ((sp,i) as x) -> - lowercase_first_char (id_of_global env (ConstructRef x)) + lowercase_first_char (id_of_global (ConstructRef x)) | Var id -> lowercase_first_char id | Sort s -> sort_hdchar s | Rel n -> @@ -671,12 +683,12 @@ let occur_rel p env id = let occur_id env nenv id0 c = let rec occur n c = match kind_of_term c with | Var id when id=id0 -> raise Occur - | Const kn when id_of_global env (ConstRef kn) = id0 -> raise Occur + | Const kn when id_of_global (ConstRef kn) = id0 -> raise Occur | Ind ind_sp - when id_of_global env (IndRef ind_sp) = id0 -> + when id_of_global (IndRef ind_sp) = id0 -> raise Occur | Construct cstr_sp - when id_of_global env (ConstructRef cstr_sp) = id0 -> + when id_of_global (ConstructRef cstr_sp) = id0 -> raise Occur | Rel p when p>n & occur_rel (p-n) nenv id0 -> raise Occur | _ -> iter_constr_with_binders succ occur n c diff --git a/pretyping/termops.mli b/pretyping/termops.mli index 0cd757ced..cd5d7def0 100644 --- a/pretyping/termops.mli +++ b/pretyping/termops.mli @@ -114,7 +114,6 @@ val eta_eq_constr : constr -> constr -> bool (* finding "intuitive" names to hypotheses *) val first_char : identifier -> string val lowercase_first_char : identifier -> string -(*val id_of_global : env -> Libnames.global_reference -> identifier*) val sort_hdchar : sorts -> string val hdchar : env -> types -> string val id_of_name_using_hdchar : @@ -144,6 +143,10 @@ val ids_of_named_context : named_context -> identifier list val ids_of_context : env -> identifier list val names_of_rel_context : env -> names_context +(* Set of local names *) +val vars_of_env: env -> Idset.t +val add_vname : Idset.t -> name -> Idset.t + (* sets of free identifiers *) type used_idents = identifier list val occur_rel : int -> name list -> identifier -> bool diff --git a/proofs/tacexpr.ml b/proofs/tacexpr.ml index 2761fbbdf..1c6d7a6f6 100644 --- a/proofs/tacexpr.ml +++ b/proofs/tacexpr.ml @@ -146,7 +146,7 @@ type ('constr,'cst,'ind,'id) gen_atomic_tactic_expr = (* Constructors *) | TacLeft of 'constr substitution | TacRight of 'constr substitution - | TacSplit of 'constr substitution + | TacSplit of bool * 'constr substitution | TacAnyConstructor of raw_tactic_expr option | TacConstructor of int or_metaid * 'constr substitution diff --git a/scripts/coqmktop.ml b/scripts/coqmktop.ml index ad6623328..0a1144714 100644 --- a/scripts/coqmktop.ml +++ b/scripts/coqmktop.ml @@ -36,6 +36,7 @@ let proofs = split_cmo Tolink.proofs let tactics = split_cmo Tolink.tactics let toplevel = split_cmo Tolink.toplevel let highparsing = split_cmo Tolink.highparsing +let highparsingnew = split_cmo Tolink.highparsingnew let core_objs = libobjs @ lib @ kernel @ library @ pretyping @ interp @ parsing @ @@ -63,6 +64,7 @@ let top = ref false let searchisos = ref false let coqide = ref false let echo = ref false +let newsyntax = ref false let src_dirs = [ []; [ "config" ]; [ "toplevel" ] ] @@ -92,9 +94,10 @@ let module_of_file name = let files_to_link userfiles = let dyn_objs = if not !opt then dynobjs else [] in let command_objs = if !searchisos then coqsearch else [] in - let toplevel_objs = if !top then topobjs else if !opt then notopobjs else [] - in - let objs = core_objs @ dyn_objs @ toplevel @ highparsing @ + let toplevel_objs = + if !top then topobjs else if !opt then notopobjs else [] in + let parsobjs = if !newsyntax then highparsingnew else highparsing in + let objs = core_objs @ dyn_objs @ toplevel @ parsobjs @ command_objs @ hightactics @ toplevel_objs in let tolink = if !opt then @@ -140,7 +143,8 @@ Options are: -top Build Coq on a ocaml toplevel (incompatible with -opt) -searchisos Build a toplevel for SearchIsos -ide Build a toplevel for the Coq IDE - -R dir Specify recursively directories for Ocaml\n"; + -R dir Specify recursively directories for Ocaml + -v8 Link with V8 grammar\n"; exit 1 (* parsing of the command line *) @@ -156,6 +160,8 @@ let parse_args () = searchisos := true; parse (op,fl) rem | "-ide" :: rem -> coqide := true; parse (op,fl) rem + | "-v8" :: rem -> + newsyntax := true; parse (op,fl) rem | "-echo" :: rem -> echo := true ; parse (op,fl) rem | ("-cclib"|"-ccopt"|"-I"|"-o" as o) :: rem' -> begin diff --git a/states/MakeInitialNew.v b/states/MakeInitialNew.v new file mode 100644 index 000000000..8a321839b --- /dev/null +++ b/states/MakeInitialNew.v @@ -0,0 +1,10 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* basename (sp_of_global None (reference_of_constr c)) + | None -> id_of_global (reference_of_constr c) | Some n -> n in (n,c) in add_resolves env sigma (List.map f lhints) dbnames @@ -440,7 +440,7 @@ let add_hints dbnames h = let f (n,c) = let c = Constrintern.interp_constr sigma env c in let n = match n with - | None -> basename (sp_of_global None (reference_of_constr c)) + | None -> id_of_global (reference_of_constr c) | Some n -> n in (n,c) in add_trivials env sigma (List.map f lhints) dbnames @@ -448,7 +448,7 @@ let add_hints dbnames h = let f (n,locqid) = let r = Nametab.global locqid in let n = match n with - | None -> basename (sp_of_global None r) + | None -> id_of_global r | Some n -> n in (n,r) in add_unfolds (List.map f lhints) dbnames @@ -900,7 +900,7 @@ let default_superauto g = superauto !default_search_depth [] [] g let interp_to_add gl locqid = let r = Nametab.global locqid in - let id = basename (sp_of_global None r) in + let id = id_of_global r in (next_ident_away id (pf_ids_of_hyps gl), constr_of_reference r) let gen_superauto nopt l a b gl = diff --git a/tactics/eauto.ml4 b/tactics/eauto.ml4 index bb965213b..442f18c45 100644 --- a/tactics/eauto.ml4 +++ b/tactics/eauto.ml4 @@ -47,22 +47,28 @@ let e_resolve_with_bindings_tac (c,lbind) gl = let e_resolve_constr c gls = e_resolve_with_bindings_tac (c,NoBindings) gls -TACTIC EXTEND EExact +(* V8 TACTIC EXTEND eexact +| [ "eexact" constr(c) ] -> [ e_give_exact c ] +END*) +TACTIC EXTEND Eexact | [ "EExact" constr(c) ] -> [ e_give_exact c ] END -let e_give_exact_constr = h_eExact +let e_give_exact_constr = h_eexact let registered_e_assumption gl = tclFIRST (List.map (fun id gl -> e_give_exact_constr (mkVar id) gl) (pf_ids_of_hyps gl)) gl (* This automatically define h_eApply (among other things) *) -TACTIC EXTEND EApply +(*V8 TACTIC EXTEND eapply + [ "eapply" constr_with_bindings(c) ] -> [ e_resolve_with_bindings_tac c ] +END*) +TACTIC EXTEND eapply [ "EApply" constr_with_bindings(c) ] -> [ e_resolve_with_bindings_tac c ] END -let vernac_e_resolve_constr c = h_eApply (c,NoBindings) +let vernac_e_resolve_constr c = h_eapply (c,NoBindings) (************************************************************************) (* PROLOG tactic *) @@ -89,6 +95,9 @@ let prolog_tac l n gl = with UserError ("Refiner.tclFIRST",_) -> errorlabstrm "Prolog.prolog" (str "Prolog failed") +(* V8 TACTIC EXTEND prolog +| [ "prolog" "[" constr_list(l) "]" int_or_var(n) ] -> [ prolog_tac l n ] +END*) TACTIC EXTEND Prolog | [ "Prolog" "[" constr_list(l) "]" int_or_var(n) ] -> [ prolog_tac l n ] END @@ -370,16 +379,26 @@ let pr_hintbases = function | Some l -> str " with " ++ Util.prlist str l let _ = - Pptactic.declare_extra_genarg_pprule + Pptactic.declare_extra_genarg_pprule + true (rawwit_hintbases, pr_hintbases) (wit_hintbases, pr_hintbases) +(* V8 TACTIC EXTEND eauto +| [ "eauto" int_or_var_opt(n) int_or_var_opt(p) hintbases(db) ] -> + [ gen_eauto false (make_dimension n p) db ] +END + +V8 TACTIC EXTEND eautodebug +| [ "eautod" int_or_var_opt(n) int_or_var_opt(p) hintbases(db) ] -> + [ gen_eauto true (make_dimension n p) db ] +END*) TACTIC EXTEND EAuto | [ "EAuto" int_or_var_opt(n) int_or_var_opt(p) hintbases(db) ] -> [ gen_eauto false (make_dimension n p) db ] END -TACTIC EXTEND EAutoDebug +V7 TACTIC EXTEND EAutodebug | [ "EAutod" int_or_var_opt(n) int_or_var_opt(p) hintbases(db) ] -> [ gen_eauto true (make_dimension n p) db ] END diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4 index 2d89b84f5..cc61d4ea8 100644 --- a/tactics/extratactics.ml4 +++ b/tactics/extratactics.ml4 @@ -18,6 +18,9 @@ open Extraargs (* Equality *) open Equality +(* V8 TACTIC EXTEND rewrite + [ "rewrite" orient(b) constr_with_bindings(c) ] -> [general_rewrite_bindings b c] +END*) TACTIC EXTEND Rewrite [ "Rewrite" orient(b) constr_with_bindings(c) ] -> [general_rewrite_bindings b c] END @@ -104,6 +107,13 @@ TACTIC EXTEND Inversion -> [ dinv (Some false) c id ] END +(* V8 TACTIC EXTEND inversionclear +| [ "inversion_clear" quantified_hypothesis(id) ] -> [ inv (Some true) id ] +| [ "inversion_clear" quantified_hypothesis(id) "in" ne_ident_list(l) ] + -> [ invIn_gen (Some true) id l] +| [ "dependent" "inversion_clear" quantified_hypothesis(id) with_constr(c) ] + -> [ dinv (Some true) c id ] +END*) TACTIC EXTEND InversionClear | [ "Inversion_clear" quantified_hypothesis(id) ] -> [ inv (Some true) id ] | [ "Inversion_clear" quantified_hypothesis(id) "in" ne_ident_list(l) ] diff --git a/tactics/hiddentac.ml b/tactics/hiddentac.ml index 749a5b12d..3e6542853 100644 --- a/tactics/hiddentac.ml +++ b/tactics/hiddentac.ml @@ -72,7 +72,7 @@ let h_rename id1 id2 = (* Constructors *) let h_left l = abstract_tactic (TacLeft l) (left l) let h_right l = abstract_tactic (TacLeft l) (right l) -let h_split l = abstract_tactic (TacSplit l) (split l) +let h_split l = abstract_tactic (TacSplit (false,l)) (split l) (* Moved to tacinterp because of dependence in Tacinterp.interp let h_any_constructor t = abstract_tactic (TacAnyConstructor t) (any_constructor t) diff --git a/tactics/setoid_replace.ml b/tactics/setoid_replace.ml index 4c2166e4e..389943f30 100644 --- a/tactics/setoid_replace.ml +++ b/tactics/setoid_replace.ml @@ -530,7 +530,7 @@ let add_morphism lem_name (m,profil) = lem2 = None})))); Options.if_verbose ppnl (prterm m ++str " is registered as a morphism") let morphism_hook stre ref = - let pf_id = basename (sp_of_global None ref) in + let pf_id = id_of_global ref in if (is_edited pf_id) then (add_morphism pf_id (what_edited pf_id); no_more_edited pf_id) diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml index 50dab07ad..87d756e1f 100644 --- a/tactics/tacinterp.ml +++ b/tactics/tacinterp.ml @@ -278,6 +278,10 @@ let glob_name l ist = function | Anonymous -> Anonymous | Name id -> Name (glob_ident l ist id) +let vars_of_ist (lfun,_,_,env) = + List.fold_left (fun s id -> Idset.add id s) + (vars_of_env env) lfun + let get_current_context () = try Pfedit.get_current_goal_context () with e when Logic.catchable_exception e -> @@ -310,12 +314,14 @@ let glob_hyp_or_metanum ist = function | MetaNum (_loc,n) -> MetaNum (loc,glob_metanum ist loc n) let glob_qualid_or_metanum ist = function - | AN qid -> AN (Qualid(loc,qualid_of_sp (sp_of_global None (Nametab.global qid)))) + | AN qid -> AN (Qualid(loc, + shortest_qualid_of_global (vars_of_ist ist) (Nametab.global qid))) | MetaNum (_loc,n) -> MetaNum (loc,glob_metanum ist loc n) let glob_reference ist = function | Ident (loc,id) as r when find_ident id ist -> r - | r -> Qualid (loc,qualid_of_sp (sp_of_global None (Nametab.global r))) + | r -> Qualid (loc, + shortest_qualid_of_global (vars_of_ist ist) (Nametab.global r)) let glob_ltac_qualid ist ref = let (loc,qid) = qualid_of_reference ref in @@ -534,7 +540,7 @@ let rec glob_atomic lf (_,_,_,_ as ist) = function (* Constructors *) | TacLeft bl -> TacLeft (glob_bindings ist bl) | TacRight bl -> TacRight (glob_bindings ist bl) - | TacSplit bl -> TacSplit (glob_bindings ist bl) + | TacSplit (b,bl) -> TacSplit (b,glob_bindings ist bl) | TacAnyConstructor t -> TacAnyConstructor (option_app (glob_tactic ist) t) | TacConstructor (n,bl) -> TacConstructor (n, glob_bindings ist bl) @@ -827,7 +833,7 @@ let constr_to_id loc c = else invalid_arg_loc (loc, "Not an identifier") let constr_to_qid loc c = - try qualid_of_sp (sp_of_global None (reference_of_constr c)) + try shortest_qualid_of_global Idset.empty (reference_of_constr c) with _ -> invalid_arg_loc (loc, "Not a global reference") (* Check for LetTac *) @@ -1556,7 +1562,7 @@ and interp_atomic ist gl = function (* Constructors *) | TacLeft bl -> h_left (bindings_interp ist gl bl) | TacRight bl -> h_right (bindings_interp ist gl bl) - | TacSplit bl -> h_split (bindings_interp ist gl bl) + | TacSplit (_,bl) -> h_split (bindings_interp ist gl bl) | TacAnyConstructor t -> abstract_tactic (TacAnyConstructor t) (Tactics.any_constructor (option_app (tactic_interp ist) t)) diff --git a/theories/Arith/Between.v b/theories/Arith/Between.v index 112edaa56..5711a81a7 100755 --- a/theories/Arith/Between.v +++ b/theories/Arith/Between.v @@ -10,6 +10,7 @@ Require Le. Require Lt. +Import nat_scope. Section Between. Variables P,Q : nat -> Prop. diff --git a/theories/Arith/Bool_nat.v b/theories/Arith/Bool_nat.v index 872c314f1..bbe1475f5 100644 --- a/theories/Arith/Bool_nat.v +++ b/theories/Arith/Bool_nat.v @@ -11,6 +11,7 @@ Require Export Compare_dec. Require Export Peano_dec. Require Sumbool. +Import nat_scope. (** The decidability of equality and order relations over type [nat] give some boolean functions with the adequate specification. *) diff --git a/theories/Arith/Compare.v b/theories/Arith/Compare.v index f98115e6b..5b12033dc 100755 --- a/theories/Arith/Compare.v +++ b/theories/Arith/Compare.v @@ -9,6 +9,7 @@ (*i $Id$ i*) (** Equality is decidable on [nat] *) +Import nat_scope. Lemma not_eq_sym : (A:Set)(p,q:A)(~p=q) -> ~(q=p). Proof sym_not_eq. diff --git a/theories/Arith/Compare_dec.v b/theories/Arith/Compare_dec.v index 67218de83..735d267d6 100755 --- a/theories/Arith/Compare_dec.v +++ b/theories/Arith/Compare_dec.v @@ -12,6 +12,7 @@ Require Le. Require Lt. Require Gt. Require Decidable. +Import nat_scope. Definition zerop : (n:nat){n=O}+{lt O n}. NewDestruct n; Auto with arith. diff --git a/theories/Arith/Div.v b/theories/Arith/Div.v index 2b7cfac14..e4b0300dd 100755 --- a/theories/Arith/Div.v +++ b/theories/Arith/Div.v @@ -9,6 +9,7 @@ (*i $Id$ i*) (** Euclidean division *) +Import nat_scope. Require Le. Require Euclid_def. diff --git a/theories/Arith/Div2.v b/theories/Arith/Div2.v index f039aa7a0..dd0cc8458 100644 --- a/theories/Arith/Div2.v +++ b/theories/Arith/Div2.v @@ -12,6 +12,7 @@ Require Lt. Require Plus. Require Compare_dec. Require Even. +Import nat_scope. (** Here we define [n/2] and prove some of its properties *) diff --git a/theories/Arith/EqNat.v b/theories/Arith/EqNat.v index b4a232e20..88cca274b 100755 --- a/theories/Arith/EqNat.v +++ b/theories/Arith/EqNat.v @@ -9,6 +9,7 @@ (*i $Id$ i*) (** Equality on natural numbers *) +Import nat_scope. Fixpoint eq_nat [n:nat] : nat -> Prop := [m:nat]Cases n m of diff --git a/theories/Arith/Euclid.v b/theories/Arith/Euclid.v index 5a2dd1a84..173a90638 100644 --- a/theories/Arith/Euclid.v +++ b/theories/Arith/Euclid.v @@ -11,6 +11,7 @@ Require Mult. Require Compare_dec. Require Wf_nat. +Import nat_scope. Inductive diveucl [a,b:nat] : Set diff --git a/theories/Arith/Even.v b/theories/Arith/Even.v index 074e0a03d..034a60088 100644 --- a/theories/Arith/Even.v +++ b/theories/Arith/Even.v @@ -11,6 +11,7 @@ (** Here we define the predicates [even] and [odd] by mutual induction and we prove the decidability and the exclusion of those predicates. The main results about parity are proved in the module Div2. *) +Import nat_scope. Inductive even : nat->Prop := even_O : (even O) diff --git a/theories/Arith/Gt.v b/theories/Arith/Gt.v index f31154018..6b55c5bd3 100755 --- a/theories/Arith/Gt.v +++ b/theories/Arith/Gt.v @@ -11,6 +11,7 @@ Require Le. Require Lt. Require Plus. +Import nat_scope. Theorem gt_Sn_O : (n:nat)(gt (S n) O). Proof. diff --git a/theories/Arith/Le.v b/theories/Arith/Le.v index 9bdb7dc23..7765187cf 100755 --- a/theories/Arith/Le.v +++ b/theories/Arith/Le.v @@ -9,6 +9,8 @@ (*i $Id$ i*) (** Order on natural numbers *) +Import nat_scope. +Open Scope nat_scope. Theorem le_n_S : (n,m:nat)(le n m)->(le (S n) (S m)). Proof. diff --git a/theories/Arith/Lt.v b/theories/Arith/Lt.v index 711e065da..96541cb9f 100755 --- a/theories/Arith/Lt.v +++ b/theories/Arith/Lt.v @@ -9,6 +9,7 @@ (*i $Id$ i*) Require Le. +Import nat_scope. Theorem lt_n_Sn : (n:nat)(lt n (S n)). Proof. diff --git a/theories/Arith/Max.v b/theories/Arith/Max.v index 43c683a90..8a599ed7e 100755 --- a/theories/Arith/Max.v +++ b/theories/Arith/Max.v @@ -9,6 +9,7 @@ (*i $Id$ i*) Require Arith. +Import nat_scope. (** maximum of two natural numbers *) diff --git a/theories/Arith/Min.v b/theories/Arith/Min.v index 8a5de8703..650b95380 100755 --- a/theories/Arith/Min.v +++ b/theories/Arith/Min.v @@ -9,6 +9,7 @@ (*i $Id$ i*) Require Arith. +Import nat_scope. (** minimum of two natural numbers *) diff --git a/theories/Arith/Minus.v b/theories/Arith/Minus.v index 0f435a560..42f44083a 100755 --- a/theories/Arith/Minus.v +++ b/theories/Arith/Minus.v @@ -12,6 +12,7 @@ Require Lt. Require Le. +Import nat_scope. Fixpoint minus [n:nat] : nat -> nat := [m:nat]Cases n m of diff --git a/theories/Arith/Mult.v b/theories/Arith/Mult.v index ac16d4cb9..1daf8567c 100755 --- a/theories/Arith/Mult.v +++ b/theories/Arith/Mult.v @@ -11,6 +11,7 @@ Require Export Plus. Require Export Minus. Require Export Lt. +Import nat_scope. (** Multiplication *) diff --git a/theories/Arith/Peano_dec.v b/theories/Arith/Peano_dec.v index 694351b67..e998e0b7c 100755 --- a/theories/Arith/Peano_dec.v +++ b/theories/Arith/Peano_dec.v @@ -9,6 +9,7 @@ (*i $Id$ i*) Require Decidable. +Import nat_scope. Theorem O_or_S : (n:nat)({m:nat|(S m)=n})+{O=n}. Proof. diff --git a/theories/Arith/Wf_nat.v b/theories/Arith/Wf_nat.v index 581b45851..b0c715c8b 100755 --- a/theories/Arith/Wf_nat.v +++ b/theories/Arith/Wf_nat.v @@ -12,6 +12,8 @@ Require Lt. +Import nat_scope. + Chapter Well_founded_Nat. Variable A : Set. diff --git a/theories/Init/DatatypesSyntax.v b/theories/Init/DatatypesSyntax.v index 06f28b503..0ba1a5884 100644 --- a/theories/Init/DatatypesSyntax.v +++ b/theories/Init/DatatypesSyntax.v @@ -15,10 +15,12 @@ Require Export Datatypes. Arguments Scope sum [type_scope type_scope]. Arguments Scope prod [type_scope type_scope]. -Infix "+" sum (at level 4, left associativity) : type_scope. -Infix "*" prod (at level 3, right associativity) : type_scope. +Infix LEFTA 4 "+" sum : type_scope. +Notation "x * y" := (prod x y) (at level 3, right associativity) : type_scope + V8only (at level 30, left associativity). -Notation "( x , y )" := (pair ? ? x y) (at level 0). +Notation "( x , y )" := (pair ? ? x y) (at level 0) + V8only "x , y" (at level 150, left associativity). Notation Fst := (fst ? ?). Notation Snd := (snd ? ?). @@ -26,7 +28,8 @@ Arguments Scope option [ type_scope ]. (** Parsing only of things in [Datatypes.v] *) +V7only[ Notation "< A , B > ( x , y )" := (pair A B x y) (at level 1, only parsing, A annot). Notation "< A , B > 'Fst' ( p )" := (fst A B p) (at level 1, only parsing, A annot). Notation "< A , B > 'Snd' ( p )" := (snd A B p) (at level 1, only parsing, A annot). - +]. diff --git a/theories/Init/LogicSyntax.v b/theories/Init/LogicSyntax.v index 2ad45477f..c652a7028 100644 --- a/theories/Init/LogicSyntax.v +++ b/theories/Init/LogicSyntax.v @@ -11,36 +11,49 @@ Require Export Logic. (** Symbolic notations for things in [Logic.v] *) - +V7only[ Notation "< P , Q > { p , q }" := (conj P Q p q) (P annot, at level 1). +]. -Notation "~ x" := (not x) (at level 5, right associativity). -Notation "x = y" := (eq ? x y) (at level 5, no associativity). +Notation "~ x" := (not x) (at level 5, right associativity) + V8only (at level 55, right associativity). +Notation "x = y" := (eq ? x y) (at level 5, no associativity) + V8only (at level 50, no associativity). Infix RIGHTA 6 "/\\" and. Infix RIGHTA 7 "\\/" or. Infix RIGHTA 8 "<->" iff. Notation "'IF' c1 'then' c2 'else' c3" := (IF c1 c2 c3) - (at level 1, c1, c2, c3 at level 8). + (at level 1, c1, c2, c3 at level 8) + V8only (at level 200). (* Order is important to give printing priority to fully typed ALL and EX *) +Notation "'ALL' x : t | p" := (all t [x:t]p) (at level 10, p at level 8) + V8only (at level 200). +Notation "'ALL' x | p" := (all ? [x]p) (at level 10, p at level 8) + V8only (at level 200, p at level 200). Notation All := (all ?). -Notation "'ALL' x | p" := (all ? [x]p) (at level 10, p at level 8). -Notation "'ALL' x : t | p" := (all t [x:t]p) (at level 10, p at level 8). +Notation "'EX' x : t | p" := (ex t [x:t]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). +Notation "'EX' x | p" := (ex ? [x]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). Notation Ex := (ex ?). -Notation "'EX' x | p" := (ex ? [x]p) (at level 10, p at level 8). -Notation "'EX' x : t | p" := (ex t [x:t]p) (at level 10, p at level 8). -Notation Ex2 := (ex2 ?). -Notation "'EX' x | p & q" := (ex2 ? [x]p [x]q) - (at level 10, p, q at level 8). Notation "'EX' x : t | p & q" := (ex2 t [x:t]p [x:t]q) - (at level 10, p, q at level 8). + (at level 10, p, q at level 8) + V8only "'EX2' x : t | p & q" (at level 200, x at level 80). +Notation "'EX' x | p & q" := (ex2 ? [x]p [x]q) + (at level 10, p, q at level 8) + V8only "'EX2' x | p & q" (at level 200, x at level 80). +Notation Ex2 := (ex2 ?). (** Parsing only of things in [Logic.v] *) +V7only[ Notation "< A > 'All' ( P )" := (all A P) (A annot, at level 1, only parsing). -Notation "< A > x = y" := (eq A x y) (A annot, at level 1, x at level 0, only parsing). +Notation "< A > x = y" := (eq A x y) + (A annot, at level 1, x at level 0, only parsing). +]. diff --git a/theories/Init/Logic_TypeSyntax.v b/theories/Init/Logic_TypeSyntax.v index ed915d119..08d4f47c5 100644 --- a/theories/Init/Logic_TypeSyntax.v +++ b/theories/Init/Logic_TypeSyntax.v @@ -12,29 +12,39 @@ Require Logic_Type. (** Symbolic notations for things in [Logic_type.v] *) -Notation "x == y" := (eqT ? x y) (at level 5, no associativity). -Notation "x === y" := (identityT ? x y) (at level 5, no associativity). +Notation "x == y" := (eqT ? x y) (at level 5, no associativity) + V8only (at level 50, no associativity). +Notation "x === y" := (identityT ? x y) (at level 5, no associativity) + V8only (at level 50, no associativity). (* Order is important to give printing priority to fully typed ALL and EX *) +Notation "'ALLT' x : t | p" := (allT t [x:t]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). +Notation "'ALLT' x | p" := (allT ? [x]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). Notation AllT := (allT ?). -Notation "'ALLT' x : t | p" := (allT t [x:t]p) (at level 10, p at level 8). -Notation "'ALLT' x | p" := (allT ? [x]p) (at level 10, p at level 8). +Notation "'EXT' x : t | p" := (exT t [x:t]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). +Notation "'EXT' x | p" := (exT ? [x]p) (at level 10, p at level 8) + V8only (at level 200, x at level 80). Notation ExT := (exT ?). -Notation "'EXT' x : t | p" := (exT t [x:t]p) (at level 10, p at level 8). -Notation "'EXT' x | p" := (exT ? [x]p) (at level 10, p at level 8). -Notation ExT2 := (exT2 ?). Notation "'EXT' x : t | p & q" := (exT2 t [x:t]p [x:t]q) - (at level 10, p, q at level 8). + (at level 10, p, q at level 8) + V8only "'EXT2' x : t | p & q" (at level 200, x at level 80). Notation "'EXT' x | p & q" := (exT2 ? [x]p [x]q) - (at level 10, p, q at level 8). + (at level 10, p, q at level 8) + V8only "'EXT2' x | p & q" (at level 200, x at level 80). +Notation ExT2 := (exT2 ?). (** Parsing only of things in [Logic_type.v] *) +V7only[ Notation "< A > x == y" := (eqT A x y) (A annot, at level 1, x at level 0, only parsing). Notation "< A > x === y" := (identityT A x y) (A annot, at level 1, x at level 0, only parsing). +]. diff --git a/theories/Init/PeanoSyntax.v b/theories/Init/PeanoSyntax.v index e76dbe665..2f53ca9b2 100644 --- a/theories/Init/PeanoSyntax.v +++ b/theories/Init/PeanoSyntax.v @@ -10,32 +10,26 @@ Require Datatypes. Require Peano. +V7only[ Syntax constr level 0: S [ (S $p) ] -> [$p:"nat_printer":9] | O [ O ] -> ["(0)"]. - +]. (* Outside the module to be able to parse the grammar for 0,1,2... !! *) Delimits Scope nat_scope with N. (* For parsing/printing based on scopes *) Module nat_scope. -Infix "+" plus (at level 4) : nat_scope. -Infix "*" mult (at level 3): nat_scope. -Infix "<=" le (at level 5, no associativity) : nat_scope. -Infix "<" lt (at level 5, no associativity) : nat_scope. -Infix ">=" ge (at level 5, no associativity) : nat_scope. -Infix ">" gt (at level 5, no associativity) : nat_scope. - (* Warning: this hides sum and prod and breaks sumor symbolic notation *) Open Scope nat_scope. -(* -Syntax constr - level 0: - S' [ (S $p) ] -> [$p:"nat_printer_S":9] -| O' [ O ] -> [ _:"nat_printer_O" ] -. -*) +Infix 4 "+" plus : nat_scope V8only 40. +Infix 3 "*" mult : nat_scope V8only 30. +Infix NONA 5 "<=" le : nat_scope V8only 50. +Infix NONA 5 "<" lt : nat_scope V8only 50. +Infix NONA 5 ">=" ge : nat_scope V8only 50. +Infix NONA 5 ">" gt : nat_scope V8only 50. + End nat_scope. diff --git a/theories/Init/SpecifSyntax.v b/theories/Init/SpecifSyntax.v index 9d868cb54..b8bb1dec6 100644 --- a/theories/Init/SpecifSyntax.v +++ b/theories/Init/SpecifSyntax.v @@ -11,12 +11,34 @@ Require DatatypesSyntax. Require Specif. +(** Extra factorization of parsing rules *) + +(* Factorizing "sumor" at level 4 to parse B+{x:A|P} without parentheses *) + +Notation "B + { x : A | P }" := B + (sig A [x:A]P) + (at level 4, left associativity, only parsing) + V8only (at level 40, x at level 80, left associativity, only parsing). + +Notation "B + { x : A | P & Q }" := B + (sig2 A [x:A]P [x:A]Q) + (at level 4, left associativity, only parsing) + V8only (at level 40, x at level 80, left associativity, only parsing). + +Notation "B + { x : A & P }" := B + (sigS A [x:A]P) + (at level 4, left associativity, only parsing) + V8only (at level 40, x at level 80, left associativity, only parsing). + +Notation "B + { x : A & P & Q }" := B + (sigS2 A [x:A]P [x:A]Q) + (at level 4, left associativity, only parsing) + V8only (at level 40, x at level 80, left associativity, only parsing). + (** Symbolic notations for things in [Specif.v] *) (* At level 1 to factor with {x:A|P} etc *) -Notation "{ A } + { B }" := (sumbool A B) (at level 1). +Notation "{ A } + { B }" := (sumbool A B) (at level 1) + V8only (at level 10, A at level 80). -Notation "A + { B }" := (sumor A B) (at level 4, left associativity). +Notation "A + { B }" := (sumor A B) (at level 4, left associativity) + V8only (at level 40, B at level 80, left associativity). Notation ProjS1 := (projS1 ? ?). Notation ProjS2 := (projS2 ? ?). @@ -27,28 +49,15 @@ Notation Value := (value ?). Arguments Scope sig [type_scope type_scope]. Arguments Scope sig2 [type_scope type_scope type_scope]. -Notation "{ x : A | P }" := (sig A [x:A]P) (at level 1). -Notation "{ x : A | P & Q }" := (sig2 A [x:A]P [x:A]Q) (at level 1). +Notation "{ x : A | P }" := (sig A [x:A]P) (at level 1) + V8only (at level 10, x at level 80). +Notation "{ x : A | P & Q }" := (sig2 A [x:A]P [x:A]Q) (at level 1) + V8only (at level 10, x at level 80). Arguments Scope sigS [type_scope type_scope]. Arguments Scope sigS2 [type_scope type_scope type_scope]. -Notation "{ x : A & P }" := (sigS A [x:A]P) (at level 1). -Notation "{ x : A & P & Q }" := (sigS2 A [x:A]P [x:A]Q) (at level 1). - -(** Extra factorization of parsing rules *) - -(* Factorizing "sumor" at level 4 to parse B+{x:A|P} without parentheses *) - -Notation "B + { x : A | P }" := B + ({x:A | P}) - (at level 4, left associativity, only parsing). - -Notation "B + { x : A | P & Q }" := B + ({x:A | P & Q}) - (at level 4, left associativity, only parsing). - -Notation "B + { x : A & P }" := B + ({x:A & P}) - (at level 4, left associativity, only parsing). - -Notation "B + { x : A & P & Q }" := B + ({x:A & P & Q}) - (at level 4, left associativity, only parsing). - +Notation "{ x : A & P }" := (sigS A [x:A]P) (at level 1) + V8only (at level 10, x at level 80). +Notation "{ x : A & P & Q }" := (sigS2 A [x:A]P [x:A]Q) (at level 1) + V8only (at level 10, x at level 80). diff --git a/theories/Init/Wf.v b/theories/Init/Wf.v index 745af5ef6..4fb640827 100755 --- a/theories/Init/Wf.v +++ b/theories/Init/Wf.v @@ -107,7 +107,7 @@ Apply F_ext; Auto. Qed. -Lemma fix_eq : (x:A)(fix x)=(F x [y:A][p:(R y x)](fix y)). +Lemma Fix_eq : (x:A)(fix x)=(F x [y:A][p:(R y x)](fix y)). Intro; Unfold fix. Case (Fix_F_eq x). Apply F_ext; Intros. diff --git a/theories/Lists/PolyList.v b/theories/Lists/PolyList.v index cd56ba1de..5baa76891 100644 --- a/theories/Lists/PolyList.v +++ b/theories/Lists/PolyList.v @@ -38,7 +38,8 @@ Fixpoint app [l:list] : list -> list | (cons a l1) => (cons a (app l1 m)) end. -Infix RIGHTA 7 "^" app. +Infix RIGHTA 7 "^" app + V8only 30. Lemma app_nil_end : (l:list)l=(l^nil). Proof. diff --git a/theories/Lists/PolyListSyntax.v b/theories/Lists/PolyListSyntax.v index dd7aba8c9..b74da123f 100644 --- a/theories/Lists/PolyListSyntax.v +++ b/theories/Lists/PolyListSyntax.v @@ -12,4 +12,5 @@ Require PolyList. -Infix RIGHTA 7 "^" app. +Infix RIGHTA 7 "^" app + V8only 30. diff --git a/theories/Reals/Alembert.v b/theories/Reals/Alembert.v index c87053772..e86a4d003 100644 --- a/theories/Reals/Alembert.v +++ b/theories/Reals/Alembert.v @@ -14,6 +14,7 @@ Require Rseries. Require SeqProp. Require PartSum. Require Max. +Import R_scope. (***************************************************) (* Various versions of the criterion of D'Alembert *) @@ -544,4 +545,4 @@ Unfold Rdiv; Apply Rmult_lt_pos. Assumption. Apply Rlt_Rinv; Apply Rabsolu_pos_lt. Red; Intro H7; Rewrite H7 in r; Elim (Rlt_antirefl ? r). -Qed. \ No newline at end of file +Qed. diff --git a/theories/Reals/AltSeries.v b/theories/Reals/AltSeries.v index d9477c50d..5c7f02d4d 100644 --- a/theories/Reals/AltSeries.v +++ b/theories/Reals/AltSeries.v @@ -14,6 +14,7 @@ Require Rseries. Require SeqProp. Require PartSum. Require Max. +Import R_scope. (**********) Definition tg_alt [Un:nat->R] : nat->R := [i:nat]``(pow (-1) i)*(Un i)``. diff --git a/theories/Reals/ArithProp.v b/theories/Reals/ArithProp.v index 3436a2558..18e833e69 100644 --- a/theories/Reals/ArithProp.v +++ b/theories/Reals/ArithProp.v @@ -12,6 +12,7 @@ Require Rbase. Require Rbasic_fun. Require Even. Require Div2. +Import R_scope. Lemma minus_neq_O : (n,i:nat) (lt i n) -> ~(minus n i)=O. Intros; Red; Intro. @@ -128,4 +129,4 @@ Replace (plus (S n) O) with (S n); [Apply le_n_Sn | Ring]. Replace (plus (S n) (S i)) with (S (plus (S n) i)). Apply le_S; Assumption. Apply INR_eq; Rewrite S_INR; Do 2 Rewrite plus_INR; Do 2 Rewrite S_INR; Ring. -Qed. \ No newline at end of file +Qed. diff --git a/theories/Reals/Binomial.v b/theories/Reals/Binomial.v index 2654a5f60..c97f98474 100644 --- a/theories/Reals/Binomial.v +++ b/theories/Reals/Binomial.v @@ -11,6 +11,7 @@ Require Rbase. Require Rfunctions. Require PartSum. +Import R_scope. Definition C [n,p:nat] : R := ``(INR (fact n))/((INR (fact p))*(INR (fact (minus n p))))``. diff --git a/theories/Reals/Cauchy_prod.v b/theories/Reals/Cauchy_prod.v index cf962a0e2..c6926ec18 100644 --- a/theories/Reals/Cauchy_prod.v +++ b/theories/Reals/Cauchy_prod.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require Rseries. Require PartSum. +Import R_scope. (**********) Lemma sum_N_predN : (An:nat->R;N:nat) (lt O N) -> (sum_f_R0 An N)==``(sum_f_R0 An (pred N)) + (An N)``. @@ -342,4 +343,4 @@ Simpl; Symmetry; Apply S_pred with O; Assumption. Inversion H. Left; Reflexivity. Right; Apply lt_le_trans with (1); [Apply lt_n_Sn | Exact H1]. -Qed. \ No newline at end of file +Qed. diff --git a/theories/Reals/Cos_plus.v b/theories/Reals/Cos_plus.v index 9a4e04283..e366adb38 100644 --- a/theories/Reals/Cos_plus.v +++ b/theories/Reals/Cos_plus.v @@ -14,6 +14,8 @@ Require SeqSeries. Require Rtrigo_def. Require Cos_rel. Require Max. +Import nat_scope. +Import R_scope. Definition Majxy [x,y:R] : nat->R := [n:nat](Rdiv (pow (Rmax R1 (Rmax (Rabsolu x) (Rabsolu y))) (mult (4) (S n))) (INR (fact n))). diff --git a/theories/Reals/Cos_rel.v b/theories/Reals/Cos_rel.v index f1275c3db..fb41d9927 100644 --- a/theories/Reals/Cos_rel.v +++ b/theories/Reals/Cos_rel.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require SeqSeries. Require Rtrigo_def. +Import R_scope. Definition A1 [x:R] : nat->R := [N:nat](sum_f_R0 [k:nat]``(pow (-1) k)/(INR (fact (mult (S (S O)) k)))*(pow x (mult (S (S O)) k))`` N). diff --git a/theories/Reals/DiscrR.v b/theories/Reals/DiscrR.v index 94307dc85..791954d06 100644 --- a/theories/Reals/DiscrR.v +++ b/theories/Reals/DiscrR.v @@ -10,6 +10,7 @@ Require RIneq. Require Omega. +Import R_scope. Lemma Rlt_R0_R2 : ``0<2``. Replace ``2`` with (INR (2)); [Apply lt_INR_0; Apply lt_O_Sn | Reflexivity]. diff --git a/theories/Reals/Exp_prop.v b/theories/Reals/Exp_prop.v index 61df2b49f..8a16bfdf1 100644 --- a/theories/Reals/Exp_prop.v +++ b/theories/Reals/Exp_prop.v @@ -17,6 +17,7 @@ Require PSeries_reg. Require Div2. Require Even. Require Max. +Import R_scope. Definition E1 [x:R] : nat->R := [N:nat](sum_f_R0 [k:nat]``/(INR (fact k))*(pow x k)`` N). diff --git a/theories/Reals/MVT.v b/theories/Reals/MVT.v index b271d408b..520ccc788 100644 --- a/theories/Reals/MVT.v +++ b/theories/Reals/MVT.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require Ranalysis1. Require Rtopology. +Import R_scope. (* The Mean Value Theorem *) Theorem MVT : (f,g:R->R;a,b:R;pr1:(c:R)``a(derivable_pt f c);pr2:(c:R)``a(derivable_pt g c)) ``a ((c:R)``a<=c<=b``->(continuity_pt f c)) -> ((c:R)``a<=c<=b``->(continuity_pt g c)) -> (EXT c : R | (EXT P : ``a Prop := [y:R]``(Rabsolu (y-x))R;N:nat) ((n:nat)``(le n N)``->``0<(An n)``) -> ``0 < (sum_f_R0 An N)``. Intros; Induction N. diff --git a/theories/Reals/RIneq.v b/theories/Reals/RIneq.v index 361bf9b22..c07302f3f 100644 --- a/theories/Reals/RIneq.v +++ b/theories/Reals/RIneq.v @@ -16,6 +16,9 @@ Require Export Raxioms. Require Export ZArithRing. Require Omega. Require Export Field. +Import nat_scope. +Import Z_scope. +Import R_scope. (***************************************************************************) (** Instantiating Ring tactic on reals *) @@ -397,6 +400,8 @@ Qed. (***********) Definition Rsqr:R->R:=[r:R]``r*r``. +Notation "x ˛" := (Rsqr x) (at level 2,left associativity) + V8only (at level 20, left associativity). (***********) Lemma Rsqr_O:(Rsqr ``0``)==``0``. diff --git a/theories/Reals/RList.v b/theories/Reals/RList.v index 82c920621..ce022d7c1 100644 --- a/theories/Reals/RList.v +++ b/theories/Reals/RList.v @@ -10,6 +10,7 @@ Require Rbase. Require Rfunctions. +Import R_scope. Inductive Rlist : Type := | nil : Rlist diff --git a/theories/Reals/R_sqr.v b/theories/Reals/R_sqr.v index 933149427..9e8b81d3e 100644 --- a/theories/Reals/R_sqr.v +++ b/theories/Reals/R_sqr.v @@ -10,6 +10,7 @@ Require Rbase. Require Rbasic_fun. +Import R_scope. (****************************************************) (* Rsqr : some results *) diff --git a/theories/Reals/R_sqrt.v b/theories/Reals/R_sqrt.v index 35890cde8..01f0662a9 100644 --- a/theories/Reals/R_sqrt.v +++ b/theories/Reals/R_sqrt.v @@ -11,6 +11,7 @@ Require Rbase. Require Rfunctions. Require Rsqrt_def. +Import R_scope. (* Here is a continuous extension of Rsqrt on R *) Definition sqrt : R->R := [x:R](Cases (case_Rabsolu x) of diff --git a/theories/Reals/Ranalysis.v b/theories/Reals/Ranalysis.v index 2575b8508..40d484dcd 100644 --- a/theories/Reals/Ranalysis.v +++ b/theories/Reals/Ranalysis.v @@ -28,6 +28,7 @@ Require Export RList. Require Export Sqrt_reg. Require Export Ranalysis4. Require Export Rpower. +Import R_scope. Axiom AppVar : R. @@ -473,4 +474,4 @@ Let aux = (RewTerm trm) In IntroHypL aux ?2; Try (Change (continuity_pt aux ?2); | [|-(eqT ? (derive_pt ?1 ?2 ?3) ?4)] -> Let trm = Eval Cbv Beta in (?1 AppVar) In Let aux = (RewTerm trm) In -IntroHypL aux ?2; Let aux2 = (ConsProof aux ?2) In Try (Replace (derive_pt ?1 ?2 ?3) with (derive_pt aux ?2 aux2); [SimplifyDerive aux ?2; Try Unfold plus_fct minus_fct mult_fct div_fct id fct_cte inv_fct opp_fct; Try Ring | Try Apply pr_nu]) Orelse IsDiff_pt. \ No newline at end of file +IntroHypL aux ?2; Let aux2 = (ConsProof aux ?2) In Try (Replace (derive_pt ?1 ?2 ?3) with (derive_pt aux ?2 aux2); [SimplifyDerive aux ?2; Try Unfold plus_fct minus_fct mult_fct div_fct id fct_cte inv_fct opp_fct; Try Ring | Try Apply pr_nu]) Orelse IsDiff_pt. diff --git a/theories/Reals/Ranalysis1.v b/theories/Reals/Ranalysis1.v index 0c1515d9f..b4c2c8cce 100644 --- a/theories/Reals/Ranalysis1.v +++ b/theories/Reals/Ranalysis1.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require Export Rlimit. Require Export Rderiv. +Import R_scope. (****************************************************) (** Basic operations on functions *) @@ -1020,4 +1021,4 @@ Unfold Rdiv; Rewrite <- Ropp_mul1; Apply Rmult_lt_pos. Apply Rlt_anti_compatibility with l. Unfold Rminus; Rewrite Rplus_Ropp_r; Rewrite Rplus_Or; Assumption. Apply Rlt_Rinv; Sup0. -Qed. \ No newline at end of file +Qed. diff --git a/theories/Reals/Ranalysis2.v b/theories/Reals/Ranalysis2.v index 83f9409cd..e2bc87ffe 100644 --- a/theories/Reals/Ranalysis2.v +++ b/theories/Reals/Ranalysis2.v @@ -11,6 +11,7 @@ Require Rbase. Require Rfunctions. Require Ranalysis1. +Import R_scope. (**********) Lemma formule : (x,h,l1,l2:R;f1,f2:R->R) ``h<>0`` -> ``(f2 x)<>0`` -> ``(f2 (x+h))<>0`` -> ``((f1 (x+h))/(f2 (x+h))-(f1 x)/(f2 x))/h-(l1*(f2 x)-l2*(f1 x))/(Rsqr (f2 x))`` == ``/(f2 (x+h))*(((f1 (x+h))-(f1 x))/h-l1) + l1/((f2 x)*(f2 (x+h)))*((f2 x)-(f2 (x+h))) - (f1 x)/((f2 x)*(f2 (x+h)))*(((f2 (x+h))-(f2 x))/h-l2) + (l2*(f1 x))/((Rsqr (f2 x))*(f2 (x+h)))*((f2 (x+h))-(f2 x))``. diff --git a/theories/Reals/Ranalysis3.v b/theories/Reals/Ranalysis3.v index 9cc34059b..0ccda909d 100644 --- a/theories/Reals/Ranalysis3.v +++ b/theories/Reals/Ranalysis3.v @@ -12,7 +12,7 @@ Require Rbase. Require Rfunctions. Require Ranalysis1. Require Ranalysis2. - +Import R_scope. (* Division *) Theorem derivable_pt_lim_div : (f1,f2:R->R;x,l1,l2:R) (derivable_pt_lim f1 x l1) -> (derivable_pt_lim f2 x l2) -> ~``(f2 x)==0``-> (derivable_pt_lim (div_fct f1 f2) x ``(l1*(f2 x)-l2*(f1 x))/(Rsqr (f2 x))``). diff --git a/theories/Reals/Ranalysis4.v b/theories/Reals/Ranalysis4.v index 5c81a076e..5dc3991e2 100644 --- a/theories/Reals/Ranalysis4.v +++ b/theories/Reals/Ranalysis4.v @@ -15,6 +15,7 @@ Require Rtrigo. Require Ranalysis1. Require Ranalysis3. Require Exp_prop. +Import R_scope. (**********) Lemma derivable_pt_inv : (f:R->R;x:R) ``(f x)<>0`` -> (derivable_pt f x) -> (derivable_pt (inv_fct f) x). diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v index cda4273a5..455511633 100644 --- a/theories/Reals/Raxioms.v +++ b/theories/Reals/Raxioms.v @@ -15,6 +15,7 @@ Require Export ZArith_base. Require Export Rsyntax. Require Export TypeSyntax. +Import R_scope. (*********************************************************) (* Field axioms *) @@ -111,6 +112,8 @@ Fixpoint INR [n:nat]:R:=(Cases n of |(S O) => ``1`` |(S n) => ``(INR n)+1`` end). +Arguments Scope INR [nat_scope]. + (**********************************************************) (** Injection from [Z] to [R] *) @@ -122,6 +125,7 @@ Definition IZR:Z->R:=[z:Z](Cases z of |(POS n) => (INR (convert n)) |(NEG n) => ``-(INR (convert n))`` end). +Arguments Scope IZR [Z_scope]. (**********************************************************) (** [R] Archimedian *) diff --git a/theories/Reals/Rbasic_fun.v b/theories/Reals/Rbasic_fun.v index 0fb879095..5de4bd1fd 100644 --- a/theories/Reals/Rbasic_fun.v +++ b/theories/Reals/Rbasic_fun.v @@ -16,6 +16,7 @@ Require Rbase. Require R_Ifp. Require Fourier. +Import R_scope. (*******************************) (** Rmin *) diff --git a/theories/Reals/Rcomplete.v b/theories/Reals/Rcomplete.v index 0fff7c42a..151f2b263 100644 --- a/theories/Reals/Rcomplete.v +++ b/theories/Reals/Rcomplete.v @@ -13,6 +13,7 @@ Require Rfunctions. Require Rseries. Require SeqProp. Require Max. +Import R_scope. (****************************************************) (* R is complete : *) diff --git a/theories/Reals/Rderiv.v b/theories/Reals/Rderiv.v index 234fb3760..b5c0c6b1c 100644 --- a/theories/Reals/Rderiv.v +++ b/theories/Reals/Rderiv.v @@ -20,6 +20,7 @@ Require Fourier. Require Classical_Prop. Require Classical_Pred_Type. Require Omega. +Import R_scope. (*********) Definition D_x:(R->Prop)->R->R->Prop:=[D:R->Prop][y:R][x:R] diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v index 64d5a613b..e4cfa330a 100644 --- a/theories/Reals/Rfunctions.v +++ b/theories/Reals/Rfunctions.v @@ -25,7 +25,7 @@ Require Export SplitRmult. Require Export ArithProp. Require Omega. Require Zpower. - +Import R_scope. (*******************************) (** Factorial *) diff --git a/theories/Reals/Rgeom.v b/theories/Reals/Rgeom.v index 9588df9d3..b36b9ad35 100644 --- a/theories/Reals/Rgeom.v +++ b/theories/Reals/Rgeom.v @@ -13,7 +13,7 @@ Require Rfunctions. Require SeqSeries. Require Rtrigo. Require R_sqrt. - +Import R_scope. Definition dist_euc [x0,y0,x1,y1:R] : R := ``(sqrt ((Rsqr (x0-x1))+(Rsqr (y0-y1))))``. diff --git a/theories/Reals/RiemannInt.v b/theories/Reals/RiemannInt.v index aa43801a9..b10b1407e 100644 --- a/theories/Reals/RiemannInt.v +++ b/theories/Reals/RiemannInt.v @@ -16,6 +16,7 @@ Require RiemannInt_SF. Require Classical_Prop. Require Classical_Pred_Type. Require Max. +Import R_scope. Implicit Arguments On. @@ -27,7 +28,7 @@ Definition Riemann_integrable [f:R->R;a,b:R] : Type := (eps:posreal) (SigT ? [ph Definition phi_sequence [un:nat->posreal;f:R->R;a,b:R;pr:(Riemann_integrable f a b)] := [n:nat](projT1 ? ? (pr (un n))). -Lemma phi_sequence_prop : (un:nat->posreal;f:R->R;a,b:R;pr:(Riemann_integrable f a b);N:nat) (SigT ? [psi:(StepFun a b)]((t:R)``(Rmin a b)<=t<=(Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence un f a b pr N) t)))<=(psi t)``)/\``(Rabsolu (RiemannInt_SF psi))<(un N)``). +Lemma phi_sequence_prop : (un:nat->posreal;f:R->R;a,b:R;pr:(Riemann_integrable f a b);N:nat) (SigT ? [psi:(StepFun a b)]((t:R)``(Rmin a b)<=t<=(Rmax a b)``->``(Rabsolu ((f t)-[(phi_sequence un pr N t)]))<=(psi t)``)/\``(Rabsolu (RiemannInt_SF psi))<(un N)``). Intros; Apply (projT2 ? ? (pr (un N))). Qed. @@ -94,17 +95,17 @@ Qed. Lemma RiemannInt_P4 : (f:R->R;a,b,l:R;pr1,pr2:(Riemann_integrable f a b);un,vn:nat->posreal) (Un_cv un R0) -> (Un_cv vn R0) -> (Un_cv [N:nat](RiemannInt_SF (phi_sequence un pr1 N)) l) -> (Un_cv [N:nat](RiemannInt_SF (phi_sequence vn pr2 N)) l). Unfold Un_cv; Unfold R_dist; Intros; Assert H3 : ``0R;a,b:R;pr1:(Riemann_integrable f a b);pr2:(Riemann_integrable f b a)) ``(RiemannInt pr1)==-(RiemannInt pr2)``. Intros; EApply UL_sequence. Unfold RiemannInt; Case (RiemannInt_exists pr1 5!RinvN RinvN_cv); Intros; Apply u. -Unfold RiemannInt; Case (RiemannInt_exists pr2 5!RinvN RinvN_cv); Intros; Cut (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence RinvN f a b pr1 n) t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). -Cut (EXT psi2:nat->(StepFun b a) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence RinvN f b a pr2 n) t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). +Unfold RiemannInt; Case (RiemannInt_exists pr2 5!RinvN RinvN_cv); Intros; Cut (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr1 n)] t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). +Cut (EXT psi2:nat->(StepFun b a) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr2 n)] t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). Intros; Elim H; Clear H; Intros psi2 H; Elim H0; Clear H0; Intros psi1 H0; Assert H1 := RinvN_cv; Unfold Un_cv; Intros; Assert H3 : ``0``(RinvN n)``(RinvN n)< eps/(5*(Rabsolu l))``. Intros; Replace (pos (RinvN n)) with ``(Rabsolu ((RinvN n)-0))``; [Unfold RinvN; Apply H5; Assumption | Unfold Rminus; Rewrite Ropp_O; Rewrite Rplus_Or; Apply Rabsolu_right; Left; Apply (cond_pos (RinvN n))]. Clear H5; Assert H5 := H7; Clear H7; Exists N; Intros; Unfold R_dist. -Apply Rle_lt_trans with ``(Rabsolu ((RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))-((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))+l*(RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n)))))+(Rabsolu ((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))-x0))+(Rabsolu l)*(Rabsolu ((RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n))-x))``. -Apply Rle_trans with ``(Rabsolu ((RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))-((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))+l*(RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n)))))+(Rabsolu (((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))-x0)+l*((RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n))-x)))``. -Replace ``(RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))-(x0+l*x)`` with ``(((RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))-((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))+l*(RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n)))))+(((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))-x0)+l*((RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n))-x))``; [Apply Rabsolu_triang | Ring]. -Rewrite Rplus_assoc; Apply Rle_compatibility; Rewrite <- Rabsolu_mult; Replace ``(RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))-x0+l*((RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n))-x)`` with ``((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))-x0)+(l*((RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n))-x))``; [Apply Rabsolu_triang | Ring]. +Apply Rle_lt_trans with ``(Rabsolu ((RiemannInt_SF [(phi_sequence RinvN pr3 n)])-((RiemannInt_SF [(phi_sequence RinvN pr1 n)])+l*(RiemannInt_SF [(phi_sequence RinvN pr2 n)]))))+(Rabsolu ((RiemannInt_SF [(phi_sequence RinvN pr1 n)])-x0))+(Rabsolu l)*(Rabsolu ((RiemannInt_SF [(phi_sequence RinvN pr2 n)])-x))``. +Apply Rle_trans with ``(Rabsolu ((RiemannInt_SF [(phi_sequence RinvN pr3 n)])-((RiemannInt_SF [(phi_sequence RinvN pr1 n)])+l*(RiemannInt_SF [(phi_sequence RinvN pr2 n)]))))+(Rabsolu (((RiemannInt_SF [(phi_sequence RinvN pr1 n)])-x0)+l*((RiemannInt_SF [(phi_sequence RinvN pr2 n)])-x)))``. +Replace ``(RiemannInt_SF [(phi_sequence RinvN pr3 n)])-(x0+l*x)`` with ``(((RiemannInt_SF [(phi_sequence RinvN pr3 n)])-((RiemannInt_SF [(phi_sequence RinvN pr1 n)])+l*(RiemannInt_SF [(phi_sequence RinvN pr2 n)]))))+(((RiemannInt_SF [(phi_sequence RinvN pr1 n)])-x0)+l*((RiemannInt_SF [(phi_sequence RinvN pr2 n)])-x))``; [Apply Rabsolu_triang | Ring]. +Rewrite Rplus_assoc; Apply Rle_compatibility; Rewrite <- Rabsolu_mult; Replace ``(RiemannInt_SF [(phi_sequence RinvN pr1 n)])-x0+l*((RiemannInt_SF [(phi_sequence RinvN pr2 n)])-x)`` with ``((RiemannInt_SF [(phi_sequence RinvN pr1 n)])-x0)+(l*((RiemannInt_SF [(phi_sequence RinvN pr2 n)])-x))``; [Apply Rabsolu_triang | Ring]. Replace eps with ``3*eps/5+eps/5+eps/5``. Repeat Apply Rplus_lt. -Assert H7 : (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence RinvN ? ? ? pr1 n) t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). +Assert H7 : (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr1 n)] t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr1 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr1 n0)). -Assert H8 : (EXT psi2:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((g t)-((phi_sequence RinvN ? ? ? pr2 n) t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). +Assert H8 : (EXT psi2:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((g t)-([(phi_sequence RinvN pr2 n)] t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr2 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr2 n0)). -Assert H9 : (EXT psi3:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu (((f t)+l*(g t))-((phi_sequence RinvN ? ? ? pr3 n) t)))<= (psi3 n t)``)/\``(Rabsolu (RiemannInt_SF (psi3 n))) < (RinvN n)``). +Assert H9 : (EXT psi3:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu (((f t)+l*(g t))-([(phi_sequence RinvN pr3 n)] t)))<= (psi3 n t)``)/\``(Rabsolu (RiemannInt_SF (psi3 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr3 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr3 n0)). -Elim H7; Clear H7; Intros psi1 H7; Elim H8; Clear H8; Intros psi2 H8; Elim H9; Clear H9; Intros psi3 H9; Replace ``(RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))-((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))+l*(RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n)))`` with ``(RiemannInt_SF (phi_sequence RinvN ? ? ? pr3 n))+(-1)*((RiemannInt_SF (phi_sequence RinvN ? ? ? pr1 n))+l*(RiemannInt_SF (phi_sequence RinvN ? ? ? pr2 n)))``; [Idtac | Ring]; Do 2 Rewrite <- StepFun_P30; Assert H10 : (Rmin a b)==a. +Elim H7; Clear H7; Intros psi1 H7; Elim H8; Clear H8; Intros psi2 H8; Elim H9; Clear H9; Intros psi3 H9; Replace ``(RiemannInt_SF [(phi_sequence RinvN pr3 n)])-((RiemannInt_SF [(phi_sequence RinvN pr1 n)])+l*(RiemannInt_SF [(phi_sequence RinvN pr2 n)]))`` with ``(RiemannInt_SF [(phi_sequence RinvN pr3 n)])+(-1)*((RiemannInt_SF [(phi_sequence RinvN pr1 n)])+l*(RiemannInt_SF [(phi_sequence RinvN pr2 n)]))``; [Idtac | Ring]; Do 2 Rewrite <- StepFun_P30; Assert H10 : (Rmin a b)==a. Unfold Rmin; Case (total_order_Rle a b); Intro; [Reflexivity | Elim n0; Assumption]. Assert H11 : (Rmax a b)==b. Unfold Rmax; Case (total_order_Rle a b); Intro; [Reflexivity | Elim n0; Assumption]. @@ -629,13 +630,13 @@ Apply StepFun_P34; Assumption. Apply Rle_lt_trans with (RiemannInt_SF (mkStepFun (StepFun_P28 R1 (psi3 n) (mkStepFun (StepFun_P28 (Rabsolu l) (psi1 n) (psi2 n)))))). Apply StepFun_P37; Try Assumption. Intros; Simpl; Rewrite Rmult_1l. -Apply Rle_trans with ``(Rabsolu (((phi_sequence RinvN ? ? ? pr3 n) x1)-((f x1)+l*(g x1))))+(Rabsolu (((f x1)+l*(g x1))+ -1*(((phi_sequence RinvN ? ? ? pr1 n) x1)+l*((phi_sequence RinvN ? ? ? pr2 n) x1))))``. -Replace ``((phi_sequence RinvN ? ? ? pr3 n) x1)+ -1*(((phi_sequence RinvN ? ? ? pr1 n) x1)+l*((phi_sequence RinvN ? ? ? pr2 n) x1))`` with ``(((phi_sequence RinvN ? ? ? pr3 n) x1)-((f x1)+l*(g x1)))+(((f x1)+l*(g x1))+ -1*(((phi_sequence RinvN ? ? ? pr1 n) x1)+l*((phi_sequence RinvN ? ? ? pr2 n) x1)))``; [Apply Rabsolu_triang | Ring]. +Apply Rle_trans with ``(Rabsolu (([(phi_sequence RinvN pr3 n)] x1)-((f x1)+l*(g x1))))+(Rabsolu (((f x1)+l*(g x1))+ -1*(([(phi_sequence RinvN pr1 n)] x1)+l*([(phi_sequence RinvN pr2 n)] x1))))``. +Replace ``([(phi_sequence RinvN pr3 n)] x1)+ -1*(([(phi_sequence RinvN pr1 n)] x1)+l*([(phi_sequence RinvN pr2 n)] x1))`` with ``(([(phi_sequence RinvN pr3 n)] x1)-((f x1)+l*(g x1)))+(((f x1)+l*(g x1))+ -1*(([(phi_sequence RinvN pr1 n)] x1)+l*([(phi_sequence RinvN pr2 n)] x1)))``; [Apply Rabsolu_triang | Ring]. Rewrite Rplus_assoc; Apply Rplus_le. Elim (H9 n); Intros; Rewrite <- Rabsolu_Ropp; Rewrite Ropp_distr2; Apply H13. Elim H12; Intros; Split; Left; Assumption. -Apply Rle_trans with ``(Rabsolu ((f x1)-((phi_sequence RinvN ? ? ? pr1 n) x1)))+(Rabsolu l)*(Rabsolu ((g x1)-((phi_sequence RinvN ? ? ? pr2 n) x1)))``. -Rewrite <- Rabsolu_mult; Replace ``((f x1)+(l*(g x1)+ -1*(((phi_sequence RinvN ? ? ? pr1 n) x1)+l*((phi_sequence RinvN ? ? ? pr2 n) x1))))`` with ``((f x1)-((phi_sequence RinvN ? ? ? pr1 n) x1))+l*((g x1)-((phi_sequence RinvN ? ? ? pr2 n) x1))``; [Apply Rabsolu_triang | Ring]. +Apply Rle_trans with ``(Rabsolu ((f x1)-([(phi_sequence RinvN pr1 n)] x1)))+(Rabsolu l)*(Rabsolu ((g x1)-([(phi_sequence RinvN pr2 n)] x1)))``. +Rewrite <- Rabsolu_mult; Replace ``((f x1)+(l*(g x1)+ -1*(([(phi_sequence RinvN pr1 n)] x1)+l*([(phi_sequence RinvN pr2 n)] x1))))`` with ``((f x1)-([(phi_sequence RinvN pr1 n)] x1))+l*((g x1)-([(phi_sequence RinvN pr2 n)] x1))``; [Apply Rabsolu_triang | Ring]. Apply Rplus_le. Elim (H7 n); Intros; Apply H13. Elim H12; Intros; Split; Left; Assumption. @@ -672,7 +673,7 @@ Qed. Lemma RiemannInt_P15 : (a,b,c:R;pr:(Riemann_integrable (fct_cte c) a b)) ``(RiemannInt pr)==c*(b-a)``. Intros; Unfold RiemannInt; Case (RiemannInt_exists 1!(fct_cte c) 2!a 3!b pr 5!RinvN RinvN_cv); Intros; EApply UL_sequence. Apply u. -Pose phi1 := [N:nat](phi_sequence RinvN 2!(fct_cte c) 3!a 4!b pr N); Change (Un_cv [N:nat](RiemannInt_SF (phi1 N)) ``c*(b-a)``); Pose f := (fct_cte c); Assert H1 : (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence RinvN f a b pr n) t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). +Pose phi1 := [N:nat](phi_sequence RinvN 2!(fct_cte c) 3!a 4!b pr N); Change (Un_cv [N:nat](RiemannInt_SF (phi1 N)) ``c*(b-a)``); Pose f := (fct_cte c); Assert H1 : (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr n)] t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr n)). Elim H1; Clear H1; Intros psi1 H1; Pose phi2 := [n:nat](mkStepFun (StepFun_P4 a b c)); Pose psi2 := [n:nat](mkStepFun (StepFun_P4 a b R0)); Apply RiemannInt_P11 with f RinvN phi2 psi2 psi1; Try Assumption. Apply RinvN_cv. @@ -1113,27 +1114,27 @@ Intros f a b c pr1 pr2 pr3 Hyp1 Hyp2; Unfold RiemannInt; Case (RiemannInt_exists Apply u. Unfold Un_cv; Intros; Assert H0 : ``0(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-((phi_sequence RinvN f a b pr1 n) t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). +Clear x u x0 x1 eps H H0 N1 H1 N2 H2; Assert H1 : (EXT psi1:nat->(StepFun a b) | (n:nat) ((t:R)``(Rmin a b) <= t``/\``t <= (Rmax a b)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr1 n)] t)))<= (psi1 n t)``)/\``(Rabsolu (RiemannInt_SF (psi1 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr1 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr1 n)). -Assert H2 : (EXT psi2:nat->(StepFun b c) | (n:nat) ((t:R)``(Rmin b c) <= t``/\``t <= (Rmax b c)``->``(Rabsolu ((f t)-((phi_sequence RinvN f b c pr2 n) t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). +Assert H2 : (EXT psi2:nat->(StepFun b c) | (n:nat) ((t:R)``(Rmin b c) <= t``/\``t <= (Rmax b c)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr2 n)] t)))<= (psi2 n t)``)/\``(Rabsolu (RiemannInt_SF (psi2 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr2 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr2 n)). -Assert H3 : (EXT psi3:nat->(StepFun a c) | (n:nat) ((t:R)``(Rmin a c) <= t``/\``t <= (Rmax a c)``->``(Rabsolu ((f t)-((phi_sequence RinvN f a c pr3 n) t)))<= (psi3 n t)``)/\``(Rabsolu (RiemannInt_SF (psi3 n))) < (RinvN n)``). +Assert H3 : (EXT psi3:nat->(StepFun a c) | (n:nat) ((t:R)``(Rmin a c) <= t``/\``t <= (Rmax a c)``->``(Rabsolu ((f t)-([(phi_sequence RinvN pr3 n)] t)))<= (psi3 n t)``)/\``(Rabsolu (RiemannInt_SF (psi3 n))) < (RinvN n)``). Split with [n:nat](projT1 ? ? (phi_sequence_prop RinvN pr3 n)); Intro; Apply (projT2 ? ? (phi_sequence_prop RinvN pr3 n)). Elim H1; Clear H1; Intros psi1 H1; Elim H2; Clear H2; Intros psi2 H2; Elim H3; Clear H3; Intros psi3 H3; Assert H := RinvN_cv; Unfold Un_cv; Intros; Assert H4 : ``0 Prop) (x, y : R) (P x) -> (P y) -> (P (Rmin x y)). Intros P x y H1 H2; Unfold Rmin; Case (total_order_Rle x y); Intro; Assumption. diff --git a/theories/Reals/Rprod.v b/theories/Reals/Rprod.v index 198309911..7089ac26c 100644 --- a/theories/Reals/Rprod.v +++ b/theories/Reals/Rprod.v @@ -14,6 +14,7 @@ Require Rfunctions. Require Rseries. Require PartSum. Require Binomial. +Import R_scope. (* TT Ak; 1<=k<=N *) Fixpoint prod_f_SO [An:nat->R;N:nat] : R := Cases N of diff --git a/theories/Reals/Rsigma.v b/theories/Reals/Rsigma.v index 4a5e35fbf..21cf09264 100644 --- a/theories/Reals/Rsigma.v +++ b/theories/Reals/Rsigma.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require Rseries. Require PartSum. +Import R_scope. Set Implicit Arguments. diff --git a/theories/Reals/Rsqrt_def.v b/theories/Reals/Rsqrt_def.v index 8c63c9eb5..6496c9982 100644 --- a/theories/Reals/Rsqrt_def.v +++ b/theories/Reals/Rsqrt_def.v @@ -13,6 +13,7 @@ Require Rbase. Require Rfunctions. Require SeqSeries. Require Ranalysis1. +Import R_scope. Fixpoint Dichotomy_lb [x,y:R;P:R->bool;N:nat] : R := Cases N of diff --git a/theories/Reals/Rsyntax.v b/theories/Reals/Rsyntax.v index 44c2638e8..5d2ec7aa7 100644 --- a/theories/Reals/Rsyntax.v +++ b/theories/Reals/Rsyntax.v @@ -12,6 +12,7 @@ Require Export Rdefinitions. Axiom NRplus : R->R. Axiom NRmult : R->R. +V7only[ Grammar rnatural ident := nat_id [ prim:var($id) ] -> [$id] @@ -194,43 +195,51 @@ Syntax constr | Rodd_inside [<<(REXPR <<(Rplus R1 $r)>>)>>] -> [ $r:"r_printer_odd" ] | Reven_inside [<<(REXPR <<(Rmult (Rplus R1 R1) $r)>>)>>] -> [ $r:"r_printer_even" ] . - +]. (* For parsing/printing based on scopes *) Module R_scope. Delimits Scope R_scope with R. -Infix "<=" Rle (at level 5, no associativity) : R_scope. -Infix "<" Rlt (at level 5, no associativity) : R_scope. -Infix ">=" Rge (at level 5, no associativity) : R_scope. -Infix ">" Rgt (at level 5, no associativity) : R_scope. -Infix "+" Rplus (at level 4) : R_scope. -Infix "-" Rminus (at level 4) : R_scope. -Infix "*" Rmult (at level 3) : R_scope. -Infix "/" Rdiv (at level 3) : R_scope. -Distfix 0 "- _" Ropp : R_scope. + +(* Warning: this hides sum and prod and breaks sumor symbolic notation *) +Open Scope R_scope. + +Infix "<=" Rle (at level 5, no associativity) : R_scope + V8only (at level 50, no associativity). +Infix "<" Rlt (at level 5, no associativity) : R_scope + V8only (at level 50, no associativity). +Infix ">=" Rge (at level 5, no associativity) : R_scope + V8only (at level 50, no associativity). +Infix ">" Rgt (at level 5, no associativity) : R_scope + V8only (at level 50, no associativity). +Infix "+" Rplus (at level 4) : R_scope + V8only (at level 40, left associativity). +Infix "-" Rminus (at level 4) : R_scope + V8only (at level 40, left associativity). +Infix "*" Rmult (at level 3) : R_scope + V8only (at level 30, left associativity). +Infix "/" Rdiv (at level 3) : R_scope + V8only (at level 30, left associativity). +Notation "- x" := (Ropp x) (at level 0) : R_scope + V8only (at level 40, left associativity). Notation "x == y == z" := (eqT R x y)/\(eqT R y z) - (at level 5, y at level 4): R_scope. + (at level 5, y at level 4, no associtivity): R_scope + V8only (at level 50, y at level 49, no associativity). Notation "x <= y <= z" := (Rle x y)/\(Rle y z) - (at level 5, y at level 4) : R_scope. + (at level 5, y at level 4) : R_scope + V8only (at level 50, y at level 49, no associativity). Notation "x <= y < z" := (Rle x y)/\(Rlt y z) - (at level 5, y at level 4) : R_scope. + (at level 5, y at level 4) : R_scope + V8only (at level 50, y at level 49, no associativity). Notation "x < y < z" := (Rlt x y)/\(Rlt y z) - (at level 5, y at level 4) : R_scope. + (at level 5, y at level 4) : R_scope + V8only (at level 50, y at level 49, no associativity). Notation "x < y <= z" := (Rlt x y)/\(Rle y z) - (at level 5, y at level 4) : R_scope. -Notation "x <> y" := ~(eqT R x y) (at level 5) : R_scope. -Distfix 0 "/ _" Rinv : R_scope. - -(* Warning: this hides sum and prod and breaks sumor symbolic notation *) -Open Scope R_scope. - -(* -Syntax constr - level 0: - Rzero' [ R0 ] -> [ _:"r_printer_R0" ] - | Rone' [ R1 ] -> [ _:"r_printer_R1" ] - | Rconst' [(Rplus R1 $r)] -> [$r:"r_printer_Rplus1"] - | Ropp' [(Ropp $n)] -> [ $n:"r_printer_Ropp" ] -*) + (at level 5, y at level 4) : R_scope + V8only (at level 50, y at level 49, no associativity). +Notation "x <> y" := ~(eqT R x y) (at level 5) : R_scope + V8only (at level 50). +Notation "/ x" := (Rinv x) (at level 0): R_scope + V8only (at level 30, left associativity). End R_scope. diff --git a/theories/Reals/Rtopology.v b/theories/Reals/Rtopology.v index b019bff37..5a170b5e5 100644 --- a/theories/Reals/Rtopology.v +++ b/theories/Reals/Rtopology.v @@ -14,6 +14,7 @@ Require Ranalysis1. Require RList. Require Classical_Prop. Require Classical_Pred_Type. +Import R_scope. Definition included [D1,D2:R->Prop] : Prop := (x:R)(D1 x)->(D2 x). Definition disc [x:R;delta:posreal] : R->Prop := [y:R]``(Rabsolu (y-x))Prop] : Prop := (included D1 D2)/\(included D2 D1). -Infix "=_D" eq_Dom (at level 5, no associativity). +Infix "=_D" eq_Dom (at level 5, no associativity) + V8only (at level 50, no associativity). Lemma open_set_P1 : (D:R->Prop) (open_set D) <-> D =_D (interior D). Intro; Split. diff --git a/theories/Reals/Rtrigo.v b/theories/Reals/Rtrigo.v index 8210b35f1..f265b4a77 100644 --- a/theories/Reals/Rtrigo.v +++ b/theories/Reals/Rtrigo.v @@ -19,6 +19,7 @@ Require Export Cos_plus. Require ZArith_base. Require Zcomplements. Require Classical_Prop. +Import R_scope. (** sin_PI2 is the only remaining axiom **) Axiom sin_PI2 : ``(sin (PI/2))==1``. diff --git a/theories/Reals/Rtrigo_alt.v b/theories/Reals/Rtrigo_alt.v index be9a6f872..d147d68ac 100644 --- a/theories/Reals/Rtrigo_alt.v +++ b/theories/Reals/Rtrigo_alt.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require SeqSeries. Require Rtrigo_def. +Import R_scope. (*****************************************************************) (* Using series definitions of cos and sin *) diff --git a/theories/Reals/Rtrigo_calc.v b/theories/Reals/Rtrigo_calc.v index 4e9244dd3..cfa65824b 100644 --- a/theories/Reals/Rtrigo_calc.v +++ b/theories/Reals/Rtrigo_calc.v @@ -13,6 +13,7 @@ Require Rfunctions. Require SeqSeries. Require Rtrigo. Require R_sqrt. +Import R_scope. Lemma tan_PI : ``(tan PI)==0``. Unfold tan; Rewrite sin_PI; Rewrite cos_PI; Unfold Rdiv; Apply Rmult_Ol. diff --git a/theories/Reals/Rtrigo_def.v b/theories/Reals/Rtrigo_def.v index 554a2059e..cf2ba3dd2 100644 --- a/theories/Reals/Rtrigo_def.v +++ b/theories/Reals/Rtrigo_def.v @@ -13,6 +13,7 @@ Require Rfunctions. Require SeqSeries. Require Rtrigo_fun. Require Max. +Import R_scope. (*****************************) (* Definition of exponential *) diff --git a/theories/Reals/Rtrigo_reg.v b/theories/Reals/Rtrigo_reg.v index ac5c9ed20..aa1c2b9d7 100644 --- a/theories/Reals/Rtrigo_reg.v +++ b/theories/Reals/Rtrigo_reg.v @@ -14,6 +14,7 @@ Require SeqSeries. Require Rtrigo. Require Ranalysis1. Require PSeries_reg. +Import R_scope. Lemma CVN_R_cos : (fn:nat->R->R) (fn == [N:nat][x:R]``(pow (-1) N)/(INR (fact (mult (S (S O)) N)))*(pow x (mult (S (S O)) N))``) -> (CVN_R fn). Unfold CVN_R; Intros. diff --git a/theories/Reals/SeqProp.v b/theories/Reals/SeqProp.v index 6f489f37f..e64d9cc4f 100644 --- a/theories/Reals/SeqProp.v +++ b/theories/Reals/SeqProp.v @@ -13,6 +13,7 @@ Require Rfunctions. Require Rseries. Require Classical. Require Max. +Import R_scope. Definition Un_decreasing [Un:nat->R] : Prop := (n:nat) (Rle (Un (S n)) (Un n)). Definition opp_seq [Un:nat->R] : nat->R := [n:nat]``-(Un n)``. diff --git a/theories/Reals/SeqSeries.v b/theories/Reals/SeqSeries.v index 0d4bfbfad..14d438077 100644 --- a/theories/Reals/SeqSeries.v +++ b/theories/Reals/SeqSeries.v @@ -21,6 +21,7 @@ Require Export Rsigma. Require Export Rprod. Require Export Cauchy_prod. Require Export Alembert. +Import R_scope. (**********) Lemma sum_maj1 : (fn:nat->R->R;An:nat->R;x,l1,l2:R;N:nat) (Un_cv [n:nat](SP fn n x) l1) -> (Un_cv [n:nat](sum_f_R0 An n) l2) -> ((n:nat)``(Rabsolu (fn n x))<=(An n)``) -> ``(Rabsolu (l1-(SP fn N x)))<=l2-(sum_f_R0 An N)``. diff --git a/theories/Reals/Sqrt_reg.v b/theories/Reals/Sqrt_reg.v index 9cf2ac003..ccda09a0b 100644 --- a/theories/Reals/Sqrt_reg.v +++ b/theories/Reals/Sqrt_reg.v @@ -12,6 +12,7 @@ Require Rbase. Require Rfunctions. Require Ranalysis1. Require R_sqrt. +Import R_scope. (**********) Lemma sqrt_var_maj : (h:R) ``(Rabsolu h) <= 1`` -> ``(Rabsolu ((sqrt (1+h))-1))<=(Rabsolu h)``. diff --git a/theories/Wellfounded/Lexicographic_Exponentiation.v b/theories/Wellfounded/Lexicographic_Exponentiation.v index 5c73cdae0..8efa124c3 100644 --- a/theories/Wellfounded/Lexicographic_Exponentiation.v +++ b/theories/Wellfounded/Lexicographic_Exponentiation.v @@ -32,7 +32,10 @@ Notation List := (list A). Notation Nil := (nil A). (* useless but symmetric *) Notation Cons := (cons 1!A). +Notation "<< x , y >>" := (exist List Descl x y) (at level 0) + V8only (at level 0, x,y at level 100). +V7only[ Syntax constr level 1: List [ (list A) ] -> ["List"] @@ -42,12 +45,11 @@ Syntax constr level 10: Cons2 [ (cons A $e $l) ] -> ["Cons " $e:L " " $l:L ]. -Hints Resolve d_one d_nil t_step. - Syntax constr level 1: pair_sig [ (exist (list A) Desc $e $d) ] -> ["<<" $e:L "," $d:L ">>"]. - +]. +Hints Resolve d_one d_nil t_step. Lemma left_prefix : (x,y,z:List)(ltl x^y z)-> (ltl x z). Proof. diff --git a/theories/ZArith/Wf_Z.v b/theories/ZArith/Wf_Z.v index be8820481..cd74d0bad 100644 --- a/theories/ZArith/Wf_Z.v +++ b/theories/ZArith/Wf_Z.v @@ -12,6 +12,7 @@ Require fast_integer. Require zarith_aux. Require auxiliary. Require Zsyntax. +Import Z_scope. (** Our purpose is to write an induction shema for {0,1,2,...} similar to the [nat] schema (Theorem [Natlike_rec]). For that the diff --git a/theories/ZArith/ZArith_dec.v b/theories/ZArith/ZArith_dec.v index 7fbc56ee8..7037a5cac 100644 --- a/theories/ZArith/ZArith_dec.v +++ b/theories/ZArith/ZArith_dec.v @@ -14,6 +14,8 @@ Require fast_integer. Require zarith_aux. Require auxiliary. Require Zsyntax. +Import Z_scope. + Lemma Dcompare_inf : (r:relation) {r=EGAL} + {r=INFERIEUR} + {r=SUPERIEUR}. Proof. diff --git a/theories/ZArith/Zcomplements.v b/theories/ZArith/Zcomplements.v index 78f537098..ccbad9de2 100644 --- a/theories/ZArith/Zcomplements.v +++ b/theories/ZArith/Zcomplements.v @@ -12,6 +12,7 @@ Require ZArith_base. Require ZArithRing. Require Omega. Require Wf_nat. +Import Z_scope. (** Multiplication by a number >0 preserves [Zcompare]. It also perserves [Zle], [Zlt], [Zge], [Zgt] *) diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v index c3ebf133d..aa0991583 100644 --- a/theories/ZArith/Zdiv.v +++ b/theories/ZArith/Zdiv.v @@ -23,6 +23,7 @@ Require Export ZArith_base. Require Omega. Require ZArithRing. Require Zcomplements. +Import Z_scope. (** @@ -267,6 +268,7 @@ Save. (** Syntax *) +V7only[ Grammar znatural expr2 : constr := expr_div [ expr2($p) "/" expr2($c) ] -> [ (Zdiv $p $c) ] | expr_mod [ expr2($p) "%" expr2($c) ] -> [ (Zmod $p $c) ] @@ -283,8 +285,13 @@ Syntax constr -> [ (ZEXPR $n1):E "/" [0 0] (ZEXPR $n2):L ] | Zmod_inside [ << (ZEXPR <<(Zmod $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E "%" [0 0] (ZEXPR $n2):L ] + -> [ (ZEXPR $n1):E " %" [1 0] (ZEXPR $n2):L ] . +]. +Infix 3 "/" Zdiv : Z_scope + V8only 30. +Infix 3 "mod" Zmod : Z_scope + V8only 30. (** Other lemmas (now using the syntax for [Zdiv] and [Zmod]). *) @@ -302,7 +309,7 @@ Absurd `b-a >= 1`. Omega. Rewrite -> H0. Rewrite -> H2. -Assert `c*(b/c)+b%c-(c*(a/c)+a%c) = c*(b/c - a/c) + b%c - a%c`. +Assert `c*(b/c)+b % c-(c*(a/c)+a % c) = c*(b/c - a/c) + b % c - a % c`. Ring. Rewrite H3. Assert `c*(b/c-a/c) >= c*1`. @@ -314,7 +321,7 @@ Ring. Omega. Save. -Lemma Z_mod_plus : (a,b,c:Z)`c > 0`->`(a+b*c)%c = a%c`. +Lemma Z_mod_plus : (a,b,c:Z)`c > 0`->`(a+b*c) % c = a % c`. Proof. Intros a b c cPos. Generalize (Z_div_mod_eq a c cPos). @@ -323,9 +330,9 @@ Generalize (Z_div_mod_eq `a+b*c` c cPos). Generalize (Z_mod_lt `a+b*c` c cPos). Intros. -Assert `(a+b*c)%c - a%c = c*(b+a/c - (a+b*c)/c)`. -Replace `(a+b*c)%c` with `a+b*c - c*((a+b*c)/c)`. -Replace `a%c` with `a - c*(a/c)`. +Assert `(a+b*c) % c - a % c = c*(b+a/c - (a+b*c)/c)`. +Replace `(a+b*c) % c` with `a+b*c - c*((a+b*c)/c)`. +Replace `a % c` with `a - c*(a/c)`. Ring. Omega. Omega. @@ -358,7 +365,7 @@ Generalize (Z_div_mod_eq `a+b*c` c cPos). Generalize (Z_mod_lt `a+b*c` c cPos). Intros. Apply Zmult_reg_left with c. Omega. -Replace `c*((a+b*c)/c)` with `a+b*c-(a+b*c)%c`. +Replace `c*((a+b*c)/c)` with `a+b*c-(a+b*c) % c`. Rewrite (Z_mod_plus a b c cPos). Pattern 1 a; Rewrite H2. Ring. @@ -375,7 +382,7 @@ Pattern 2 a; Rewrite H. Omega. Save. -Lemma Z_mod_same : (a:Z)`a>0`->`a%a=0`. +Lemma Z_mod_same : (a:Z)`a>0`->`a % a=0`. Proof. Intros a aPos. Generalize (Z_mod_plus `0` `1` a aPos). diff --git a/theories/ZArith/Zlogarithm.v b/theories/ZArith/Zlogarithm.v index 660a58ed6..52e2f552e 100644 --- a/theories/ZArith/Zlogarithm.v +++ b/theories/ZArith/Zlogarithm.v @@ -23,6 +23,7 @@ Require ZArith_base. Require Omega. Require Zcomplements. Require Zpower. +Import Z_scope. Section Log_pos. (* Log of positive integers *) diff --git a/theories/ZArith/Zmisc.v b/theories/ZArith/Zmisc.v index ad5db4b53..2ed638696 100644 --- a/theories/ZArith/Zmisc.v +++ b/theories/ZArith/Zmisc.v @@ -13,6 +13,7 @@ Require zarith_aux. Require auxiliary. Require Zsyntax. Require Bool. +Import Z_scope. (** Overview of the sections of this file: - logic: Logic complements. @@ -37,8 +38,10 @@ End logic. Section numbers. -Definition entier_of_Z := [z:Z]Case z of Nul Pos Pos end. -Definition Z_of_entier := [x:entier]Case x of ZERO POS end. +Definition entier_of_Z := + [z:Z]Cases z of ZERO => Nul | (POS p) => (Pos p) | (NEG p) => (Pos p) end. +Definition Z_of_entier := + [x:entier]Cases x of Nul => ZERO | (Pos p) => (POS p) end. (* Coercion Z_of_entier : entier >-> Z. *) @@ -370,7 +373,7 @@ Qed. Theorem Zcompare_elim : (c1,c2,c3:Prop)(x,y:Z) ((x=y) -> c1) ->(`x < y` -> c2) ->(`x > y`-> c3) - -> Case `x ?= y`of c1 c2 c3 end. + -> Cases `x ?= y`of EGAL => c1 | INFERIEUR => c2 | SUPERIEUR => c3 end. Intros. Apply rename with x:=`x ?= y`; Intro r; Elim r; @@ -380,7 +383,7 @@ Apply rename with x:=`x ?= y`; Intro r; Elim r; Qed. Lemma Zcompare_x_x : (x:Z) `x ?= x` = EGAL. -Intro; Apply Case (Zcompare_EGAL x x) of [h1,h2: ?]h2 end. +Intro; Apply let (h1,h2) = (Zcompare_EGAL x x) in h2. Apply refl_equal. Qed. @@ -397,7 +400,8 @@ Discriminate H. Qed. Lemma Zcompare_eq_case : - (c1,c2,c3:Prop)(x,y:Z) c1 -> x=y -> (Case `x ?= y` of c1 c2 c3 end). + (c1,c2,c3:Prop)(x,y:Z) c1 -> x=y -> + Cases `x ?= y` of EGAL => c1 | INFERIEUR => c2 | SUPERIEUR => c3 end. Intros. Rewrite -> (Case (Zcompare_EGAL x y) of [h1,h2: ?]h2 end H0). Assumption. @@ -405,22 +409,26 @@ Qed. (** Four very basic lemmas about [Zle], [Zlt], [Zge], [Zgt] *) Lemma Zle_Zcompare : - (x,y:Z)`x <= y` -> Case `x ?= y` of True True False end. + (x,y:Z)`x <= y` -> + Cases `x ?= y` of EGAL => True | INFERIEUR => True | SUPERIEUR => False end. Intros x y; Unfold Zle; Elim `x ?=y`; Auto with arith. Qed. Lemma Zlt_Zcompare : - (x,y:Z)`x < y` -> Case `x ?= y` of False True False end. + (x,y:Z)`x < y` -> + Cases `x ?= y` of EGAL => False | INFERIEUR => True | SUPERIEUR => False end. Intros x y; Unfold Zlt; Elim `x ?=y`; Intros; Discriminate Orelse Trivial with arith. Qed. Lemma Zge_Zcompare : - (x,y:Z)` x >= y`-> Case `x ?= y` of True False True end. + (x,y:Z)` x >= y`-> + Cases `x ?= y` of EGAL => True | INFERIEUR => False | SUPERIEUR => True end. Intros x y; Unfold Zge; Elim `x ?=y`; Auto with arith. Qed. Lemma Zgt_Zcompare : - (x,y:Z)`x > y` -> Case `x ?= y` of False False True end. + (x,y:Z)`x > y` -> + Cases `x ?= y` of EGAL => False | INFERIEUR => False | SUPERIEUR => True end. Intros x y; Unfold Zgt; Elim `x ?= y`; Intros; Discriminate Orelse Trivial with arith. Qed. diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v index 11ec071c1..8b68b223d 100644 --- a/theories/ZArith/Zpower.v +++ b/theories/ZArith/Zpower.v @@ -11,6 +11,7 @@ Require ZArith_base. Require Omega. Require Zcomplements. +Import Z_scope. Section section1. @@ -312,7 +313,7 @@ Elim (convert p); Simpl; [ Trivial with zarith | Intro n; Rewrite (two_power_nat_S n); Unfold 2 Zdiv_rest_aux; - Elim (iter_nat n (Z*Z)*Z Zdiv_rest_aux ((x,`0`),`1`)); + Elim (iter_nat n 'T:(Z*Z)*Z ' Zdiv_rest_aux ((x,`0`),`1`)); NewDestruct a; Intros; Apply f_equal with f:=[z:Z]`2*z`; Assumption ]. Qed. @@ -374,7 +375,7 @@ Lemma Zdiv_rest_correct : (x:Z)(p:positive)(Zdiv_rest_proofs x p). Intros x p. Generalize (Zdiv_rest_correct1 x p); Generalize (Zdiv_rest_correct2 x p). -Elim (iter_pos p (Z*Z)*Z Zdiv_rest_aux ((x,`0`),`1`)). +Elim (iter_pos p 'T:(Z*Z)*Z ' Zdiv_rest_aux ((x,`0`),`1`)). Induction a. Intros. Elim H; Intros H1 H2; Clear H. diff --git a/theories/ZArith/Zsqrt.v b/theories/ZArith/Zsqrt.v index f9f8288d0..a5ddfff4f 100644 --- a/theories/ZArith/Zsqrt.v +++ b/theories/ZArith/Zsqrt.v @@ -11,6 +11,7 @@ Require Export ZArith_base. Require Export ZArithRing. Require Export Omega. +Import Z_scope. (** The following tactic replaces all instances of (POS (xI ...)) by `2*(POS ...)+1` , but only when ... is not made only with xO, XI, or xH. *) diff --git a/theories/ZArith/Zsyntax.v b/theories/ZArith/Zsyntax.v index 9ac3d03d3..3f1abc8d7 100644 --- a/theories/ZArith/Zsyntax.v +++ b/theories/ZArith/Zsyntax.v @@ -11,6 +11,7 @@ Require Export fast_integer. Require Export zarith_aux. +V7only[ Grammar znatural ident := nat_id [ prim:var($id) ] -> [$id] @@ -114,9 +115,9 @@ Syntax constr | Zlt_Zlt [ (Zlt $n1 $n2)/\(Zlt $n2 $n3) ] -> [[ "`" (ZEXPR $n1) [1 0] "< " (ZEXPR $n2) [1 0] "< " (ZEXPR $n3) "`"]] - | ZZero [ ZERO ] -> [ "`0`" ] - | ZPos [ (POS $r) ] -> [$r:"positive_printer":9] - | ZNeg [ (NEG $r) ] -> [$r:"negative_printer":9] + | ZZero_v7 [ ZERO ] -> [ "`0`" ] + | ZPos_v7 [ (POS $r) ] -> [$r:"positive_printer":9] + | ZNeg_v7 [ (NEG $r) ] -> [$r:"negative_printer":9] ; level 7: @@ -217,6 +218,7 @@ Syntax constr | ZNeg_inside [ << (ZEXPR <<(NEG $p)>>) >>] -> [$p:"negative_printer_inside":9] . +]. (* For parsing/printing based on scopes *) Module Z_scope. @@ -226,44 +228,46 @@ Module Z_scope. Delimits Scope positive_scope with P. Delimits Scope Z_scope with Z. -Infix "+" Zplus (at level 4) : Z_scope. -Infix "-" Zminus (at level 4) : Z_scope. -Infix "*" Zmult (at level 3) : Z_scope. -Distfix 0 "- _" Zopp : Z_scope. -Infix "<=" Zle (at level 5, no associativity) : Z_scope. -Infix "<" Zlt (at level 5, no associativity) : Z_scope. -Infix ">=" Zge (at level 5, no associativity) : Z_scope. -Infix ">" Zgt (at level 5, no associativity) : Z_scope. -Infix "? =" Zcompare (at level 5, no associativity) : Z_scope. +(* Warning: this hides sum and prod and breaks sumor symbolic notation *) +Open Scope Z_scope. + +Arguments Scope POS [positive_scope]. +Arguments Scope NEG [positive_scope]. + +Infix LEFTA 4 "+" Zplus : Z_scope. +Infix LEFTA 4 "-" Zminus : Z_scope. +Infix LEFTA 3 "*" Zmult : Z_scope. +Notation "- x" := (Zopp x) (at level 0) : Z_scope + V8only (at level 40). +Infix NONA 5 "<=" Zle : Z_scope. +Infix NONA 5 "<" Zlt : Z_scope. +Infix NONA 5 ">=" Zge : Z_scope. +Infix NONA 5 ">" Zgt : Z_scope. +Infix NONA 5 "?=" Zcompare : Z_scope. Notation "x <= y <= z" := (Zle x y)/\(Zle y z) - (at level 5, y at level 4):Z_scope. + (at level 5, y at level 4):Z_scope + V8only (at level 50, y at level 49). Notation "x <= y < z" := (Zle x y)/\(Zlt y z) - (at level 5, y at level 4):Z_scope. + (at level 5, y at level 4):Z_scope + V8only (at level 50, y at level 49). Notation "x < y < z" := (Zlt x y)/\(Zlt y z) - (at level 5, y at level 4):Z_scope. + (at level 5, y at level 4):Z_scope + V8only (at level 50, y at level 49). Notation "x < y <= z" := (Zlt x y)/\(Zle y z) - (at level 5, y at level 4):Z_scope. + (at level 5, y at level 4):Z_scope + V8only (at level 50, y at level 49). Notation "x = y = z" := x=y/\y=z - (at level 5, y at level 4):Z_scope. + (at level 5, y at level 4) : Z_scope + V8only (at level 50, y at level 49). -Notation "x <> y" := ~(eq Z x y) (at level 5) : Z_scope. +Notation "x <> y" := ~(eq Z x y) (at level 5, no associativity) : Z_scope + V8only (at level 50, no associativity). (* Notation "| x |" (Zabs x) : Z_scope.(* "|" conflicts with THENS *)*) +V7only[ (* Overwrite the printing of "`x = y`" *) Syntax constr level 0: Zeq [ (eq Z $n1 $n2) ] -> [[ $n1 [1 0] "= " $n2 ]]. - - -(* Warning: this hides sum and prod and breaks sumor symbolic notation *) -Open Scope Z_scope. - -(* -Syntax constr - level 0: - | ZZero' [ ZERO ] -> [dummy:"z_printer_ZERO"] - | ZPos' [ (POS $r) ] -> [$r:"z_printer_POS":9] - | ZNeg' [ (NEG $r) ] -> [$r:"z_printer_NEG":9] -. -*) +]. End Z_scope. diff --git a/theories/ZArith/Zwf.v b/theories/ZArith/Zwf.v index 56baf1782..d93a6ea44 100644 --- a/theories/ZArith/Zwf.v +++ b/theories/ZArith/Zwf.v @@ -11,6 +11,7 @@ Require ZArith_base. Require Export Wf_nat. Require Omega. +Import Z_scope. (** Well-founded relations on Z. *) diff --git a/theories/ZArith/fast_integer.v b/theories/ZArith/fast_integer.v index fe16e8dff..8b73960be 100644 --- a/theories/ZArith/fast_integer.v +++ b/theories/ZArith/fast_integer.v @@ -75,7 +75,7 @@ with add_carry [x,y:positive]:positive := end end. -Infix "+" add (at level 4, left associativity) : positive_scope. +Infix LEFTA 4 "+" add : positive_scope. Open Scope positive_scope. @@ -1020,7 +1020,8 @@ Fixpoint times [x:positive] : positive -> positive:= | xH => y end. -Infix "*" times (at level 3, left associativity) : positive_scope. +Infix LEFTA 3 "*" times : positive_scope + V8only 30. (** Correctness of multiplication on positive *) Theorem times_convert : @@ -1065,7 +1066,8 @@ Definition Zmult := [x,y:Z] end end. -Infix "*" Zmult (at level 3, left associativity) : Z_scope. +Infix LEFTA 3 "*" Zmult : Z_scope + V8only 30. Open Scope Z_scope. diff --git a/toplevel/class.ml b/toplevel/class.ml index 1a252dc13..e1ba8e736 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -335,7 +335,7 @@ let try_add_new_coercion_with_source ref stre ~source = let add_coercion_hook stre ref = try_add_new_coercion ref stre; Options.if_verbose message - (string_of_qualid (shortest_qualid_of_global None ref) + (string_of_qualid (shortest_qualid_of_global Idset.empty ref) ^ " is now a coercion") let add_subclass_hook stre ref = diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index 4ee405318..130271aef 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -166,7 +166,11 @@ let parse_args is_ide = | "-compile-verbose" :: [] -> usage () | "-translate" :: rem -> make_translate true; parse rem - | "-ftranslate" :: rem -> make_translate true; translate_file := true; parse rem + | "-ftranslate" :: rem -> + make_translate true; + translate_file := true; + Constrextern.print_coercions := true; + parse rem | "-unsafe" :: f :: rem -> add_unsafe f; parse rem | "-unsafe" :: [] -> usage () diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 46cadeb0a..b6dac63a9 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -316,7 +316,7 @@ let explain_unsolvable_implicit env = function str "Cannot infer a type of this anonymous abstraction" | ImplicitArg (c,n) -> str "Cannot infer the " ++ pr_ord n ++ - str " implicit argument of " ++ Nametab.pr_global_env None c + str " implicit argument of " ++ Nametab.pr_global_env Idset.empty c | InternalHole -> str "Cannot infer a term for an internal placeholder" | TomatchTypeParameter (tyi,n) -> diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml index fbbae0a1b..36734041a 100644 --- a/toplevel/metasyntax.ml +++ b/toplevel/metasyntax.ml @@ -42,15 +42,15 @@ let constr_parser_with_glob = Pcoq.map_entry constr_to_ast Constr.constr let globalize_ref vars ref = match Constrintern.interp_reference (vars,[]) ref with - | RRef (loc,a) -> Constrextern.extern_reference loc a + | RRef (loc,a) -> Constrextern.extern_reference loc Idset.empty a | RVar (loc,x) -> Ident (loc,x) | _ -> anomaly "globalize_ref: not a reference" let globalize_ref_term vars ref = match Constrintern.interp_reference (vars,[]) ref with - | RRef (loc,a) -> CRef (Constrextern.extern_reference loc a) + | RRef (loc,a) -> CRef (Constrextern.extern_reference loc Idset.empty a) | RVar (loc,x) -> CRef (Ident (loc,x)) - | c -> Constrextern.extern_rawconstr c + | c -> Constrextern.extern_rawconstr Idset.empty c let rec globalize_constr_expr vars = function | CRef ref -> globalize_ref_term vars ref @@ -100,6 +100,7 @@ let (inPPSyntax,outPPSyntax) = * Syntax objects in compiled modules are not re-checked. *) let add_syntax_obj whatfor sel = + if not !Options.v7_only then Lib.add_anonymous_leaf (inPPSyntax (interp_syntax_entry whatfor sel)) @@ -304,7 +305,7 @@ let make_hunks_ast symbols etyps from = in make NoBreak symbols -let add_break n l = UnpCut (PpBrk(n,1)) :: l +let add_break n l = UnpCut (PpBrk(n,0)) :: l let make_hunks etyps symbols = let vars,typs = List.split etyps in @@ -319,25 +320,21 @@ let make_hunks etyps symbols = u :: make CanBreak prods | Terminal s :: prods when List.exists is_non_terminal prods -> - let protect = - is_letter s.[0] || - (is_non_terminal (List.hd prods) && - (is_letter (s.[String.length s -1])) || - (is_digit (s.[String.length s -1]))) in - if is_comma s || is_right_bracket s then - UnpTerminal s :: add_break 0 (make NoBreak prods) - else if (is_operator s || is_left_bracket s) && ws = CanBreak then - add_break (if protect then 1 else 0) - (UnpTerminal (if protect then s^" " else s) :: make CanBreak prods) - else - if protect then - (if ws = CanBreak then add_break 1 else (fun x -> x)) - (UnpTerminal (s^" ") :: make CanBreak prods) - else - UnpTerminal s :: make CanBreak prods + if ws = CanBreak then + if is_comma s || is_right_bracket s then + UnpTerminal s :: add_break 1 (make NoBreak prods) + else if is_operator s then + UnpTerminal (" "^s) :: add_break 1 (make NoBreak prods) + else + add_break 1 (UnpTerminal (s^" ") :: make CanBreak prods) + else + UnpTerminal (s^" ") :: make CanBreak prods | Terminal s :: prods -> - UnpTerminal s :: make NoBreak prods + if ws = CanBreak then + UnpTerminal (" "^s) :: make NoBreak prods + else + UnpTerminal s :: make NoBreak prods | Break n :: prods -> add_break n (make NoBreak prods) @@ -365,17 +362,28 @@ let make_symbolic n symbols etyps = (n,List.map assoc_of_type etyps), (String.concat " " (List.flatten (List.map string_of_symbol symbols))) +let rec define_keywords = function + NonTerm(_,Some(_,(ETConstr _|ETOther("constr","binder_constr")))) as n1 :: + Term("IDENT",k) :: l when not !Options.v7 -> + prerr_endline ("Defining '"^k^"' as keyword"); + Lexer.add_token("",k); + n1 :: Term("",k) :: define_keywords l + | n :: l -> n :: define_keywords l + | [] -> [] + let make_production etyps symbols = - List.fold_right - (fun t l -> match t with - | NonTerminal m -> - let typ = List.assoc m etyps in - NonTerm (ProdPrimitive typ, Some (m,typ)) :: l - | Terminal s -> - Term (Extend.terminal s) :: l - | Break _ -> - l) - symbols [] + let prod = + List.fold_right + (fun t l -> match t with + | NonTerminal m -> + let typ = List.assoc m etyps in + NonTerm (ProdPrimitive typ, Some (m,typ)) :: l + | Terminal s -> + Term (Extend.terminal s) :: l + | Break _ -> + l) + symbols [] in + define_keywords prod let strip s = let n = String.length s in @@ -550,9 +558,11 @@ let recompute_assoc typs = let add_syntax_extension df modifiers = let (assoc,n,etyps,onlyparse) = interp_notation_modifiers modifiers in + let inner = if !Options.v7 then (10,InternalProd) else + (200,InternalProd) in let (typs,symbs) = find_symbols - (n,BorderProd(true,assoc)) (10,InternalProd) (n,BorderProd(false,assoc)) + (n,BorderProd(true,assoc)) inner (n,BorderProd(false,assoc)) [] (split df) in let typs = List.map (set_entry_type etyps) typs in let assoc = recompute_assoc typs in @@ -570,6 +580,7 @@ let load_notation _ (_,(_,prec,ntn,scope,pat,onlyparse,_)) = Symbols.declare_scope scope let open_notation i (_,(oldse,prec,ntn,scope,pat,onlyparse,df)) = +(*print_string ("Open notation "^ntn^" at "^string_of_int (fst prec)^"\n");*) if i=1 then begin let b = Symbols.exists_notation_in_scope scope prec ntn pat in (* Declare the old printer rule and its interpretation *) @@ -622,11 +633,14 @@ let make_old_pp_rule n symbols typs r ntn scope vars = let rule_name = ntn^"_"^scope^"_notation" in make_syntax_rule n rule_name symbols typs ast ntn scope -let add_notation_in_scope df c (assoc,n,etyps,onlyparse) sc toks = +let add_notation_in_scope df c (assoc,n,etyps,onlyparse) omodv8 sc toks = + let onlyparse = onlyparse or !Options.v7_only in let scope = match sc with None -> Symbols.default_scope | Some sc -> sc in + let inner = + if !Options.v7 then (10,InternalProd) else (200,InternalProd) in let (typs,symbols) = find_symbols - (n,BorderProd(true,assoc)) (10,InternalProd) (n,BorderProd(false,assoc)) + (n,BorderProd(true,assoc)) inner (n,BorderProd(false,assoc)) [] toks in let vars = List.map fst typs in (* To globalize... *) @@ -635,21 +649,36 @@ let add_notation_in_scope df c (assoc,n,etyps,onlyparse) sc toks = let assoc = recompute_assoc typs in (* Declare the parsing and printing rules if not already done *) let (prec,notation) = make_symbolic n symbols typs in + let (ppprec,ppn,pptyps,ppsymbols) = + match omodv8 with + Some(toks8,(a8,n8,typs8,_)) when Options.do_translate() -> + let (typs,symbols) = + find_symbols + (n8,BorderProd(true,a8)) (200,InternalProd) + (n8,BorderProd(false,a8)) + [] toks8 in + let typs = List.map (set_entry_type typs8) typs in + let (prec,notation) = make_symbolic n8 symbols typs in + (prec, n8, typs, symbols) + | _ -> (prec, n, typs, symbols) in let gram_rule = make_grammar_rule n assoc typs symbols notation in - let pp_rule = if onlyparse then None else Some (make_pp_rule typs symbols) in - Lib.add_anonymous_leaf (inSyntaxExtension(prec,notation,gram_rule,pp_rule)); + let pp_rule = + if onlyparse then None + else Some (make_pp_rule pptyps ppsymbols) in + Lib.add_anonymous_leaf + (inSyntaxExtension(ppprec,notation,gram_rule,pp_rule)); let old_pp_rule = if onlyparse then None - else - let r = - interp_rawconstr_gen false Evd.empty (Global.env()) [] false (vars,[]) c in - Some (make_old_pp_rule n symbols typs r notation scope vars) in + else + let r = interp_rawconstr_gen + false Evd.empty (Global.env()) [] false (vars,[]) c in + Some (make_old_pp_rule ppn ppsymbols pptyps r notation scope vars) in (* Declare the interpretation *) let vars = List.map (fun id -> id,[] (* insert the right scope *)) vars in Lib.add_anonymous_leaf - (inNotation(old_pp_rule,prec,notation,scope,a,onlyparse,df)) + (inNotation(old_pp_rule,ppprec,notation,scope,a,onlyparse,df)) -let add_notation df a modifiers sc = +let add_notation df a modifiers mv8 sc = let toks = split df in match toks with | [String x] when quote(strip x) = x & modifiers = [] -> @@ -658,7 +687,13 @@ let add_notation df a modifiers sc = let c = snd (interp_aconstr [] a) in Syntax_def.declare_syntactic_definition ident c | _ -> - add_notation_in_scope df a (interp_notation_modifiers modifiers) sc toks + add_notation_in_scope + df a (interp_notation_modifiers modifiers) + (option_app (fun (s8,ml8) -> + let toks8 = split s8 in + let im8 = interp_notation_modifiers ml8 in + (toks8,im8)) mv8) + sc toks (* TODO add boxes information in the expression *) @@ -682,12 +717,12 @@ let add_distfix assoc n df r sc = let df = String.concat " " l in let a = mkAppC (mkRefC r, vars) in let assoc = match assoc with None -> Gramext.LeftA | Some a -> a in - add_notation_in_scope df a (Some assoc,n,[],false) sc (split df) + add_notation_in_scope df a (Some assoc,n,[],false) None sc (split df) -let add_infix assoc n inf pr onlyparse sc = +let add_infix assoc n inf pr onlyparse mv8 sc = (* let pr = Astterm.globalize_qualid pr in*) (* check the precedence *) - if n<1 or n>10 then + if !Options.v7 & (n<1 or n>10) then errorlabstrm "Metasyntax.infix_grammar_entry" (str"Precedence must be between 1 and 10."); (* @@ -698,7 +733,11 @@ let add_infix assoc n inf pr onlyparse sc = let metas = [inject_var "x"; inject_var "y"] in let a = mkAppC (mkRefC pr,metas) in let df = "x "^(quote inf)^" y" in - add_notation_in_scope df a (assoc,n,[],onlyparse) sc (split df) + let mv8 = match mv8 with + None -> Some(split df,(assoc,n*10,[],false)) + | Some(a8,n8,s8) -> + Some(split ("x "^quote s8^" y"),(a8,n8,[],false)) in + add_notation_in_scope df a (assoc,n,[],onlyparse) mv8 sc (split df) (* Delimiters *) let load_delimiters _ (_,(scope,dlm)) = diff --git a/toplevel/metasyntax.mli b/toplevel/metasyntax.mli index 08a4d3c06..0c84a3ead 100644 --- a/toplevel/metasyntax.mli +++ b/toplevel/metasyntax.mli @@ -29,15 +29,17 @@ val add_tactic_grammar : (string * (string * grammar_production list) * raw_tactic_expr) list -> unit val add_infix : - grammar_associativity -> precedence -> string -> reference -> bool - -> scope_name option -> unit + grammar_associativity -> precedence -> string -> reference -> bool -> + (grammar_associativity * precedence * string) option -> + scope_name option -> unit val add_distfix : grammar_associativity -> precedence -> string -> reference -> scope_name option -> unit val add_delimiters : scope_name -> string -> unit val add_notation : string -> constr_expr - -> syntax_modifier list -> scope_name option -> unit + -> syntax_modifier list -> (string * syntax_modifier list) option + -> scope_name option -> unit val add_syntax_extension : string -> syntax_modifier list -> unit diff --git a/toplevel/mltop.ml4 b/toplevel/mltop.ml4 index 553da3493..d531d1d2e 100644 --- a/toplevel/mltop.ml4 +++ b/toplevel/mltop.ml4 @@ -221,9 +221,10 @@ let load_object mname fname= (* Summary of declared ML Modules *) -let loaded_modules = ref Stringset.empty -let get_loaded_modules () = Stringset.elements !loaded_modules -let add_loaded_module md = loaded_modules := Stringset.add md !loaded_modules +(* List and not Strinset because order is important *) +let loaded_modules = ref [] +let get_loaded_modules () = !loaded_modules +let add_loaded_module md = loaded_modules := md :: !loaded_modules let orig_loaded_modules = ref !loaded_modules let init_ml_modules () = loaded_modules := !orig_loaded_modules diff --git a/toplevel/recordobj.ml b/toplevel/recordobj.ml index 58fbb9781..9a8579aa0 100755 --- a/toplevel/recordobj.ml +++ b/toplevel/recordobj.ml @@ -32,7 +32,7 @@ let typ_lams_of t = let objdef_err ref = errorlabstrm "object_declare" - (pr_id (id_of_global None ref) ++ + (pr_id (id_of_global ref) ++ str" is not a structure object") let objdef_declare ref = diff --git a/toplevel/toplevel.ml b/toplevel/toplevel.ml index 399134d11..33f8e488f 100644 --- a/toplevel/toplevel.ml +++ b/toplevel/toplevel.ml @@ -237,10 +237,13 @@ let print_toplevel_error exc = (if is_pervasive_exn exc then (mt ()) else locstrm) ++ Cerrors.explain_exn exc +let is_term s = + if !Options.v7 then s="." else s=";" + (* Read the input stream until a dot is encountered *) let parse_to_dot = let rec dot st = match Stream.next st with - | ("", ".") -> () + | ("", c) when is_term c -> () | ("EOI", "") -> raise End_of_input | _ -> dot st in diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index c26741416..6eab9d36e 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -119,25 +119,38 @@ let rec vernac_com interpfun (loc,com) = System.fmt_time_difference tstart tend) (* To be interpreted in v7 or translator input only *) - | VernacV7only v -> if !Options.v7 then interp v + | VernacV7only v -> + if !Options.v7 || Options.do_translate() then interp v (* To be interpreted in translator output only *) - | VernacV8only v -> if true (* !translate *) then interp v + | VernacV8only v -> + if not !Options.v7 && not (do_translate()) then + interp v | v -> if not !just_parsing then interpfun v in try + Options.v7_only := false; if do_translate () then - let _ = Format.set_formatter_out_channel stdout in - let _ = interp com in - let _ = Format.set_formatter_out_channel !chan_translate in - if !translate_file then - msgnl (pr_comments !comments ++ pr_vernac com ++ sep_end) - else - msgnl (hov 4 (str"New Syntax:" ++ fnl() ++ pr_comments !comments ++ pr_vernac com ++ sep_end)); comments := None - else interp com - with e -> + (Format.set_formatter_out_channel !chan_translate; + Format.set_max_boxes max_int; + (match com with + | VernacV7only _ -> + Options.v7_only := true; + if !translate_file then msgnl (pr_comments !comments) + | _ -> + if !translate_file then + msgnl + (pr_comments !comments ++ hov 0 (pr_vernac com) ++ sep_end) + else + msgnl + (hov 4 (str"New Syntax:" ++ fnl() ++ pr_comments !comments ++ + pr_vernac com ++ sep_end))); + comments := None; + Format.set_formatter_out_channel stdout); + interp com; + with e -> Format.set_formatter_out_channel stdout; raise (DuringCommandInterp (loc, e)) @@ -197,8 +210,8 @@ let compile verbosely f = *) let ldir,long_f_dot_v = Library.start_library f in if !dump then dump_string ("F" ^ Names.string_of_dirpath ldir ^ "\n"); - chan_translate := if Options.do_translate () then open_out (f^".v8") else stdout; - let _ = Format.set_formatter_out_channel !chan_translate in + chan_translate := + if Options.do_translate () then open_out (f^".v8") else stdout; let _ = load_vernac verbosely long_f_dot_v in let _ = close_out !chan_translate in if Pfedit.get_all_proof_names () <> [] then diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index ef5430e3b..ed1c52de9 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -229,7 +229,7 @@ let print_located_qualid r = | ConstructRef _ -> "Constructor" | VarRef _ -> "Variable" in - let sp = Nametab.sp_of_global None ref in + let sp = Nametab.sp_of_global ref in str ref_str ++ spc () ++ pr_sp sp with Not_found -> try @@ -643,7 +643,10 @@ let vernac_set_end_tac tac = (* Auxiliary file management *) let vernac_require_from export spec filename = - Library.require_library_from_file spec None filename export + match export with + Some exp -> + Library.require_library_from_file spec None filename exp + | None -> Library.read_library_from_file filename let vernac_add_loadpath isrec pdir ldiropt = let alias = match ldiropt with @@ -1126,9 +1129,11 @@ let interp c = match c with | VernacDelimiters (sc,lr) -> vernac_delimiters sc lr | VernacOpenScope sc -> vernac_open_scope sc | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope qid scl - | VernacInfix (assoc,n,inf,qid,b,sc) -> vernac_infix assoc n inf qid b sc + | VernacInfix (assoc,n,inf,qid,b,mv8,sc) -> + vernac_infix assoc n inf qid b mv8 sc | VernacDistfix (assoc,n,inf,qid,sc) -> vernac_distfix assoc n inf qid sc - | VernacNotation (inf,c,pl,sc) -> vernac_notation inf c pl sc + | VernacNotation (inf,c,pl,mv8,sc) -> + vernac_notation inf c pl mv8 sc (* Gallina *) | VernacDefinition (k,id,d,f,_) -> vernac_definition k id d f diff --git a/toplevel/vernacexpr.ml b/toplevel/vernacexpr.ml index 53434bb8e..8204a1ead 100644 --- a/toplevel/vernacexpr.ml +++ b/toplevel/vernacexpr.ml @@ -156,17 +156,19 @@ type vernac_expr = (string * (string * grammar_production list) * raw_tactic_expr) list | VernacSyntax of string * raw_syntax_entry list | VernacSyntaxExtension of string * syntax_modifier list + | VernacDistfix of + grammar_associativity * precedence * string * reference * + scope_name option | VernacOpenScope of scope_name | VernacDelimiters of scope_name * string | VernacArgumentsScope of reference * scope_name option list | VernacInfix of grammar_associativity * precedence * string * reference * bool * - scope_name option - | VernacDistfix of - grammar_associativity * precedence * string * reference * + (grammar_associativity * precedence * string) option * scope_name option | VernacNotation of - string * constr_expr * syntax_modifier list * scope_name option + string * constr_expr * syntax_modifier list * + (string * syntax_modifier list) option * scope_name option (* Gallina *) | VernacDefinition of definition_kind * identifier * definition_expr * @@ -207,7 +209,7 @@ type vernac_expr = | VernacSolveExistential of int * constr_expr (* Auxiliary file and library management *) - | VernacRequireFrom of export_flag * specif_flag option * string + | VernacRequireFrom of export_flag option * specif_flag option * string | VernacAddLoadPath of rec_flag * string * dir_path option | VernacRemoveLoadPath of string | VernacAddMLPath of rec_flag * string diff --git a/translate/ppconstrnew.ml b/translate/ppconstrnew.ml index 07bd444f8..14375e04e 100644 --- a/translate/ppconstrnew.ml +++ b/translate/ppconstrnew.ml @@ -20,55 +20,12 @@ open Ppextend open Topconstr open Term -let sep = fun _ -> spc() +let sep = fun _ -> brk(1,0) let sep_p = fun _ -> str"." -let sep_v = fun _ -> str"," +let sep_v = fun _ -> str"," ++ spc() let sep_pp = fun _ -> str":" -let sep_pv = fun _ -> str"; " - -let dfltpr ast = (str"#GENTERM " ++ print_ast ast);; - -let pr_global ref = pr_global_env None ref - -let global_const_name kn = - try pr_global (ConstRef kn) - with Not_found -> (str ("CONST("^(string_of_kn kn)^")")) - -let global_var_name id = - try pr_global (VarRef id) - with Not_found -> (str ("SECVAR("^(string_of_id id)^")")) - -let global_ind_name (kn,tyi) = - try pr_global (IndRef (kn,tyi)) - with Not_found -> (str ("IND("^(string_of_kn kn)^","^(string_of_int tyi)^")")) - -let global_constr_name ((kn,tyi),i) = - try pr_global (ConstructRef ((kn,tyi),i)) - with Not_found -> (str ("CONSTRUCT("^(string_of_kn kn)^","^(string_of_int tyi)^","^(string_of_int i)^")")) - -let globpr gt = match gt with - | Nvar(_,s) -> (pr_id s) - | Node(_,"EVAR", [Num (_,ev)]) -> (str ("?" ^ (string_of_int ev))) - | Node(_,"CONST",[Path(_,sl)]) -> - global_const_name (section_path sl) - | Node(_,"SECVAR",[Nvar(_,s)]) -> - global_var_name s - | Node(_,"MUTIND",[Path(_,sl); Num(_,tyi)]) -> - global_ind_name (section_path sl, tyi) - | Node(_,"MUTCONSTRUCT",[Path(_,sl); Num(_,tyi); Num(_,i)]) -> - global_constr_name ((section_path sl, tyi), i) - | Dynamic(_,d) -> - if (Dyn.tag d) = "constr" then (str"") - else dfltpr gt - | gt -> dfltpr gt - -let wrap_exception = function - Anomaly (s1,s2) -> - warning ("Anomaly ("^s1^")"); pp s2; - str"" - | Failure _ | UserError _ | Not_found -> - str"" - | s -> raise s +let sep_bar = fun _ -> spc() ++ str"| " +let pr_tight_coma () = str "," ++ cut () let latom = 0 let lannot = 100 @@ -76,18 +33,23 @@ let lprod = 200 let llambda = 200 let lif = 200 let lletin = 200 -let lcases = 0 -let larrow = 90 -let lcast = 0 -let lapp = 10 +let lfix = 200 +let larrow = 80 +let lnegint = 40 +let lcast = 100 +let lapp = 1 let ltop = (200,E) +let lsimple = (0,E) let prec_less child (parent,assoc) = - match assoc with - | E -> (<=) child parent - | L -> (<) child parent - | Prec n -> child<=n - | Any -> true + 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) @@ -106,13 +68,7 @@ let pr_notation pr s env = prlist (print_hunk level pr env) unpl, level let pr_delimiters key strm = - let left = "`"^key^":" and right = "`" in - let lspace = - if is_letter (left.[String.length left -1]) then str " " else mt () in - let rspace = - let c = right.[0] in - if is_letter c or is_digit c or c = '\'' then str " " else mt () in - str left ++ lspace ++ strm ++ rspace ++ str right + strm ++ str ("%"^key) open Rawterm @@ -129,7 +85,7 @@ let pr_sort = function let pr_explicitation = function | None -> mt () - | Some n -> int n ++ str "!" + | Some n -> str "@" ++ int n ++ str ":=" let pr_expl_args pr (a,expl) = pr_explicitation expl ++ pr (lapp,L) a @@ -138,7 +94,9 @@ let pr_opt_type pr = function | CHole _ -> mt () | t -> cut () ++ str ":" ++ pr ltop t -let pr_tight_coma () = str "," ++ cut () +let pr_opt_type_spc pr = function + | CHole _ -> mt () + | t -> str " :" ++ brk(1,2) ++ pr ltop t let pr_name = function | Anonymous -> str"_" @@ -146,26 +104,50 @@ let pr_name = function let pr_located pr (loc,x) = pr x -let pr_let_binder pr x a = - hov 0 ( - str "let" ++ spc () ++ - pr_name x ++ spc () ++ - str "=" ++ spc () ++ - pr ltop a ++ spc () ++ - str "in") +let las = 2 + +let rec pr_patt inh p = + let (strm,prec) = match p with + | CPatAlias (_,p,id) -> + pr_patt (las,E) p ++ str " as " ++ pr_id id, las + | CPatCstr (_,c,[]) -> pr_reference c, latom + | CPatCstr (_,c,args) -> + pr_reference c ++ spc() ++ + prlist_with_sep sep (pr_patt (lapp,L)) args, lapp + | CPatAtom (_,None) -> str "_", latom + | CPatAtom (_,Some r) -> pr_reference r, latom + | CPatNumeral (_,i) -> Bignat.pr_bigint i, latom + | CPatDelimiters (_,k,p) -> + pr_delimiters k (pr_patt (latom,E) p), latom + in + if prec_less prec inh then strm + else str"(" ++ strm ++ str")" + +let pr_eqn pr (_,pl,rhs) = + spc() ++ hov 4 + (str "| " ++ + hv 0 (prlist_with_sep sep_v (pr_patt ltop) pl ++ str " =>") ++ + spc() ++ pr ltop rhs) + +(* let pr_binder pr (nal,t) = - h 0 ( + hov 0 ( prlist_with_sep sep (pr_located pr_name) nal ++ pr_opt_type pr t) +*) +let pr_oneb pr t na = + match t with + CHole _ -> pr_located pr_name na + | _ -> hov 1 + (str "(" ++ pr_located pr_name na ++ pr_opt_type pr t ++ str ")") +let pr_binder pr (nal,t) = + hov 0 (prlist_with_sep sep (pr_oneb pr t) nal) let pr_binders pr bl = - h 0 (prlist_with_sep sep (pr_binder pr) bl) + hv 0 (prlist_with_sep sep (pr_binder pr) bl) -let pr_recursive_decl pr id b t c = - pr_id id ++ - brk (1,2) ++ str ": " ++ pr ltop t ++ str ":=" ++ - brk (1,2) ++ pr ltop c +let pr_global vars ref = pr_global_env vars ref let split_lambda = function | CLambdaN (loc,[[na],t],c) -> (na,t,c) @@ -188,109 +170,152 @@ let rec split_fix n typ def = let (bl,typ,def) = split_fix (n-1) typ def in (([na],t)::bl,typ,def) +let pr_recursive_decl pr id b t c = + pr_id id ++ b ++ pr_opt_type_spc pr t ++ str " :=" ++ + brk(1,2) ++ pr ltop c + let pr_fixdecl pr (id,n,t,c) = let (bl,t,c) = split_fix (n+1) t c in - pr_recursive_decl pr id - (brk (1,2) ++ str "[" ++ pr_binders pr bl ++ str "]") t c + let annot = + let ids = List.flatten (List.map fst bl) in + if List.length ids > 1 then + spc() ++ str "{struct " ++ pr_name (snd (list_last ids)) ++ str"}" + else mt() in + pr_recursive_decl pr id (str" " ++ hov 0 (pr_binders pr bl) ++ annot) t c let pr_cofixdecl pr (id,t,c) = pr_recursive_decl pr id (mt ()) t c -let pr_recursive s pr_decl id = function +let pr_recursive pr_decl id = function | [] -> anomaly "(co)fixpoint with no definition" - | d1::dl -> - hov 0 ( - str "Fix " ++ pr_id id ++ brk (0,2) ++ str "{" ++ - (v 0 ( - (hov 0 (pr_decl d1)) ++ - (prlist (fun fix -> fnl () ++ hov 0 (str "with" ++ pr_decl fix)) - dl))) ++ - str "}") - -let pr_fix pr = pr_recursive "Fix" (pr_fixdecl pr) -let pr_cofix pr = pr_recursive "Fix" (pr_cofixdecl pr) + | [d1] -> pr_decl d1 + | dl -> + prlist_with_sep (fun () -> fnl() ++ str "with ") pr_decl dl ++ + fnl() ++ str "for " ++ pr_id id let rec pr_arrow pr = function - | CArrow (_,a,b) -> pr (larrow,L) a ++ str "->" ++ pr_arrow pr b - | a -> pr (larrow,E) a + | CArrow (_,a,b) -> pr (larrow,L) a ++ str " ->" ++ brk(1,0) ++ pr_arrow pr b + | a -> pr (-larrow,E) a -let pr_annotation pr t = str "<" ++ pr ltop t ++ str ">" - -let pr_cases _ _ _ = str "" +let pr_annotation pr po = + match po with + None -> mt() + | Some p -> spc() ++ str "=> " ++ hov 0 (pr ltop p) let rec pr inherited a = let (strm,prec) = match a with | CRef r -> pr_reference r, latom - | CFix (_,id,fix) -> pr_fix pr (snd id) fix, latom - | CCoFix (_,id,cofix) -> pr_cofix pr (snd id) cofix, latom - | CArrow _ -> h 0 (pr_arrow pr a), larrow + | CFix (_,id,fix) -> + hov 0 (str "fix " ++ pr_recursive (pr_fixdecl pr) (snd id) fix), + lfix + | CCoFix (_,id,cofix) -> + hov 0 (str "cofix " ++ pr_recursive (pr_cofixdecl pr) (snd id) cofix), + lfix + | CArrow _ -> hv 0 (pr_arrow pr a), larrow | CProdN (_,bl,a) -> - h 0 ( - str "!" ++ pr_binders pr bl ++ str "." ++ spc () ++ pr ltop a), lprod + hv 1 ( + str "!" ++ pr_binders pr bl ++ str "." ++ spc() ++ pr ltop a), lprod | CLambdaN (_,bl,a) -> - hov 0 ( - str "fun" ++ spc () ++ pr_binders pr bl ++ spc () ++ str "=>" ++ spc () ++ pr ltop a), + hov 2 ( + str "fun" ++ spc () ++ pr_binders pr bl ++ + str " =>" ++ spc() ++ pr ltop a), llambda | CLetIn (_,x,a,b) -> - h 0 (pr_let_binder pr (snd x) a ++ spc () ++ pr ltop b), lletin + hv 0 ( + hov 2 (str "let " ++ pr_name (snd x) ++ str " :=" ++ spc() ++ + pr ltop a ++ str " in") ++ + spc () ++ pr ltop b), + lletin | CAppExpl (_,f,l) -> - h 0 ( + hov 2 ( str "@" ++ pr_reference f ++ - prlist (fun a -> spc () ++ pr (lapp,L) a) l), lapp + prlist (fun a -> spc () ++ pr (lapp,L) a) l), + lapp | CApp (_,a,l) -> - h 0 ( - pr (lapp,E) a ++ - prlist (fun a -> spc () ++ pr_expl_args pr a) l), lapp + hov 2 ( + pr (lapp,L) a ++ + prlist (fun a -> spc () ++ pr_expl_args pr a) l), + lapp | CCases (_,po,c,eqns) -> - pr_cases po c eqns, lcases - | COrderedCase (_,IfStyle,po,c,[b1;b2]) -> + v 0 + (hov 4 (str "match " ++ prlist_with_sep sep_v (pr ltop) c ++ + pr_annotation pr po ++ str " with") ++ + prlist (pr_eqn pr) eqns ++ + spc() ++ str "end"), + latom + | COrderedCase (_,_,po,c,[b1;b2]) -> (* On force les parenthčses autour d'un "if" sous-terme (męme si le parsing est lui plus tolérant) *) - hov 0 ( - pr_opt (pr_annotation pr) po ++ - hv 0 ( - str "if" ++ pr ltop c ++ spc () ++ - hov 0 (str "then" ++ brk (1,1) ++ pr ltop b1) ++ spc () ++ - hov 0 (str "else" ++ brk (1,1) ++ pr ltop b2))), lif - | COrderedCase (_,LetStyle,po,c,[CLambdaN(_,[_,_ as bd],b)]) -> - hov 0 ( - pr_opt (pr_annotation pr) po ++ - hv 0 ( - str "let" ++ brk (1,1) ++ - hov 0 (str "(" ++ pr_binder pr bd ++ str ")") ++ - str "=" ++ brk (1,1) ++ - pr ltop c ++ spc () ++ - str "in " ++ pr ltop b)), lletin - | COrderedCase (_,(MatchStyle|RegularStyle as style),po,c,bl) -> - hov 0 ( - hov 0 ( - pr_opt (pr_annotation pr) po ++ brk (0,2) ++ - hov 0 ( - str (if style=RegularStyle then "Case" else "match") ++ - brk (1,1) ++ pr ltop c ++ spc () ++ - str (if style=RegularStyle then "of" else "with") ++ - brk (1,3) ++ - hov 0 (prlist (fun b -> pr ltop b ++ fnl ()) bl)) ++ - str "end")), lcases - | COrderedCase (_,_,_,_,_) -> - anomaly "malformed if or destructuring let" + hv 0 ( + str "if " ++ pr ltop c ++ pr_annotation pr po ++ spc () ++ + hov 0 (str "then" ++ brk (1,1) ++ pr ltop b1) ++ spc () ++ + hov 0 (str "else" ++ brk (1,1) ++ pr ltop b2)), + lif + | COrderedCase (_,_,po,c,[CLambdaN(_,[nal,_],b)]) -> + hv 0 ( + str "let " ++ + hov 0 (str "(" ++ + prlist_with_sep sep_v (fun (_,n) -> pr_name n) nal ++ + str ")" ++ + pr_annotation pr po ++ str " :=" ++ + spc() ++ pr ltop c ++ str " in") ++ + spc() ++ pr ltop b), + lletin + | COrderedCase (_,style,po,c,bl) -> + hv 0 ( + str (if style=MatchStyle then "old_match " else "match ") ++ + pr ltop c ++ + pr_annotation pr po ++ + str " with" ++ brk (1,0) ++ + hov 0 (prlist + (fun b -> str "| ??? =>" ++ spc() ++ pr ltop b ++ fnl ()) bl) ++ + str "end"), + latom | CHole _ -> str "_", latom - | CMeta (_,p) -> str "@" ++ int p, latom + | CMeta (_,p) -> str "?" ++ int p, latom | CSort (_,s) -> pr_sort s, latom | CCast (_,a,b) -> - hv 0 (pr (lcast,L) a ++ cut () ++ str "::" ++ pr (lcast,E) b), lcast + hv 0 (pr (lcast,L) a ++ cut () ++ str ":" ++ pr (-lcast,E) b), lcast | CNotation (_,s,env) -> pr_notation pr s env - | CNumeral (_,p) -> Bignat.pr_bigint p, latom - | CDelimiters (_,sc,a) -> pr_delimiters sc (pr ltop a), latom + | CNumeral (_,(Bignat.POS _ as p)) -> Bignat.pr_bigint p, latom + | CNumeral (_,(Bignat.NEG _ as p)) -> Bignat.pr_bigint p, lnegint + | CDelimiters (_,sc,a) -> pr_delimiters sc (pr (latom,E) a), latom | CDynamic _ -> str "", latom in if prec_less prec inherited then strm else str"(" ++ strm ++ str")" - -let pr_constr c = pr ltop (Constrextern.extern_rawconstr (Constrintern.for_grammar (Constrintern.interp_rawconstr Evd.empty (Global.env())) c)) + + +let rec fact_constr c = + match c with + CLambdaN(loc,bl1,CLambdaN(_,bl2,b)) -> + fact_constr (CLambdaN(loc,bl1@bl2,b)) + | CProdN(loc,bl1,CProdN(_,bl2,b)) -> + fact_constr (CProdN(loc,bl1@bl2,b)) + | _ -> map_constr_expr_with_binders + (fun _ -> fact_constr) (fun _ _ -> ()) () c +let pr l c = pr l (fact_constr c) + + +let transf env c = + if Options.do_translate() then + Constrextern.extern_rawconstr (Termops.vars_of_env env) + (Constrintern.for_grammar + (Constrintern.interp_rawconstr Evd.empty env) c) + else c + +let pr_constr_env env c = pr lsimple (transf env c) +let pr_lconstr_env env c = pr ltop (transf env 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_binders pr +let pr_cases_pattern = pr_patt ltop let pr_pattern = pr_constr -let pr_occurrences prc (nl,c) = prlist (fun n -> int n ++ spc ()) nl ++ prc c +let pr_occurrences prc (nl,c) = + prlist (fun n -> int n ++ spc ()) nl ++ + str"(" ++ prc c ++ str")" let pr_qualid qid = str (string_of_qualid qid) @@ -299,14 +324,14 @@ 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.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" + if r.rDelta then pr_arg str "delta" else mt () else - pr_arg str "Delta" ++ (if r.rDelta then str "-" else mt ()) ++ + pr_arg str "delta " ++ (if r.rDelta then str "-" else mt ()) ++ hov 0 (str "[" ++ prlist_with_sep spc pr r.rConst ++ str "]")) open Genarg @@ -317,54 +342,38 @@ let pr_metanum pr = function let pr_metaid id = str"?" ++ pr_id id -let pr_red_expr (pr_constr,pr_ref) = function - | Red false -> str "Red" - | Hnf -> str "Hnf" - | Simpl o -> str "Simpl" ++ pr_opt (pr_occurrences pr_constr) o +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_occurrences pr_lconstr) o | Cbv f -> if f = {rBeta=true;rIota=true;rZeta=true;rDelta=true;rConst=[]} then - str "Compute" + str "compute" else - hov 1 (str "Cbv" ++ spc () ++ pr_red_flag pr_ref f) + hov 1 (str "cbv" ++ pr_red_flag pr_ref f) | Lazy f -> - hov 1 (str "Lazy" ++ spc () ++ pr_red_flag pr_ref f) + hov 1 (str "lazy" ++ pr_red_flag pr_ref f) | Unfold l -> - hov 1 (str "Unfold" ++ + hov 1 (str "unfold" ++ prlist (fun (nl,qid) -> prlist (pr_arg int) nl ++ spc () ++ pr_ref qid) l) - | Fold l -> hov 1 (str "Fold" ++ prlist (pr_arg pr_constr) l) + | Fold l -> hov 1 (str "fold" ++ prlist (pr_arg pr_constr) l) | Pattern l -> - hov 1 (str "Pattern" ++ prlist (pr_occurrences pr_constr) l) + hov 1 (str "pattern" ++ pr_arg (prlist (pr_occurrences pr_lconstr)) l) | Red true -> error "Shouldn't be accessible from user" | ExtraRedExpr (s,c) -> hov 1 (str s ++ pr_arg pr_constr c) -let rec pr_may_eval pr = function +let rec pr_may_eval prc prlc = function | ConstrEval (r,c) -> hov 0 - (str "Eval" ++ brk (1,1) ++ pr_red_expr (pr,pr_metanum pr_reference) r ++ - spc () ++ str "in" ++ brk (1,1) ++ pr c) + (str "eval" ++ brk (1,1) ++ + pr_red_expr (prc,prlc,pr_metanum pr_reference) r ++ + str " in" ++ spc() ++ prlc c) | ConstrContext ((_,id),c) -> hov 0 - (str "Inst " ++ brk (1,1) ++ pr_id id ++ spc () ++ - str "[" ++ pr c ++ str "]") - | ConstrTypeOf c -> hov 0 (str "Check " ++ brk (1,1) ++ pr c) - | ConstrTerm c -> pr c - -(**********************************************************************) -let constr_syntax_universe = "constr" -(* This is starting precedence for printing constructions or tactics *) -(* Level 9 means no parentheses except for applicative terms (at level 10) *) -let constr_initial_prec = Some (9,Ppextend.L) - -let gentermpr_fail gt = - Esyntax.genprint globpr constr_syntax_universe constr_initial_prec gt - -let gentermpr gt = - try gentermpr_fail gt - with s -> wrap_exception s - -(* [at_top] means ids of env must be avoided in bound variables *) -let gentermpr_core at_top env t = - gentermpr (Termast.ast_of_constr at_top env t) + (str "inst " ++ pr_id id ++ spc () ++ + str "[" ++ prlc c ++ str "]") + | ConstrTypeOf c -> hov 1 (str "check" ++ spc() ++ prlc c) + | ConstrTerm c -> prlc c diff --git a/translate/ppconstrnew.mli b/translate/ppconstrnew.mli index 1488a9d17..27e1307c5 100644 --- a/translate/ppconstrnew.mli +++ b/translate/ppconstrnew.mli @@ -20,22 +20,29 @@ open Topconstr open Names open Util +val prec_less : int -> int * Ppextend.parenRelation -> bool + val split_fix : int -> constr_expr -> constr_expr -> (name located list * constr_expr) list * constr_expr * constr_expr -val pr_global : global_reference -> std_ppcmds - -val gentermpr : Coqast.t -> std_ppcmds -val gentermpr_core : bool -> env -> constr -> std_ppcmds +val pr_global : Idset.t -> global_reference -> std_ppcmds val pr_opt : ('a -> std_ppcmds) -> 'a option -> std_ppcmds val pr_name : name -> std_ppcmds val pr_qualid : qualid -> std_ppcmds val pr_red_expr : - ('a -> std_ppcmds) * ('b -> std_ppcmds) -> + ('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_may_eval : ('a -> std_ppcmds) -> 'a may_eval -> 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_cases_pattern : cases_pattern_expr -> std_ppcmds +val pr_binders : (name located list * constr_expr) list -> std_ppcmds +val pr_may_eval : + ('a -> std_ppcmds) -> ('a -> std_ppcmds) -> 'a may_eval -> std_ppcmds +val pr_metanum : ('a -> std_ppcmds) -> 'a or_metanum -> std_ppcmds +val pr_metaid : identifier -> std_ppcmds diff --git a/translate/pptacticnew.ml b/translate/pptacticnew.ml index 1cab30aa0..b80ae29d6 100644 --- a/translate/pptacticnew.ml +++ b/translate/pptacticnew.ml @@ -11,33 +11,17 @@ open Pp open Names open Nameops +open Environ open Util open Extend -open Ppconstr +open Ppextend +open Ppconstrnew open Tacexpr open Rawterm open Topconstr open Genarg open Libnames -(* Extensions *) -let prtac_tab = Hashtbl.create 17 - -let declare_extra_tactic_pprule s f g = - Hashtbl.add prtac_tab s (f,g) - -let genarg_pprule = ref Stringmap.empty - -let declare_extra_genarg_pprule (rawwit, f) (wit, g) = - let s = match unquote wit with - | ExtraArgType s -> s - | _ -> error - "Can declare a pretty-printing rule only for extra argument types" - in - let f x = f (out_gen rawwit x) in - let g x = g (out_gen wit x) in - genarg_pprule := Stringmap.add s (f,g) !genarg_pprule - (* [pr_rawtac] is here to cheat with ML typing system, gen_tactic_expr is polymorphic but with some occurrences of its instance raw_tactic_expr in it; then pr_tactic should be @@ -47,10 +31,10 @@ let declare_extra_genarg_pprule (rawwit, f) (wit, g) = let pr_rawtac = ref (fun _ -> failwith "Printer for raw tactic expr is not defined" - : raw_tactic_expr -> std_ppcmds) + : env -> raw_tactic_expr -> std_ppcmds) let pr_rawtac0 = ref (fun _ -> failwith "Printer for raw tactic expr is not defined" - : raw_tactic_expr -> std_ppcmds) + : env -> raw_tactic_expr -> std_ppcmds) let pr_arg pr x = spc () ++ pr x @@ -66,8 +50,6 @@ let pr_or_meta pr = function | AI x -> pr x | _ -> failwith "pr_hyp_location: unexpected quotation meta-variable" -let pr_casted_open_constr = pr_constr - let pr_quantified_hypothesis = function | AnonHyp n -> int n | NamedHyp id -> pr_id id @@ -78,16 +60,24 @@ let pr_binding prc = function | NamedHyp id, c -> hov 1 (pr_id id ++ str ":=" ++ cut () ++ prc c) | AnonHyp n, c -> hov 1 (int n ++ str ":=" ++ cut () ++ prc c) -let pr_bindings prc = function +let pr_esubst prc l = + let pr_qhyp = function + (_,AnonHyp n,c) -> int n ++ str":=" ++ prc c + | (_,NamedHyp id,c) -> pr_id id ++ str":=" ++ prc c in + prlist_with_sep spc pr_qhyp l + +let pr_bindings_gen for_ex prc = function | ImplicitBindings l -> - brk (1,1) ++ str "with" ++ brk (1,1) ++ - prlist_with_sep spc prc l + spc () ++ (if for_ex then mt() else str "with ") ++ + hv 0 (prlist_with_sep spc prc l) | ExplicitBindings l -> - str"TODO" (* brk (1,1) ++ str "with" ++ brk (1,1) ++ - prlist_with_sep spc (pr_binding prc) l *) + spc () ++ (if for_ex then mt() else str "with ") ++ + hv 0 (pr_esubst prc l) | NoBindings -> mt () -let pr_with_bindings prc (c,bl) = prc c ++ hv 0 (pr_bindings prc bl) +let pr_bindings prc = pr_bindings_gen false prc + +let pr_with_bindings prc (c,bl) = prc c ++ pr_bindings prc bl let pr_with_names = function | [] -> mt () @@ -106,7 +96,7 @@ let rec pr_intro_pattern = function let pr_hyp_location pr_id = function | InHyp id -> spc () ++ pr_id id - | InHypType id -> spc () ++ str "(Type of " ++ pr_id id ++ str ")" + | InHypType id -> spc () ++ str "(type of " ++ pr_id id ++ str ")" let pr_clause pr_id = function | [] -> mt () @@ -138,30 +128,35 @@ let pr_match_hyps = function let pr_match_rule m pr = function | Pat ([],mp,t) when m -> str "[" ++ pr_match_pattern mp ++ str "]" - ++ spc () ++ str "->" ++ brk (1,2) ++ pr t + ++ spc () ++ str "=>" ++ brk (1,4) ++ pr t | Pat (rl,mp,t) -> - str "[" ++ prlist_with_sep pr_semicolon pr_match_hyps rl ++ spc () ++ - str "|-" ++ spc () ++ pr_match_pattern mp ++ spc () ++ str "]" ++ - spc () ++ str "->" ++ brk (1,2) ++ pr t - | All t -> str "_" ++ spc () ++ str "->" ++ brk (1,2) ++ pr t + str "[" ++ prlist_with_sep (fun () -> str",") pr_match_hyps rl ++ + spc () ++ str "|-" ++ spc () ++ pr_match_pattern 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) -> hv 0(str k ++ pr_id id ++ str " =" ++ brk (1,1) ++ pr t) - | ((_,id),Some c,t) -> str "TODO(LETCLAUSE)" + | ((_,id),None,t) -> + hov 0 (str k ++ pr_id id ++ str " :=" ++ brk (1,1) ++ + pr (TacArg t)) + | ((_,id),Some c,t) -> + hv 0 (str k ++ pr_id id ++ str" :" ++ brk(1,2) ++ + pr_may_eval pr_constr pr_constr c ++ + str " :=" ++ brk (1,1) ++ pr (TacArg t)) let pr_let_clauses pr = function | hd::tl -> hv 0 (pr_let_clause "let " pr hd ++ spc () ++ - prlist_with_sep spc (pr_let_clause "and " pr) tl) + prlist_with_sep spc (pr_let_clause "with " pr) tl) | [] -> anomaly "LetIn must declare at least one binding" let pr_rec_clause pr ((_,id),(l,t)) = - pr_id id ++ prlist pr_funvar l ++ str "->" ++ spc () ++ pr t + pr_id id ++ prlist pr_funvar l ++ str "=>" ++ spc () ++ pr t let pr_rec_clauses pr l = prlist_with_sep (fun () -> fnl () ++ str "and ") (pr_rec_clause pr) l @@ -170,7 +165,7 @@ let pr_hintbases = function | None -> spc () ++ str "with *" | Some [] -> mt () | Some l -> - spc () ++ str "with" ++ hv 0 (prlist (fun s -> spc () ++ str s) l) + spc () ++ hov 2 (str "with" ++ prlist (fun s -> spc () ++ str s) l) let pr_autoarg_adding = function | [] -> mt () @@ -186,410 +181,335 @@ let pr_autoarg_usingTDB = function | true -> spc () ++ str "Using TDB" | false -> mt () -let rec pr_rawgen prtac x = - match Genarg.genarg_tag x with - | BoolArgType -> pr_arg str (if out_gen rawwit_bool x then "true" else "false") - | IntArgType -> pr_arg int (out_gen rawwit_int x) - | IntOrVarArgType -> pr_arg (pr_or_var pr_int) (out_gen rawwit_int_or_var x) - | StringArgType -> spc () ++ str "\"" ++ str (out_gen rawwit_string x) ++ str "\"" - | PreIdentArgType -> pr_arg str (out_gen rawwit_pre_ident x) - | IdentArgType -> pr_arg pr_id (out_gen rawwit_ident x) - | RefArgType -> pr_arg pr_reference (out_gen rawwit_ref x) - | SortArgType -> pr_arg pr_sort (out_gen rawwit_sort x) - | ConstrArgType -> pr_arg pr_constr (out_gen rawwit_constr x) - | ConstrMayEvalArgType -> - pr_arg (pr_may_eval pr_constr) (out_gen rawwit_constr_may_eval x) - | QuantHypArgType -> - pr_arg pr_quantified_hypothesis (out_gen rawwit_quant_hyp x) - | RedExprArgType -> - pr_arg (pr_red_expr (pr_constr,pr_metanum pr_reference)) (out_gen rawwit_red_expr x) - | TacticArgType -> pr_arg prtac (out_gen rawwit_tactic x) - | CastedOpenConstrArgType -> - pr_arg pr_casted_open_constr (out_gen rawwit_casted_open_constr x) - | ConstrWithBindingsArgType -> - pr_arg (pr_with_bindings pr_constr) - (out_gen rawwit_constr_with_bindings x) - | List0ArgType _ -> - hov 0 (fold_list0 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) - | List1ArgType _ -> - hov 0 (fold_list1 (fun x a -> pr_rawgen prtac x ++ a) x (mt())) - | OptArgType _ -> hov 0 (fold_opt (pr_rawgen prtac) (mt()) x) - | PairArgType _ -> - hov 0 - (fold_pair - (fun a b -> pr_rawgen prtac a ++ spc () ++ pr_rawgen prtac b) - x) - | ExtraArgType s -> - try fst (Stringmap.find s !genarg_pprule) x - with Not_found -> str " [no printer for " ++ str s ++ str "] " - -let rec pr_raw_tacarg_using_rule pr_gen = function - | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_raw_tacarg_using_rule pr_gen (l,al) - | Egrammar.TacNonTerm _ :: l, a :: al -> pr_gen a ++ pr_raw_tacarg_using_rule pr_gen (l,al) - | [], [] -> mt () - | _ -> failwith "Inconsistent arguments of extended tactic" - -let pr_raw_extend prt s l = - try - let (s,pl) = fst (Hashtbl.find prtac_tab s) l in - str s ++ pr_raw_tacarg_using_rule (pr_rawgen prt) (pl,l) - with Not_found -> - str "TODO(" ++ str s ++ prlist (pr_rawgen prt) l ++ str ")" - open Closure -let pr_evaluable_reference = function +let pr_evaluable_reference vars = function | EvalVarRef id -> pr_id id - | EvalConstRef sp -> pr_global (Libnames.ConstRef sp) - -let pr_inductive ind = pr_global (Libnames.IndRef ind) - -let rec pr_generic prtac x = - match Genarg.genarg_tag x with - | BoolArgType -> pr_arg str (if out_gen wit_bool x then "true" else "false") - | IntArgType -> pr_arg int (out_gen wit_int x) - | IntOrVarArgType -> pr_arg (pr_or_var pr_int) (out_gen wit_int_or_var x) - | StringArgType -> spc () ++ str "\"" ++ str (out_gen wit_string x) ++ str "\"" - | PreIdentArgType -> pr_arg str (out_gen wit_pre_ident x) - | IdentArgType -> pr_arg pr_id (out_gen wit_ident x) - | RefArgType -> pr_arg pr_global (out_gen wit_ref x) - | SortArgType -> pr_arg Printer.prterm (Term.mkSort (out_gen wit_sort x)) - | ConstrArgType -> pr_arg Printer.prterm (out_gen wit_constr x) - | ConstrMayEvalArgType -> - pr_arg Printer.prterm (out_gen wit_constr_may_eval x) - | QuantHypArgType -> - pr_arg pr_quantified_hypothesis (out_gen wit_quant_hyp x) - | RedExprArgType -> - pr_arg (pr_red_expr (Printer.prterm,pr_evaluable_reference)) (out_gen wit_red_expr x) - | TacticArgType -> pr_arg prtac (out_gen wit_tactic x) - | CastedOpenConstrArgType -> - pr_arg Printer.prterm (snd (out_gen wit_casted_open_constr x)) - | ConstrWithBindingsArgType -> - pr_arg (pr_with_bindings Printer.prterm) - (out_gen wit_constr_with_bindings x) - | List0ArgType _ -> - hov 0 (fold_list0 (fun x a -> pr_generic prtac x ++ a) x (mt())) - | List1ArgType _ -> - hov 0 (fold_list1 (fun x a -> pr_generic prtac x ++ a) x (mt())) - | OptArgType _ -> hov 0 (fold_opt (pr_generic prtac) (mt()) x) - | PairArgType _ -> - hov 0 - (fold_pair - (fun a b -> pr_generic prtac a ++ spc () ++ pr_generic prtac b) - x) - | ExtraArgType s -> - try snd (Stringmap.find s !genarg_pprule) x - with Not_found -> str "[no printer for " ++ str s ++ str "]" - -let rec pr_tacarg_using_rule prt = function - | Egrammar.TacTerm s :: l, al -> spc () ++ str s ++ pr_tacarg_using_rule prt (l,al) - | Egrammar.TacNonTerm _ :: l, a :: al -> pr_generic prt a ++ pr_tacarg_using_rule prt (l,al) - | [], [] -> mt () - | _ -> failwith "Inconsistent arguments of extended tactic" - -let pr_extend prt s l = - try - let (s,pl) = snd (Hashtbl.find prtac_tab s) l in - str s ++ pr_tacarg_using_rule prt (pl,l) - with Not_found -> - str s ++ prlist (pr_generic prt) l - -let make_pr_tac (pr_constr,pr_cst,pr_ind,pr_ident,pr_extend) = - -let pr_bindings = pr_bindings pr_constr in -let pr_with_bindings = pr_with_bindings pr_constr in -let pr_eliminator cb = str "using" ++ pr_arg pr_with_bindings cb in -let pr_constrarg c = spc () ++ pr_constr c in + | EvalConstRef sp -> pr_global vars (Libnames.ConstRef sp) + +let pr_inductive vars ind = pr_global vars (Libnames.IndRef ind) + +let make_pr_tac (pr_constr,pr_lconstr,pr_cst,pr_ind,pr_ident,pr_extend) = + +let pr_bindings env = pr_bindings (pr_constr env) in +let pr_ex_bindings env = pr_bindings_gen true (pr_constr env) in +let pr_with_bindings env = pr_with_bindings (pr_constr env) in +let pr_eliminator env cb = str "using" ++ pr_arg (pr_with_bindings env) cb 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 (* Printing tactics as arguments *) -let rec pr_atom0 = function - | TacIntroPattern [] -> str "Intros" - | TacIntroMove (None,None) -> str "Intro" - | TacAssumption -> str "Assumption" - | TacAnyConstructor None -> str "Constructor" - | TacTrivial (Some []) -> str "Trivial" - | TacAuto (None,Some []) -> str "Auto" - | TacAutoTDB None -> str "AutoTDB" - | TacDestructConcl -> str "DConcl" - | TacReflexivity -> str "Reflexivity" - | TacSymmetry -> str "Symmetry" - | t -> str "(" ++ pr_atom1 t ++ str ")" +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" + | TacAutoTDB None -> str "autotdb" + | TacDestructConcl -> str "dconcl" + | TacReflexivity -> str "reflexivity" + | TacSymmetry -> str "symmetry" + | t -> str "(" ++ pr_atom1 env t ++ str ")" (* Main tactic printer *) -and pr_atom1 = function - | TacExtend (_,s,l) -> pr_extend !pr_rawtac s l - | TacAlias (s,l,_) -> pr_extend !pr_rawtac s (List.map snd l) +and pr_atom1 env = function + | TacExtend (_,s,l) -> + pr_extend (Ppconstrnew.pr_constr_env env) (!pr_rawtac env) s l + | TacAlias (s,l,_) -> + pr_extend (Ppconstrnew.pr_constr_env env) (!pr_rawtac env) s + (List.map snd l) (* Basic tactics *) - | TacIntroPattern [] as t -> pr_atom0 t + | TacIntroPattern [] as t -> pr_atom0 env t | TacIntroPattern (_::_ as p) -> - hov 1 (str "Intros" ++ spc () ++ prlist_with_sep spc pr_intro_pattern 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 t - | TacIntroMove (Some id1,None) -> str "Intro " ++ pr_id id1 + 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_id id2) - | TacAssumption as t -> pr_atom0 t - | TacExact c -> hov 1 (str "Exact" ++ pr_arg pr_constr c) - | TacApply cb -> hov 1 (str "Apply" ++ spc () ++ pr_with_bindings cb) + (str "intro" ++ pr_opt pr_id ido1 ++ spc () ++ str "after " ++ pr_id id2) + | TacAssumption as t -> pr_atom0 env t + | TacExact c -> hov 1 (str "exact" ++ pr_lconstrarg env c) + | TacApply cb -> hov 1 (str "apply " ++ pr_with_bindings env cb) | TacElim (cb,cbo) -> - hov 1 (str "Elim" ++ pr_arg pr_with_bindings cb ++ - pr_opt pr_eliminator cbo) - | TacElimType c -> hov 1 (str "ElimType" ++ pr_arg pr_constr c) - | TacCase cb -> hov 1 (str "Case" ++ spc () ++ pr_with_bindings cb) - | TacCaseType c -> hov 1 (str "CaseType" ++ pr_arg pr_constr c) - | TacFix (ido,n) -> hov 1 (str "Fix" ++ pr_opt pr_id ido ++ pr_intarg n) + hov 1 (str "elim" ++ pr_arg (pr_with_bindings env) cb ++ + pr_opt (pr_eliminator env) cbo) + | TacElimType c -> hov 1 (str "elimtype" ++ pr_lconstrarg env c) + | TacCase cb -> hov 1 (str "case" ++ spc () ++ pr_with_bindings env cb) + | TacCaseType c -> hov 1 (str "CaseType" ++ pr_lconstrarg env c) + | TacFix (ido,n) -> hov 1 (str "fix" ++ pr_opt pr_id ido ++ pr_intarg n) | TacMutualFix (id,n,l) -> - hov 1 (str "Cofix" ++ spc () ++ pr_id id ++ pr_intarg n ++ spc () ++ + hov 1 (str "fix" ++ spc () ++ pr_id id ++ pr_intarg n ++ spc () ++ hov 0 (str "with" ++ brk (1,1) ++ prlist_with_sep spc (fun (id,n,c) -> - spc () ++ pr_id id ++ pr_intarg n ++ pr_arg pr_constr c) + spc () ++ pr_id id ++ pr_intarg n ++ pr_constrarg env c) l)) - | TacCofix ido -> hov 1 (str "Cofix" ++ pr_opt pr_id ido) + | TacCofix ido -> hov 1 (str "cofix" ++ pr_opt pr_id ido) | TacMutualCofix (id,l) -> - hov 1 (str "Cofix" ++ spc () ++ pr_id id ++ spc () ++ + hov 1 (str "cofix" ++ spc () ++ pr_id id ++ spc () ++ hov 0 (str "with" ++ brk (1,1) ++ - prlist (fun (id,c) -> spc () ++ pr_id id ++ pr_arg pr_constr c) + prlist (fun (id,c) -> spc () ++ pr_id id ++ pr_constrarg env c) l)) - | TacCut c -> hov 1 (str "Cut" ++ pr_arg pr_constr c) + | TacCut c -> hov 1 (str "cut" ++ pr_lconstrarg env c) | TacTrueCut (None,c) -> - hov 1 (str "Assert" ++ pr_arg pr_constr c) + hov 1 (str "assert" ++ pr_lconstrarg env c) | TacTrueCut (Some id,c) -> - hov 1 (str "Assert" ++ spc () ++ pr_id id ++ str ":" ++ pr_constr c) + hov 1 (str "assert" ++ spc () ++ pr_id id ++ str ":" ++ + pr_lconstr env c) | TacForward (false,na,c) -> - hov 1 (str "Assert" ++ pr_arg pr_name na ++ str ":=" ++ pr_constr c) + hov 1 (str "assert" ++ pr_arg pr_name na ++ str ":=" ++ + pr_lconstr env c) | TacForward (true,na,c) -> - hov 1 (str "Pose" ++ pr_arg pr_name na ++ str ":=" ++ pr_constr c) + hov 1 (str "pose" ++ pr_arg pr_name na ++ str ":=" ++ + pr_lconstr env c) | TacGeneralize l -> - hov 1 (str "Generalize" ++ spc () ++ prlist_with_sep spc pr_constr l) + hov 1 (str "generalize" ++ spc () ++ + prlist_with_sep spc (pr_constr env) l) | TacGeneralizeDep c -> - hov 1 (str "Generalize" ++ spc () ++ str "Dependent" ++ spc () ++ - pr_constr c) + hov 1 (str "generalize" ++ spc () ++ str "dependent" ++ + pr_lconstrarg env c) | TacLetTac (id,c,cl) -> - hov 1 (str "LetTac" ++ spc () ++ pr_id id ++ str ":=" ++ - pr_constr c ++ pr_clause_pattern pr_ident cl) + hov 1 (str "lettac" ++ spc () ++ pr_id id ++ str ":=" ++ + pr_constr env c ++ pr_clause_pattern pr_ident cl) | TacInstantiate (n,c) -> - hov 1 (str "Instantiate" ++ pr_arg int n ++ pr_arg pr_constr c) + hov 1 (str "instantiate" ++ pr_arg int n ++ pr_lconstrarg env c) (* Derived basic tactics *) | TacOldInduction h -> - hov 1 (str "Induction" ++ pr_arg pr_quantified_hypothesis h) + hov 1 (str "induction" ++ pr_arg pr_quantified_hypothesis h) | TacNewInduction (h,e,ids) -> - hov 1 (str "NewInduction" ++ spc () ++ pr_induction_arg pr_constr h ++ - pr_opt pr_eliminator e ++ pr_with_names ids) + hov 1 (str "newinduction" ++ spc () ++ + pr_induction_arg (pr_lconstr env) h ++ + pr_opt (pr_eliminator env) e ++ pr_with_names ids) | TacOldDestruct h -> - hov 1 (str "Destruct" ++ pr_arg pr_quantified_hypothesis h) + hov 1 (str "destruct" ++ pr_arg pr_quantified_hypothesis h) | TacNewDestruct (h,e,ids) -> - hov 1 (str "NewDestruct" ++ spc () ++ pr_induction_arg pr_constr h ++ - pr_opt pr_eliminator e ++ pr_with_names ids) + hov 1 (str "newdestruct" ++ spc () ++ + pr_induction_arg (pr_lconstr env) h ++ + pr_opt (pr_eliminator env) e ++ pr_with_names ids) | TacDoubleInduction (h1,h2) -> hov 1 - (str "Double Induction" ++ + (str "double induction" ++ pr_arg pr_quantified_hypothesis h1 ++ pr_arg pr_quantified_hypothesis h2) | TacDecomposeAnd c -> - hov 1 (str "Decompose Record" ++ pr_arg pr_constr c) + hov 1 (str "decompose record" ++ pr_lconstrarg env c) | TacDecomposeOr c -> - hov 1 (str "Decompose Sum" ++ pr_arg pr_constr c) + hov 1 (str "decompose sum" ++ pr_lconstrarg env c) | TacDecompose (l,c) -> - hov 1 (str "Decompose" ++ spc () ++ - hov 0 (str "[" ++ prlist_with_sep spc (pr_metanum pr_ind) l - ++ str "]")) + let vars = Termops.vars_of_env env in + hov 1 (str "decompose" ++ spc () ++ + hov 0 (str "[" ++ prlist_with_sep spc (pr_metanum (pr_ind vars)) l + ++ str "]" ++ pr_lconstrarg env c)) | TacSpecialize (n,c) -> - hov 1 (str "Specialize" ++ pr_opt int n ++ pr_with_bindings c) + hov 1 (str "specialize " ++ pr_opt int n ++ pr_with_bindings env c) | TacLApply c -> - hov 1 (str "LApply" ++ pr_constr c) + hov 1 (str "lapply" ++ pr_lconstrarg env c) (* Automation tactics *) - | TacTrivial (Some []) as x -> pr_atom0 x - | TacTrivial db -> hov 0 (str "Trivial" ++ pr_hintbases db) - | TacAuto (None,Some []) as x -> pr_atom0 x - | TacAuto (n,db) -> hov 0 (str "Auto" ++ pr_opt int n ++ pr_hintbases db) - | TacAutoTDB None as x -> pr_atom0 x - | TacAutoTDB (Some n) -> hov 0 (str "AutoTDB" ++ spc () ++ int n) - | TacDestructHyp (true,(_,id)) -> hov 0 (str "CDHyp" ++ spc () ++ pr_id id) - | TacDestructHyp (false,(_,id)) -> hov 0 (str "DHyp" ++ spc () ++ pr_id id) - | TacDestructConcl as x -> pr_atom0 x + | 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 int n ++ pr_hintbases db) + | TacAutoTDB None as x -> pr_atom0 env x + | TacAutoTDB (Some n) -> hov 0 (str "autotdb" ++ spc () ++ int n) + | TacDestructHyp (true,(_,id)) -> hov 0 (str "cdhyp" ++ spc () ++ pr_id id) + | TacDestructHyp (false,(_,id)) -> hov 0 (str "dhyp" ++ spc () ++ pr_id id) + | TacDestructConcl as x -> pr_atom0 env x | TacSuperAuto (n,l,b1,b2) -> - hov 1 (str "SuperAuto" ++ pr_opt int n ++ pr_autoarg_adding l ++ + hov 1 (str "superauto" ++ pr_opt int n ++ pr_autoarg_adding l ++ pr_autoarg_destructing b1 ++ pr_autoarg_usingTDB b2) | TacDAuto (n,p) -> - hov 1 (str "Auto" ++ pr_opt int n ++ str "Decomp" ++ pr_opt int p) + hov 1 (str "auto" ++ pr_opt int n ++ str "decomp" ++ pr_opt int p) (* Context management *) | TacClear l -> - hov 1 (str "Clear" ++ spc () ++ prlist_with_sep spc (pr_metanum pr_id) l) + hov 1 (str "clear" ++ spc () ++ prlist_with_sep spc (pr_metanum pr_id) l) | TacClearBody l -> - hov 1 (str "Clear" ++ spc () ++ prlist_with_sep spc (pr_metanum pr_id) l) + hov 1 (str "clear" ++ spc () ++ prlist_with_sep spc (pr_metanum pr_id) l) | TacMove (b,(_,id1),(_,id2)) -> (* Rem: only b = true is available for users *) assert b; hov 1 - (str "Move" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ + (str "move" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ str "after" ++ brk (1,1) ++ pr_id id2) | TacRename ((_,id1),(_,id2)) -> hov 1 - (str "Rename" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ + (str "rename" ++ brk (1,1) ++ pr_id id1 ++ spc () ++ str "into" ++ brk (1,1) ++ pr_id id2) (* Constructors *) - | TacLeft l -> hov 1 (str "Left" ++ pr_bindings l) - | TacRight l -> hov 1 (str "Right" ++ pr_bindings l) - | TacSplit l -> hov 1 (str "Split" ++ pr_bindings l) + | 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_rawtac0 t) - | TacAnyConstructor None as t -> pr_atom0 t + hov 1 (str "constructor" ++ pr_arg (!pr_rawtac0 env) t) + | TacAnyConstructor None as t -> pr_atom0 env t | TacConstructor (n,l) -> - hov 1 (str "Constructor" ++ pr_or_meta pr_intarg n ++ pr_bindings l) + hov 1 (str "constructor" ++ pr_or_meta pr_intarg n ++ + pr_bindings env l) (* Conversion *) | TacReduce (r,h) -> - hov 1 (pr_red_expr (pr_constr,pr_cst) r ++ pr_clause pr_ident h) - | TacChange (_,c,h) -> (* A Verifier *) - hov 1 (str "Change" ++ brk (1,1) ++ pr_constr c ++ pr_clause pr_ident h) + hov 1 (pr_red_expr (pr_constr env,pr_lconstr env,pr_cst) r ++ + pr_clause pr_ident h) + | TacChange (occ,c,h) -> (* A Verifier *) + hov 1 (str "change" ++ brk (1,1) ++ + (match occ with + None -> mt() + | Some(ocl,c1) -> + hov 1 (prlist (fun i -> int i ++ spc()) ocl ++ + pr_constr env c1) ++ spc() ++ str "with ") ++ + pr_constr env c ++ pr_clause pr_ident h) (* Equivalence relations *) - | (TacReflexivity | TacSymmetry) as x -> pr_atom0 x - | TacTransitivity c -> str "Transitivity" ++ pr_arg pr_constr c - -and pr_tactic_seq_body tl = + | (TacReflexivity | TacSymmetry) as x -> pr_atom0 env x + | TacTransitivity c -> str "transitivity" ++ pr_lconstrarg env c in + +let ltop = (5,E) in +let lseq = 5 in +let ltactical = 3 in +let lorelse = 2 in +let llet = 1 in +let lfun = 1 in +let labstract = 1 in +let lmatch = 1 in +let latom = 0 in +let lcall = 1 in +let leval = 1 in +let ltatom = 1 in + +let pr_seq_body pr tl = hv 0 (str "[ " ++ - prlist_with_sep (fun () -> spc () ++ str "| ") prtac tl ++ str " ]") - - (* Strictly closed atomic tactic expressions *) -and pr0 = function - | TacFirst tl -> str "First" ++ spc () ++ pr_tactic_seq_body tl - | TacSolve tl -> str "Solve" ++ spc () ++ pr_tactic_seq_body tl - | TacId -> str "Idtac" - | TacFail 0 -> str "Fail" - | TacAtom (_,t) -> pr_atom0 t - | TacArg c -> pr_tacarg c - | t -> str "(" ++ prtac t ++ str ")" - - (* Semi-closed atomic tactic expressions *) -and pr1 = function - | TacAtom (_,t) -> pr_atom1 t - | TacFail n -> str "Fail " ++ int n - | t -> pr0 t - - (* Orelse tactic expressions (printed as if parsed associating on the right - though the semantics is purely associative) *) -and pr2 = function - | TacOrelse (t1,t2) -> - hov 1 (pr1 t1 ++ str " Orelse" ++ brk (1,1) ++ pr3 t2) - | t -> pr1 t - - (* Non closed prefix tactic expressions *) -and pr3 = function - | TacTry t -> hov 1 (str "Try" ++ spc () ++ pr3 t) - | TacDo (n,t) -> hov 1 (str "Do " ++ int n ++ spc () ++ pr3 t) - | TacRepeat t -> hov 1 (str "Repeat" ++ spc () ++ pr3 t) - | TacProgress t -> hov 1 (str "Progress" ++ spc () ++ pr3 t) - | TacInfo t -> hov 1 (str "Info" ++ spc () ++ pr3 t) - | t -> pr2 t - -and pr4 = function - | t -> pr3 t - - (* THEN and THENS tactic expressions (printed as if parsed - associating on the left though the semantics is purely associative) *) -and pr5 = function - | TacThens (t,tl) -> - hov 1 (pr5 t ++ spc () ++ str "&" ++ spc () ++ pr_tactic_seq_body tl) - | TacThen (t1,t2) -> - hov 1 (pr5 t1 ++ spc () ++ str "&" ++ spc () ++ pr4 t2) - | t -> pr4 t - - (* Ltac tactic expressions *) -and pr6 = function - |(TacAtom _ - | TacThen _ - | TacThens _ - | TacFirst _ - | TacSolve _ - | TacTry _ - | TacOrelse _ - | TacDo _ - | TacRepeat _ - | TacProgress _ - | TacId - | TacFail _ - | TacInfo _) as t -> pr5 t - - | TacAbstract (t,None) -> str "Abstract " ++ pr6 t + prlist_with_sep (fun () -> spc () ++ str "| ") (pr ltop) tl ++ + str " ]") in + +let rec pr_tac env inherited tac = + let (strm,prec) = match tac with + | TacAbstract (t,None) -> + str "abstract " ++ pr_tac env (labstract,E) t, labstract | TacAbstract (t,Some s) -> hov 0 - (str "Abstract " ++ pr6 t ++ spc () ++ str "using" ++ spc () ++ pr_id s) + (str "abstract " ++ pr_tac env ltop t ++ spc () ++ + str "using " ++ pr_id s), + labstract | TacLetRecIn (l,t) -> hv 0 - (str "let rec " ++ pr_rec_clauses prtac l ++ - spc () ++ str "in" ++ fnl () ++ prtac t) + (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_tacarg0 llc ++ spc () ++ str "in") ++ fnl () ++ prtac u) - | TacLetCut llc -> - pr_let_clauses pr_tacarg0 - (List.map (fun (id,c,t) -> ((dummy_loc,id),Some c,t)) llc) - ++ fnl () + (hv 0 (pr_let_clauses (pr_tac env ltop) llc ++ str " in") ++ + fnl () ++ pr_tac env (llet,E) u), + llet + | TacLetCut llc -> assert false | TacMatch (t,lrul) -> - hov 0 (str "match" ++ spc () ++ pr_may_eval Ppconstr.pr_constr t ++ spc() - ++ str "with" + hov 0 (str "match " ++ + pr_may_eval (Ppconstrnew.pr_constr_env env) + (Ppconstrnew.pr_lconstr_env env) t + ++ str " with" ++ prlist - (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule true prtac r) - lrul) + (fun r -> fnl () ++ str "| " ++ + pr_match_rule true (pr_tac env ltop) r) + lrul + ++ fnl() ++ str "end"), + lmatch | TacMatchContext (lr,lrul) -> hov 0 ( - str (if lr then "Match Reverse Context With" else "Match Context With") + str (if lr then "match reverse context with" else "match context with") ++ prlist - (fun r -> fnl () ++ str "|" ++ spc () ++ pr_match_rule false prtac r) - lrul) + (fun r -> fnl () ++ str "| " ++ + pr_match_rule false (pr_tac env ltop) r) + lrul + ++ fnl() ++ str "end"), + lmatch | TacFun (lvar,body) -> - hov 0 (str "fun" ++ - prlist pr_funvar lvar ++ spc () ++ str "->" ++ spc () ++ prtac body) - | TacArg c -> pr_tacarg c + 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 ++ str " &" ++ spc () ++ + pr_seq_body (pr_tac env) tl), + lseq + | TacThen (t1,t2) -> + hov 1 (pr_tac env (lseq,E) t1 ++ str " &" ++ 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 " ++ 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 0 -> str "fail", latom + | TacFail n -> str "fail " ++ int n, latom + | TacFirst tl -> + str "first" ++ spc () ++ pr_seq_body (pr_tac env) tl, llet + | TacSolve tl -> + str "solve" ++ spc () ++ pr_seq_body (pr_tac env) tl, llet + | TacId -> str "idtac", latom + | TacAtom (_,t) -> hov 1 (pr_atom1 env t), ltatom + | TacArg(Tacexp e) -> !pr_rawtac0 env e, latom + | TacArg(ConstrMayEval (ConstrTerm c)) -> str "'" ++ pr_constr env c, latom + | TacArg(ConstrMayEval c) -> + pr_may_eval (pr_constr env) (pr_lconstr env) c, leval + | TacArg(Integer n) -> int n, latom + | TacArg(TacCall(_,f,l)) -> + hov 1 (pr_reference 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_tacarg0 = function +and pr_tacarg env = function | TacDynamic (_,t) -> str ("") | MetaNumArg (_,n) -> str ("?" ^ string_of_int n ) | MetaIdArg (_,s) -> str ("$" ^ s) | TacVoid -> str "()" | Reference r -> pr_reference r - | ConstrMayEval (ConstrTerm c) -> str "'" ++ pr_constr c - | ConstrMayEval c -> pr_may_eval pr_constr c - | Integer n -> int n - | (TacCall _ | Tacexp _) as t -> str "(" ++ pr_tacarg1 t ++ str ")" - -and pr_tacarg1 = function - | TacCall (_,f,l) -> - hov 0 (pr_reference f ++ spc () ++ prlist_with_sep spc pr_tacarg0 l) - | Tacexp t -> !pr_rawtac t - | t -> pr_tacarg0 t + | ConstrMayEval (ConstrTerm c) -> pr_constr env c + | (ConstrMayEval _|TacCall _|Tacexp _|Integer _) as a -> + str "'" ++ pr_tac env (latom,E) (TacArg a) -and pr_tacarg x = pr_tacarg1 x - -and prtac x = pr6 x - -in (prtac,pr0,pr_match_rule) +in ((fun env -> pr_tac env ltop), + (fun env -> pr_tac env (latom,E)), + pr_match_rule) let (pr_raw_tactic,pr_raw_tactic0,pr_match_rule) = make_pr_tac - (Ppconstrnew.pr_constr, + (Ppconstrnew.pr_constr_env, + Ppconstrnew.pr_lconstr_env, pr_metanum pr_reference, - pr_reference, + (fun _ -> pr_reference), pr_or_meta (fun (loc,id) -> pr_id id), - pr_raw_extend) + Pptactic.pr_raw_extend) let _ = pr_rawtac := pr_raw_tactic let _ = pr_rawtac0 := pr_raw_tactic0 -let (pr_tactic,_,_) = - make_pr_tac - (Printer.prterm, - pr_evaluable_reference, - pr_inductive, - pr_id, - pr_extend) +let pr_gen env = + Pptactic.pr_rawgen (Ppconstrnew.pr_constr_env env) + (pr_raw_tactic env) diff --git a/translate/pptacticnew.mli b/translate/pptacticnew.mli index 152b7e630..41fc97da9 100644 --- a/translate/pptacticnew.mli +++ b/translate/pptacticnew.mli @@ -12,22 +12,8 @@ open Pp open Genarg open Tacexpr open Proof_type +open Topconstr -val declare_extra_genarg_pprule : - ('a raw_abstract_argument_type * ('a -> std_ppcmds)) -> - ('b closed_abstract_argument_type * ('b -> std_ppcmds)) -> unit - -val declare_extra_tactic_pprule : - string -> - (raw_generic_argument list -> - string * Egrammar.grammar_tactic_production list) - -> (closed_generic_argument list -> - string * Egrammar.grammar_tactic_production list) - -> unit - -val pr_match_rule : bool -> (raw_tactic_expr -> std_ppcmds) -> - (pattern_expr,raw_tactic_expr) match_rule -> std_ppcmds - -val pr_raw_tactic : raw_tactic_expr -> std_ppcmds - -val pr_tactic : Proof_type.tactic_expr -> std_ppcmds +val pr_raw_tactic : Environ.env -> raw_tactic_expr -> std_ppcmds +val pr_gen : Environ.env -> + (constr_expr, raw_tactic_expr) generic_argument -> std_ppcmds diff --git a/translate/ppvernacnew.ml b/translate/ppvernacnew.ml index bfebda362..5f5e3b121 100644 --- a/translate/ppvernacnew.ml +++ b/translate/ppvernacnew.ml @@ -26,29 +26,51 @@ open Libnames open Ppextend open Topconstr +let pr_raw_tactic tac = + Pptacticnew.pr_raw_tactic (Global.env()) tac +let pr_raw_tactic_goal n tac = + let (_,env) = Pfedit.get_goal_context n in + Pptacticnew.pr_raw_tactic env tac +let pr_lconstr_goal n c = + let (_,env) = Pfedit.get_goal_context n in + Ppconstrnew.pr_lconstr_env env c + +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 + | (s,pargs)::rls -> + if extract_signature pargs = tys then (s,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 sep_pv = fun _ -> str"; " let pr_located pr (loc,x) = pr x +let pr_ne_sep sep pr = function + [] -> mt() + | l -> sep() ++ pr l + let pr_entry_prec = function | Some Gramext.LeftA -> spc() ++ str"LEFTA" | Some Gramext.RightA -> spc() ++ str"RIGHTA" | Some Gramext.NonA -> spc() ++ str"NONA" | None -> mt() -let pr_metaid _ = str"" -let pr_metanum _ = str"" - let pr_set_entry_type = function | ETIdent -> str"ident" | ETReference -> str"global" | ETPattern -> str"pattern" - | ETConstr _ | ETOther (_,_) -> mt () - | ETBigint -> str"TODO" + | ETConstr _ -> str"constr" + | ETOther (_,e) -> str e + | ETBigint -> str "bigint" let pr_non_terminal = function | NtQual (u,nt) -> str u ++ str" : " ++ str nt @@ -73,7 +95,7 @@ let pr_search a b pr_c = match a with | SearchHead qid -> str"Search" ++ spc() ++ pr_reference qid ++ spc() ++ pr_in_out_modules b | SearchPattern c -> str"SearchPattern" ++ spc() ++ pr_c c ++ spc() ++ pr_in_out_modules b | SearchRewrite c -> str"SearchRewrite" ++ spc() ++ pr_c c ++ spc() ++ pr_in_out_modules b - | SearchAbout _ -> str"TODO" + | SearchAbout qid -> str"SearchAbout" ++ spc() ++ pr_reference qid ++ spc() ++ pr_in_out_modules b let pr_class_rawexpr = function | FunClass -> str"FUNCLASS" @@ -110,29 +132,72 @@ let pr_hints db h pr_c = | [] -> (false , mt()) | c1::c2 -> match c1 with | None,_ -> (false , mt()) - | Some name,_ -> (true , pr_id name) - in let opth = pr_opt_hintbases db - in let pr_aux = function + | Some name,_ -> (true , pr_id name) in + let opth = pr_opt_hintbases db in + let pr_aux = function | CRef qid -> pr_reference qid - | _ -> mt () - in match h with - | HintsResolve l -> let (f,dbn) = db_name l in if f then hov 1 (str"Hint" ++ spc() ++ dbn ++ spc() ++ opth ++ spc() ++ str":=" ++ spc() ++ str"Resolve" ++ spc() ++ prlist_with_sep sep pr_c (List.map (fun (_,y) -> y) l)) else hov 1 (str"Hints Resolve" ++ spc() ++ prlist_with_sep sep pr_aux (List.map (fun (_,y) -> y) l) ++ spc() ++ opth) - | HintsImmediate l -> let (f,dbn) = db_name l in if f then hov 1 (str"Hint" ++ spc() ++ dbn ++ spc() ++ opth ++ spc() ++ str":=" ++ spc() ++ str"Immediate" ++ spc() ++ prlist_with_sep sep pr_c (List.map (fun (_,y) -> y) l)) else hov 1 (str"Hints Immediate" ++ spc() ++ prlist_with_sep sep pr_aux (List.map (fun (_,y) -> y) l) ++ spc() ++ opth) - | HintsUnfold l -> let (f,dbn) = db_name l in if f then hov 1 (str"Hint" ++ spc() ++ dbn ++ spc() ++ opth ++ spc() ++ str":=" ++ spc() ++ str"Unfold" ++ spc() ++ prlist_with_sep sep pr_reference (List.map (fun (_,y) -> y) l)) else hov 1 (str"Hints Unfold" ++ spc() ++ prlist_with_sep sep pr_reference (List.map (fun (_,y) -> y) l) ++ spc() ++ opth) - | HintsConstructors (n,c) -> hov 1 (str"Hint" ++ spc() ++ pr_id n ++ spc() ++ opth ++ spc() ++ str":=" ++ spc() ++ str"Constructors" ++ spc() ++ pr_reference c) - | HintsExtern (name,n,c,tac) -> hov 1 (str"Hint" ++ spc() ++ pr_id name ++ spc() ++ opth ++ spc() ++ str":=" ++ spc() ++ str"Extern" ++ spc() ++ int n ++ spc() ++ pr_c c ++ pr_raw_tactic tac) + | _ -> mt () in + match h with + | HintsResolve l -> + let (f,dbn) = db_name l in + if f then + hov 1 (str"Hint " ++ dbn ++ spc() ++ opth ++ + str" :=" ++ spc() ++ str"Resolve" ++ spc() ++ + prlist_with_sep sep pr_c (List.map (fun (_,y) -> y) l)) + else hov 1 + (str"Hints Resolve " ++ + prlist_with_sep sep pr_aux + (List.map (fun (_,y) -> y) l) ++ spc() ++ opth) + | HintsImmediate l -> + let (f,dbn) = db_name l in + if f then + hov 1 (str"Hint " ++ dbn ++ spc() ++ opth ++ + str" :=" ++ spc() ++ str"Immediate" ++ spc() ++ + prlist_with_sep sep pr_c (List.map (fun (_,y) -> y) l)) + else hov 1 + (str"Hints Immediate " ++ + prlist_with_sep sep pr_aux + (List.map (fun (_,y) -> y) l) ++ spc() ++ opth) + | HintsUnfold l -> + let (f,dbn) = db_name l in + if f then + hov 1 (str"Hint" ++ spc() ++ dbn ++ spc() ++ opth ++ + str" :=" ++ spc() ++ str"Unfold" ++ spc() ++ + prlist_with_sep sep pr_reference + (List.map snd l)) + else hov 1 + (str"Hints Unfold " ++ prlist_with_sep sep pr_reference + (List.map snd l) ++ spc() ++ opth) + | HintsConstructors (n,c) -> + hov 1 (str"Hint " ++ pr_id n ++ spc() ++ opth ++ str" :=" ++ + spc() ++ str"Constructors" ++ spc() ++ pr_reference c) + | HintsExtern (name,n,c,tac) -> + hov 1 (str"Hint " ++ pr_id name ++ spc() ++ opth ++ str" :=" ++ + spc() ++ str"Extern " ++ int n ++ spc() ++ pr_c c ++ + spc() ++ pr_raw_tactic tac) let pr_with_declaration pr_c = function - | CWith_Definition (id,c) -> str"Definition" ++ spc() ++ pr_id id ++ str" := " ++ pr_c c - | CWith_Module (id,qid) -> str"Module" ++ spc() ++ pr_id id ++ str" := " ++ pr_located pr_qualid qid + | CWith_Definition (id,c) -> + str"Definition" ++ spc() ++ pr_id id ++ str" := " ++ pr_c c + | CWith_Module (id,qid) -> + str"Module" ++ spc() ++ pr_id id ++ str" := " ++ + pr_located pr_qualid qid let rec pr_module_type pr_c = function | CMTEident qid -> pr_located pr_qualid qid - | CMTEwith (mty,decl) -> pr_module_type pr_c mty ++ spc() ++ str"with" ++ spc() ++ pr_with_declaration pr_c decl + | CMTEwith (mty,decl) -> + pr_module_type pr_c mty ++ spc() ++ str" with" ++ + pr_with_declaration pr_c decl -let pr_module_vardecls pr_c (l,mty) = prlist_with_sep (fun _ -> str",") pr_id l ++ spc() ++ pr_module_type pr_c mty +let pr_module_vardecls pr_c (l,mty) = + prlist + (fun id -> + spc() ++ str"(" ++ pr_id id ++ str":" ++ + pr_module_type pr_c mty ++ str")") + l -let pr_module_binders l pr_c = str"[" ++ prlist_with_sep (fun _ -> str";") (pr_module_vardecls pr_c) l ++ str"]" +let pr_module_binders l pr_c = + prlist (pr_module_vardecls pr_c) l let pr_module_binders_list l pr_c = pr_module_binders l pr_c @@ -148,15 +213,32 @@ let pr_type_option pr_c = function | CHole loc -> mt() | _ as c -> str":" ++ pr_c c +let pr_binder pr_c ty na = + match ty with + CHole _ -> pr_located pr_name na + | _ -> + hov 1 + (str "(" ++ pr_located pr_name na ++ str ":" ++ cut() ++ + pr_c ty ++ str")") + let pr_valdecls pr_c = function - | LocalRawAssum (na,c) -> prlist_with_sep (fun _ -> str",") (pr_located pr_name) na ++ spc() ++ pr_type_option pr_c c - | LocalRawDef (na,c) -> pr_located pr_name na ++ spc() ++ str":=" ++ spc() ++ pr_opt_casted_constr pr_c c + | LocalRawAssum (na,c) -> + prlist_with_sep spc (pr_binder pr_c c) na + | LocalRawDef (na,c) -> + hov 1 (str"(" ++ pr_located pr_name na ++ str":=" ++ cut() ++ + pr_c c ++ str")") -let pr_binders pr_c l = match l with -| [] -> mt() -| _ as t -> str"(" ++ prlist_with_sep (fun _ -> str") (") (pr_valdecls pr_c) t ++ str")" +let pr_vbinders pr_c l = + hv 0 (prlist_with_sep spc (pr_valdecls pr_c) l) -let pr_onescheme (id,dep,ind,s) = pr_id id ++ spc() ++ str":=" ++ spc() ++ if dep then str"Induction for" else str"Minimality for" ++ spc() ++ pr_reference ind ++ spc() ++ str"Sort" ++ spc() ++ pr_sort s +let pr_sbinders sbl = + let bl = List.map (fun (id,c) -> ([(dummy_loc,Name id)],c)) sbl in + pr_binders bl + +let pr_onescheme (id,dep,ind,s) = + pr_id id ++ str" :=" ++ spc() ++ + (if dep then str"Induction for" else str"Minimality for") + ++ spc() ++ pr_reference ind ++ spc() ++ str"Sort" ++ spc() ++ pr_sort s let pr_class_rawexpr = function | FunClass -> str"FUNCLASS" @@ -169,9 +251,12 @@ let pr_assumption_token = function | (Decl_kinds.Global,Decl_kinds.Logical) -> str"Axiom" | (Decl_kinds.Global,Decl_kinds.Definitional) -> str"Parameter" -let pr_params pr_c (a,(b,c)) = pr_id b ++ spc() ++ if a then str":>" else str":" ++ spc() ++ pr_c c +let pr_params pr_c (a,(b,c)) = + hov 2 (str"(" ++ pr_id b ++ spc() ++ + (if a then str":>" else str":") ++ + spc() ++ pr_c c ++ str")") -let pr_ne_params_list pr_c l = prlist_with_sep sep_pv (pr_params pr_c) l +let pr_ne_params_list pr_c l = prlist_with_sep spc (pr_params pr_c) l let pr_thm_token = function | Decl_kinds.Theorem -> str"Theorem" @@ -179,8 +264,15 @@ let pr_thm_token = function | Decl_kinds.Fact -> str"Fact" | Decl_kinds.Remark -> str"Remark" +let pr_require_token = function + | Some true -> str "Export" + | Some false -> str "Import" + | None -> str "Closed" + let pr_syntax_modifier = function - | SetItemLevel (l,n) -> prlist_with_sep (fun _ -> str",") str l ++ spc() ++ str"at level" ++ spc() ++ int n + | SetItemLevel (l,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" @@ -188,7 +280,10 @@ let pr_syntax_modifier = function | SetEntryType (x,typ) -> str x ++ spc() ++ pr_set_entry_type typ | SetOnlyParsing -> str"only parsing" -let pr_grammar_tactic_rule (name,(s,pil),t) = str name ++ spc() ++ str"[" ++ qs s ++ spc() ++ prlist_with_sep sep pr_production_item pil ++ str"]" ++ spc() ++ str"->" ++ spc() ++ str"[" ++ pr_raw_tactic t ++ str"]" +let pr_grammar_tactic_rule (name,(s,pil),t) = + str name ++ spc() ++ str"[" ++ qs s ++ spc() ++ + prlist_with_sep sep pr_production_item pil ++ str"]" ++ + spc() ++ str"->" ++ spc() ++ str"[" ++ pr_raw_tactic t ++ str"]" let pr_box b = let pr_boxkind = function | PpHB n -> str"h" ++ spc() ++ int n @@ -216,22 +311,26 @@ let rec pr_next_hunks = function | 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 = prlist_with_sep sep pr_next_hunks u +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 ++ spc() ++ str":" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"|") pr_syntax_rule rl +let pr_syntax_entry (p,rl) = + str"level" ++ spc() ++ int p ++ str" :" ++ fnl() ++ + prlist_with_sep (fun _ -> fnl() ++ str"| ") pr_syntax_rule rl let sep_end = str";" (**************************************) (* Pretty printer for vernac commands *) (**************************************) -let make_pr_vernac pr_constr = +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 @@ -241,7 +340,7 @@ let rec pr_vernac = function | VernacRestart -> str"Restart" | VernacSuspend -> str"Suspend" | VernacUnfocus -> str"Unfocus" - | VernacGoal c -> str"Goal" ++ pr_constrarg c + | VernacGoal c -> str"Goal" ++ pr_lconstrarg c | VernacAbort id -> str"Abort" ++ pr_opt (pr_located pr_id) id | VernacResume id -> str"Resume" ++ pr_opt (pr_located pr_id) id | VernacUndo i -> if i=1 then str"Undo" else str"Undo" ++ pr_intarg i @@ -288,172 +387,266 @@ let rec pr_vernac = function (* Syntax *) | VernacGrammar _ -> str"(* : Grammar is replaced by Notation *)" | VernacTacticGrammar l -> hov 1 (str"Grammar tactic simple_tactic :=" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"|") pr_grammar_tactic_rule l) (***) - | VernacSyntax (u,el) -> hov 1 (str"Syntax" ++ str u ++ prlist_with_sep (fun _ -> str";") pr_syntax_entry el) (***) + | VernacSyntax (u,el) -> hov 1 (str"Syntax " ++ str u ++ spc() ++ + prlist_with_sep sep_v2 pr_syntax_entry el) (***) | VernacOpenScope sc -> str"Open Scope" ++ spc() ++ str sc - | VernacDelimiters (sc,key) -> str"Delimits Scope" ++ spc () ++ str sc ++ spc () ++ str key + | VernacDelimiters (sc,key) -> + str"Delimits Scope" ++ spc () ++ str sc ++ + spc() ++ str "with " ++ str key | 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 (a,p,s,q,_,sn) -> (* A Verifier *) - h 0 (str"Infix" ++ pr_entry_prec a ++ pr_intarg p ++ spc() ++ qs s ++ spc() ++ pr_reference q ++ (match sn with + | VernacInfix (a,p,s,q,_,ov8,sn) -> (* A Verifier *) + let (a,p,s) = match ov8 with + Some mv8 -> mv8 + | None -> (a,p,s) in + hov 0 (str"Infix" ++ pr_entry_prec a ++ pr_intarg p ++ spc() ++ qs s ++ spc() ++ pr_reference q ++ (match sn with | None -> mt() | Some sc -> spc() ++ str":" ++ spc() ++ str sc)) - | VernacDistfix (a,p,s,q,sn) -> h 0 (str"Distfix" ++ pr_entry_prec a ++ pr_intarg p ++ spc() ++ qs s ++ spc() ++ pr_reference q ++ (match sn with + | VernacDistfix (a,p,s,q,sn) -> hov 0 (str"Distfix" ++ pr_entry_prec a ++ pr_intarg p ++ spc() ++ qs s ++ spc() ++ pr_reference q ++ (match sn with | None -> mt() | Some sc -> spc() ++ str":" ++ spc() ++ str sc)) - | VernacNotation (s,c,l,opt) -> str"Notation" ++ spc() ++ qs s ++ pr_constrarg c ++ (match l with - | [] -> mt() - | _ as t -> spc() ++ str"(" ++ prlist_with_sep (fun _ -> str",") pr_syntax_modifier t ++ str")") ++ (match opt with - | None -> mt() - | Some sc -> spc() ++ str":" ++ spc() ++ str sc) (***) - | VernacSyntaxExtension (a,b) -> str"Uninterpreted Notation" ++ spc() ++ qs a ++ (match b with | [] -> mt() | _ as l -> str"(" ++ prlist_with_sep (fun _ -> str",") pr_syntax_modifier l ++ str")") + | VernacNotation (s,c,l,mv8,opt) -> + let (s,l) = match mv8 with + None -> (s,l) + | Some ml -> ml in + hov 2( str"Notation" ++ spc() ++ qs s ++ + str " :=" ++ pr_constrarg c ++ + (match l with + | [] -> mt() + | _ as t -> + spc() ++ hov 0 (str"(" ++ prlist_with_sep sep_v2 pr_syntax_modifier t ++ str")")) ++ + (match opt with + | None -> mt() + | Some sc -> str" :" ++ spc() ++ str sc)) + | VernacSyntaxExtension (a,b) -> str"Uninterpreted Notation" ++ spc() ++ qs a ++ (match b with | [] -> mt() | _ as l -> str"(" ++ prlist_with_sep sep_v2 pr_syntax_modifier l ++ str")") (* Gallina *) | VernacDefinition (d,id,b,f,e) -> (* A verifier... *) let pr_def_token = function - | Decl_kinds.LCoercion -> str"Coercion Local" - | Decl_kinds.GCoercion -> str"Coercion" - | Decl_kinds.LDefinition -> str"Local" - | Decl_kinds.GDefinition -> str"Definition" - | Decl_kinds.SCanonical -> str"Canonical Structure" - in let pr_reduce = function - | None -> mt() - | Some r -> str"Eval" ++ spc() ++ pr_red_expr (pr_constr, fun _ -> str"TODO" (* pr_metanum pr_reference *)) r ++ spc() ++ str"in" ++ spc() - in let pr_binders_def = function - | [] -> mt () - | _ as l -> spc() ++ str"(" ++ prlist_with_sep (fun _ -> str") (") (pr_valdecls pr_constr) l ++ str")" - in let pr_valloc (l,c) = match l with - | [] -> mt () - | _ -> prlist_with_sep (fun _ -> str",") (pr_located pr_name) l ++ str":" ++ pr_constr c - in let pr_binders_def2 = function - | [] -> mt () - | _ as l -> spc() ++ str"(" ++ prlist_with_sep (fun _ -> str") (") pr_valloc l ++ str")" - in let pr_def_body = function - | DefineBody (bl,red,c,d) -> (match c with - | CLambdaN (_,bl2,a) -> (pr_binders_def bl ++ pr_binders_def2 bl2, (match d with - | None -> mt() - | Some t -> spc() ++ str":" ++ pr_constrarg t) , Some (pr_reduce red ++ pr_constr a)) - | _ -> (pr_binders_def bl , (match d with - | None -> mt() - | Some t -> spc() ++ str":" ++ pr_constrarg t) , Some (pr_reduce red ++ pr_constr c))) - | ProveBody (bl,t) -> (pr_binders_def bl , pr_constrarg t , None) - in let (binds,typ,c) = pr_def_body b - in h 0 (pr_def_token e ++ spc() ++ pr_id id ++ binds ++ typ ++ (match c with - | None -> mt() - | Some cc -> spc() ++ str":=" ++ spc() ++ cc)) - | VernacStartTheoremProof (ki,id,(bl,c),b,d) -> hov 1 (pr_thm_token ki ++ spc() ++ pr_id id ++ spc() ++ (match bl with | [] -> mt() | _ -> pr_binders pr_constr bl ++ spc()) ++ str":" ++ spc() ++ pr_constr c) + | Decl_kinds.LCoercion -> str"Coercion Local" + | Decl_kinds.GCoercion -> str"Coercion" + | Decl_kinds.LDefinition -> str"Local" + | Decl_kinds.GDefinition -> str"Definition" + | Decl_kinds.SCanonical -> str"Canonical Structure" in + let pr_reduce = function + | None -> mt() + | Some r -> + str"Eval" ++ spc() ++ + pr_red_expr (pr_constr, pr_lconstr, pr_metanum pr_reference) r ++ + str" in" ++ spc() in + let pr_def_body = function + | DefineBody (bl,red,c,d) -> + let (binds,body) = match c with + | CLambdaN (_,bl2,a) when d=None -> + (pr_ne_sep spc (pr_vbinders pr_lconstr) bl ++ + spc() ++ pr_binders bl2, a) + | _ -> (pr_ne_sep spc (pr_vbinders pr_lconstr) bl, c) in + let ty = + match d with + | None -> mt() + | Some t -> spc() ++ str":" ++ pr_lconstrarg t in + let ppred = Some (pr_reduce red ++ pr_lconstr body) in + (binds,ty,ppred) + | ProveBody (bl,t) -> + (pr_vbinders pr_lconstr bl, str" :" ++ pr_lconstrarg t, None) in + let (binds,typ,c) = pr_def_body b in + hov 2 (pr_def_token e ++ spc() ++ pr_id 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_id id ++ spc() ++ + (match bl with + | [] -> mt() + | _ -> pr_vbinders pr_lconstr bl ++ spc()) ++ str":" ++ spc() ++ pr_lconstr c) | VernacEndProof (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_id id | Some tok -> str"Save" ++ spc() ++ pr_thm_token tok ++ spc() ++ pr_id id)) - | VernacExactProof c -> str"Proof" ++ pr_constrarg c - | VernacAssumption (stre,l) -> hov 1 (pr_assumption_token stre ++ spc() ++ pr_ne_params_list pr_constr l) - | VernacInductive (f,l) -> let pr_constructor (coe,(id,c)) = pr_id id ++ spc() ++ (if coe then str":>" else str":") ++ pr_constrarg c in - let pr_constructor_list l = match l with - | [] -> mt() - | _ -> fnl() ++ str"| " ++ prlist_with_sep (fun _ -> fnl() ++ str"| ") pr_constructor l in - let pr_simple_binder (s,t) = pr_id s ++ str":" ++ pr_constr t in - let pr_ne_simple_binders = function - | [] -> mt() - | _ as l -> str"(" ++ prlist_with_sep (fun _ -> str") (") pr_simple_binder l ++ str")" ++ spc() in - let pr_oneind (id,indpar,s,lc) = pr_id id ++ spc() ++ pr_ne_simple_binders indpar ++ str":" ++ spc() ++ pr_constr s ++ spc() ++ str":=" ++ pr_constructor_list lc - in hov 1 ((if f then str"Inductive" else str"CoInductive") ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"with ") pr_oneind l) - | VernacFixpoint recs -> let pr_fixbinder (l,c) = prlist_with_sep (fun _ -> str",") (pr_located pr_name) l ++ str":" ++ pr_constr c in - let pr_fixbinders l = str"[" ++ prlist_with_sep (fun _ -> str";") pr_fixbinder l ++ str"]" in - let pr_onerec = function - | (id,_,CProdN(_,bl,type_),CLambdaN(_,_,def)) -> pr_id id ++ spc() ++ pr_fixbinders bl ++ spc() ++ str":" ++ spc() ++ pr_constr type_ ++ spc() ++ str":=" ++ spc() ++ pr_constr def - | _ -> mt() - in hov 1 (str"Fixpoint" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"with ") pr_onerec recs) - | VernacCoFixpoint corecs -> let pr_onecorec (id,c,def) = pr_id id ++ spc() ++ str":" ++ pr_constrarg c ++ spc() ++ str":=" ++ pr_constrarg def - in hov 1 (str"CoFixpoint" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"with ") pr_onecorec corecs) - | VernacScheme l -> hov 1 (str"Scheme" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"with") pr_onescheme l) + | VernacExactProof c -> hov 2 (str"Proof" ++ pr_lconstrarg c) + | VernacAssumption (stre,l) -> + hov 2 + (pr_assumption_token stre ++ spc() ++ pr_ne_params_list pr_lconstr l) + | VernacInductive (f,l) -> + let pr_constructor (coe,(id,c)) = + pr_id id ++ spc() ++ (if coe then str":>" else str":") + ++ pr_lconstrarg c in + let pr_constructor_list l = match l with + | [] -> mt() + | _ -> + fnl() ++ str" | " ++ + prlist_with_sep (fun _ -> fnl() ++ str" | ") pr_constructor l in + let pr_oneind (id,indpar,s,lc) = + pr_id id ++ spc() ++ pr_sbinders indpar ++ str":" ++ spc() ++ + pr_lconstr s ++ str" :=" ++ pr_constructor_list lc in + hov 1 + ((if f then str"Inductive" else str"CoInductive") ++ spc() ++ + prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_oneind l) + | VernacFixpoint recs -> + let pr_onerec = function + | (id,_,CProdN(_,bl,type_),CLambdaN(_,_,def)) -> + pr_id id ++ spc() ++ pr_binders bl ++ spc() + ++ pr_type_option (fun c -> spc() ++ pr_lconstr c) type_ + ++ str" :=" ++ brk(1,1) ++ pr_lconstr def + | _ -> mt() in + hov 1 (str"Fixpoint" ++ spc() ++ + prlist_with_sep (fun _ -> fnl() ++ str"with ") pr_onerec recs) + | VernacCoFixpoint corecs -> + let pr_onecorec (id,c,def) = + pr_id id ++ spc() ++ str":" ++ pr_lconstrarg c ++ + str" :=" ++ brk(1,1) ++ pr_lconstr def in + hov 1 (str"CoFixpoint" ++ 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 ((oc,name),ps,s,c,fs) -> let pr_simple_binder (s,t) = pr_id s ++ str":" ++ pr_constr t in - let pr_record_field = function - | (oc,AssumExpr (id,t)) -> pr_id id ++ spc() ++ if oc then str":>" else str":" ++ spc() ++ pr_constr t - | (oc,DefExpr(id,b,opt)) -> (match opt with - | Some t -> pr_id id ++ spc() ++ if oc then str":>" else str":" ++ spc() ++ pr_constr t ++ spc() ++ str":=" ++ pr_constr b - | None -> pr_id id ++ spc() ++ str":=" ++ pr_constr b) in - hov 1 (str"Record" ++ if oc then str" > " else spc() ++ pr_id name ++ spc() ++ (match ps with - | [] -> mt() - | _ as l -> str"[" ++ prlist_with_sep (fun _ -> str";") pr_simple_binder l ++ str"]") ++ spc() ++ str":" ++ spc() ++ (* pr_sort s *) str"TODO" ++ spc() ++ str":=" ++ (match c with - | None -> mt() - | Some sc -> pr_id sc) ++ spc() ++ str"{" ++ cut() ++ prlist_with_sep (fun _ -> str";" ++ brk(1,1)) pr_record_field fs ++ str"}") - | VernacBeginSection id -> h 0 (str"Section" ++ spc () ++ pr_id id) - | VernacEndSegment id -> h 0 (str"End" ++ spc() ++ pr_id id) - | VernacRequire (exp,spe,l) -> h 0 ((match exp with - | None -> str"Read Module" - | Some export -> str"Require" ++ if export then spc() ++ str"Export" else mt ()) ++ spc() ++ + | VernacRecord ((oc,name),ps,s,c,fs) -> + let pr_record_field = function + | (oc,AssumExpr (id,t)) -> + hov 1 (str"(" ++ pr_id id ++ + (if oc then str" :>" else str" :") ++ spc() ++ + pr_lconstr t ++ str ")") + | (oc,DefExpr(id,b,opt)) -> (match opt with + | Some t -> + hov 1 (str "(" ++ pr_id id ++ + (if oc then str" :>" else str" :") ++ spc() ++ + pr_lconstr t ++ str" :=" ++ pr_lconstr b ++ str ")") + | None -> + hov 1 (str "(" ++ pr_id id ++ str" :=" ++ spc() ++ + pr_lconstr b ++ str ")")) in + hov 2 + (str"Record" ++ + (if oc then str" > " else str" ") ++ pr_id name ++ spc() ++ + pr_sbinders ps ++ str" :" ++ spc() ++ pr_lconstr s ++ str" :=" ++ + (match c with + | None -> mt() + | Some sc -> pr_id sc) ++ spc() ++ str"{" ++ cut() ++ + hv 0 (prlist_with_sep fnl pr_record_field fs) + ++ str"}") + | VernacBeginSection id -> hov 2 (str"Section" ++ spc () ++ pr_id id) + | VernacEndSegment id -> hov 2 (str"End" ++ spc() ++ pr_id id) + | VernacRequire (exp,spe,l) -> hov 2 + (str "Require " ++ pr_require_token exp ++ spc() ++ (match spe with | None -> mt() - | Some flag -> (if flag then str"Specification" else str"Implementation") ++ spc ()) ++ - prlist_with_sep sep pr_reference l) - | VernacImport (f,l) -> if f then str"Export" else str"Import" ++ spc() ++ prlist_with_sep sep pr_reference l + | Some flag -> + (if flag then str"Specification" else str"Implementation") ++ + spc ()) ++ + prlist_with_sep sep pr_reference l) + | VernacImport (f,l) -> + (if f then str"Export" else str"Import") ++ spc() ++ + prlist_with_sep sep pr_reference l | VernacCanonical q -> str"Canonical Structure" ++ spc() ++ pr_reference q | VernacCoercion (s,id,c1,c2) -> hov 1 (str"Coercion" ++ (match s with | Decl_kinds.Local -> spc() ++ str"Local" ++ spc() | Decl_kinds.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 | Decl_kinds.Local -> spc() ++ str"Local" ++ spc() | Decl_kinds.Global -> spc()) ++ pr_id id ++ spc() ++ str":" ++ spc() ++ pr_class_rawexpr c1 ++ spc() ++ str">->" ++ spc() ++ pr_class_rawexpr c2) (* Modules and Module Types *) - | VernacDeclareModule (id,l,m1,m2) -> hov 1 (str"Module" ++ spc() ++ pr_id id ++ spc() ++ pr_module_binders_list l pr_constr ++ (match m1 with + | VernacDeclareModule (id,l,m1,m2) -> hov 1 (str"Module" ++ spc() ++ pr_id id ++ spc() ++ pr_module_binders_list l pr_lconstr ++ (match m1 with | None -> mt() | Some n1 -> str" : " ++ str"TODO" (* pr_module_type pr_constr n1 *) ) ++ (match m2 with | None -> mt() | Some n2 -> str" := " ++ pr_module_expr n2)) - | VernacDeclareModuleType (id,l,m) -> hov 1 (str"Module Type" ++ spc() ++ pr_id id ++ spc() ++ pr_module_binders_list l pr_constr ++ (match m with + | VernacDeclareModuleType (id,l,m) -> hov 1 (str"Module Type" ++ spc() ++ pr_id id ++ spc() ++ pr_module_binders_list l pr_lconstr ++ (match m with | None -> mt() - | Some n -> str" := " ++ pr_module_type pr_constr n)) + | Some n -> str" := " ++ pr_module_type pr_lconstr n)) (* Solving *) - | VernacSolve (i,tac,_) -> pr_raw_tactic tac (* A Verifier *) - | VernacSolveExistential (i,c) -> str"Existential" ++ spc() ++ int i ++ pr_constrarg c + | VernacSolve (i,tac,deftac) -> + (if i = 1 then mt() else int i ++ str ": ") ++ + (if deftac then mt() else str "!! ") ++ + pr_raw_tactic_goal i tac + | VernacSolveExistential (i,c) -> + str"Existential " ++ int i ++ pr_lconstrarg c (* Auxiliary file and library management *) - | VernacRequireFrom (exp,spe,f) -> h 0 (str"Require" ++ if exp then spc() ++ str"Export" ++ spc() else spc() ++ (match spe with - | None -> mt() - | Some false -> str"Implementation" ++ spc() - | Some true -> str"Specification" ++ spc ()) ++ qs f) - | VernacAddLoadPath (fl,s,d) -> hov 1 (str"Add" ++ if fl then str" Rec " else spc() ++ str"LoadPath" ++ spc() ++ qs s ++ (match d with - | None -> mt() - | Some dir -> spc() ++ str"as" ++ spc() ++ pr_dirpath dir)) + | VernacRequireFrom (exp,spe,f) -> hov 2 + (str"Require " ++ pr_require_token exp ++ spc() ++ + (match spe with + | None -> mt() + | Some false -> str"Implementation" ++ spc() + | Some true -> str"Specification" ++ spc ()) ++ + qs f) + | VernacAddLoadPath (fl,s,d) -> hov 2 + (str"Add" ++ + (if fl then str" Rec " else spc()) ++ + str"LoadPath" ++ spc() ++ qs s ++ + (match d with + | None -> mt() + | Some dir -> spc() ++ str"as" ++ spc() ++ pr_dirpath dir)) | VernacRemoveLoadPath s -> str"Remove LoadPath" ++ qs s - | VernacAddMLPath (fl,s) -> str"Add" ++ if fl then str" Rec " else spc() ++ str"ML Path" ++ qs s - | VernacDeclareMLModule l -> hov 1 (str"Declare ML Module" ++ spc() ++ prlist_with_sep sep qs l) + | VernacAddMLPath (fl,s) -> + str"Add" ++ (if fl then str" Rec " else spc()) ++ str"ML Path" ++ qs s + | VernacDeclareMLModule l -> + hov 2 (str"Declare ML Module" ++ spc() ++ prlist_with_sep sep qs l) | VernacChdir s -> str"Cd" ++ pr_opt qs s (* Commands *) - | VernacDeclareTacticDefinition (_,l) -> (match l with - | [] -> mt() - | [(id,b)] -> hov 1 (str"Tactic Definition" ++ spc() ++ pr_located pr_id id ++ spc() ++ str":=" ++ spc() ++ pr_raw_tactic b) - | _ as t -> let pr_vrec_clause (id,b) = pr_located pr_id id ++ spc() ++ str":=" ++ spc() ++ pr_raw_tactic b in hov 1 (str"Recursive Tactic Definition" ++ spc() ++ prlist_with_sep (fun _ -> brk(1,1) ++ str"And") pr_vrec_clause t)) + | VernacDeclareTacticDefinition (rc,l) -> + let pr_tac_body (id, body) = + let ppb, body = + match body with + Tacexpr.TacFun (idl,b) -> + (spc() ++ prlist_with_sep spc + (function None -> str"_" | Some id -> pr_id id) + idl), + b + | _ -> mt(), body in + pr_located pr_id id ++ ppb ++ str" :=" ++ brk(1,1) ++ + pr_raw_tactic body in + hov 1 + ((if rc then str "Recursive " else mt()) ++ + str "Tactic Definition " ++ + prlist_with_sep (fun () -> fnl() ++ str"with ") pr_tac_body l) | VernacHints (dbnames,h) -> pr_hints dbnames h pr_constr - | VernacHintDestruct (id,loc,c,i,tac) -> hov 1 (str"HintDestruct" ++ spc() ++ pr_destruct_location loc ++ spc() ++ pr_id id ++ pr_constrarg c ++ pr_intarg i ++ spc() ++ str"[" ++ pr_raw_tactic tac ++ str"]") - | VernacSyntacticDefinition (id,c,None) -> hov 1 (str"Syntactic Definition" ++ spc() ++ pr_id id ++ spc() ++ str":=" ++ pr_constrarg c) - | VernacSyntacticDefinition (id,c,Some n) -> hov 1 (str"Syntactic Definition" ++ spc() ++ pr_id id ++ spc() ++ str":=" ++ pr_constrarg c ++ spc() ++ str"|" ++ int n) - | VernacDeclareImplicits (q,None) -> hov 1 (str"Implicits" ++ spc() ++ pr_reference q) - | VernacDeclareImplicits (q,Some l) -> hov 1 (str"Implicits" ++ spc() ++ pr_reference q ++ spc() ++ str"[" ++ prlist_with_sep sep int l ++ str"]") - | 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) + | VernacHintDestruct (id,loc,c,i,tac) -> + hov 2 (str"HintDestruct " ++ pr_destruct_location loc ++ spc() ++ + pr_id id ++ pr_constrarg c ++ pr_intarg i ++ spc() ++ + str"[" ++ pr_raw_tactic tac ++ str"]") + | VernacSyntacticDefinition (id,c,None) -> + hov 2 (str"Syntactic Definition " ++ pr_id id ++ str" :=" ++ + pr_lconstrarg c) + | VernacSyntacticDefinition (id,c,Some n) -> + hov 2 (str"Syntactic Definition " ++ pr_id id ++ str" :=" ++ + pr_lconstrarg c ++ spc() ++ str"|" ++ int n) + | VernacDeclareImplicits (q,None) -> + hov 2 (str"Implicits" ++ spc() ++ pr_reference q) + | VernacDeclareImplicits (q,Some l) -> + hov 1 (str"Implicits" ++ spc() ++ pr_reference q ++ spc() ++ + str"[" ++ prlist_with_sep sep int l ++ str"]") + | 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 (Goptions.SecondaryTable ("Implicit","Arguments"),BoolValue true) -> str"Set Implicit Arguments" | VernacSetOption (Goptions.SecondaryTable ("Implicit","Arguments"),BoolValue false) -> str"Unset Implicit Arguments" - | VernacSetOption (na,v) -> hov 1 (str"Set" ++ spc() ++ pr_set_option na v) - | VernacAddOption (na,l) -> hov 1 (str"Add" ++ spc() ++ pr_printoption na (Some l)) - | VernacRemoveOption (na,l) -> hov 1 (str"Remove" ++ spc() ++ pr_printoption na (Some l)) - | VernacMemOption (na,l) -> hov 1 (str"Test" ++ spc() ++ pr_printoption na (Some l)) - | VernacPrintOption na -> hov 1 (str"Test" ++ 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 -> h 0 (str"Eval" ++ spc() ++ pr_red_expr (pr_constr,fun _ -> str"TODO" (*pr_metanum pr_reference *)) r0 ++ spc() ++ str"in" ++ spc () ++ pr_constr c) - | None -> h 0 (str"Check" ++ spc() ++ pr_constr c) + | Some r0 -> + hov 2 (str"Eval" ++ spc() ++ + pr_red_expr (pr_constr,pr_lconstr,pr_metanum pr_reference) r0 ++ + spc() ++ str"in" ++ spc () ++ pr_lconstr c) + | None -> hov 2 (str"Check" ++ spc() ++ pr_lconstr c) in pr_mayeval r c - | VernacGlobalCheck c -> hov 1 (str"Type" ++ pr_constrarg c) + | VernacGlobalCheck c -> hov 2 (str"Type" ++ pr_lconstrarg c) | VernacPrint p -> let pr_printable = function | PrintFullContext -> str"Print All" - | PrintSectionContext s -> str"Print Section" ++ spc() ++ pr_reference s - | PrintGrammar (uni,ent) -> str"Print Grammar" ++ spc() ++ str uni ++ spc() ++ str ent + | PrintSectionContext s -> + str"Print Section" ++ spc() ++ pr_reference s + | PrintGrammar (uni,ent) -> + str"Print Grammar" ++ spc() ++ str uni ++ spc() ++ str ent | PrintLoadPath -> str"Print LoadPath" | PrintModules -> str"Print Modules" | PrintMLLoadPath -> str"Print ML Path" @@ -474,7 +667,7 @@ let rec pr_vernac = function | 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 - | PrintScope s -> str"Scope" ++ spc() ++ str s + | PrintScope s -> str"Print Scope" ++ spc() ++ str s in pr_printable p | VernacSearch (sea,sea_r) -> pr_search sea sea_r pr_constr | VernacLocate loc -> @@ -484,17 +677,48 @@ let rec pr_vernac = function | LocateLibrary qid -> str"Library" ++ spc () ++ pr_reference qid | LocateNotation s -> str ("\""^s^"\"") in str"Locate" ++ spc() ++ pr_locate loc - | VernacComments l -> hov 1 (str"Comments" ++ spc() ++ prlist_with_sep sep (pr_comment pr_constr) l) + | VernacComments l -> + hov 2 + (str"Comments" ++ spc() ++ prlist_with_sep sep (pr_comment pr_constr) l) | VernacNop -> str"Proof" (* Toplevel control *) | VernacToplevelControl exn -> pr_topcmd exn (* For extension *) - | VernacExtend (s,c) -> hov 1 (str s ++ prlist_with_sep sep pr_constrarg (List.map (Genarg.out_gen Genarg.rawwit_constr) c)) - | VernacV7only _ | VernacV8only _ -> mt () - | VernacProof _ | VernacDefineModule _ -> str"TODO" + | VernacExtend (s,c) -> pr_extend s c + | VernacV7only _ -> mt() + | VernacV8only com -> pr_vernac com + | VernacProof te -> str "Proof with" ++ spc() ++ pr_raw_tactic te + | VernacDefineModule (m,bl,ty,bd) -> + hov 2 (str"Module " ++ pr_id m ++ + pr_module_binders bl pr_lconstr ++ + (match ty with + None -> mt() + | Some(t,_) -> pr_module_type pr_lconstr t) ++ + pr_opt pr_module_expr bd) + +and pr_extend s cl = + let pr_arg a = + try Pptacticnew.pr_gen (Global.env()) a + with Failure _ -> str ("") in + try + let rls = List.assoc s (Egrammar.get_extend_vernac_grammars()) in + let (hd,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 ++ Pptacticnew.pr_gen (Global.env()) (List.hd args), + List.tl args) + | Egrammar.TacTerm s -> (strm ++ spc() ++ str s, args)) + (str hd,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 = make_pr_vernac (Ppconstrnew.pr_constr) +let pr_vernac = make_pr_vernac Ppconstrnew.pr_constr Ppconstrnew.pr_lconstr -- cgit v1.2.3