From f5c81ac1bcae7c40d5e89229647c06c97b5ddc85 Mon Sep 17 00:00:00 2001 From: Thomas Sibut-Pinote Date: Wed, 1 Feb 2017 17:58:32 +0100 Subject: Added some theory on powerRZ. For this, I used a new inductive type Z_spec to reason on Z. --- theories/Reals/Rfunctions.v | 98 +++++++++++++++++++++++++++++++++++++++++++++ theories/Reals/Rpower.v | 14 +++++++ 2 files changed, 112 insertions(+) diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v index 0a49d4983..80dc36e87 100644 --- a/theories/Reals/Rfunctions.v +++ b/theories/Reals/Rfunctions.v @@ -531,6 +531,36 @@ Qed. (*******************************) (*i Due to L.Thery i*) +Section PowerRZ. + +Local Coercion Z_of_nat : nat >-> Z. + +(* the following section should probably be somewhere else, but not sure where *) +Section Z_compl. + +Local Open Scope Z_scope. + +(* Provides a way to reason directly on Z in terms of nats instead of positive *) +Inductive Z_spec (x : Z) : Z -> Type := +| ZintNull : x = 0 -> Z_spec x 0 +| ZintPos (n : nat) : x = n -> Z_spec x n +| ZintNeg (n : nat) : x = - n -> Z_spec x (- n). + +Lemma intP (x : Z) : Z_spec x x. +Proof. + destruct x as [|p|p]. + - now apply ZintNull. + - rewrite <-positive_nat_Z at 2. + apply ZintPos. + now rewrite positive_nat_Z. + - rewrite <-Pos2Z.opp_pos. + rewrite <-positive_nat_Z at 2. + apply ZintNeg. + now rewrite positive_nat_Z. +Qed. + +End Z_compl. + Definition powerRZ (x:R) (n:Z) := match n with | Z0 => 1 @@ -657,6 +687,74 @@ Proof. auto with real. Qed. +Local Open Scope Z_scope. + +Lemma pow_powerRZ (r : R) (n : nat) : + (r ^ n)%R = powerRZ r (Z_of_nat n). +Proof. + induction n; [easy|simpl]. + now rewrite SuccNat2Pos.id_succ. +Qed. + +Lemma powerRZ_ind (P : Z -> R -> R -> Prop) : + (forall x, P 0 x 1%R) -> + (forall x n, P (Z.of_nat n) x (x ^ n)%R) -> + (forall x n, P ((-(Z.of_nat n))%Z) x (Rinv (x ^ n))) -> + forall x (m : Z), P m x (powerRZ x m)%R. +Proof. + intros ? ? ? x m. + destruct (intP m) as [n Hm| n Hm|n Hm]. + - easy. + - now rewrite <- pow_powerRZ. + - unfold powerRZ. + destruct n as [|n]; [ easy |]. + rewrite Nat2Z.inj_succ, <- Zpos_P_of_succ_nat, Pos2Z.opp_pos. + now rewrite <- Pos2Z.opp_pos, <- positive_nat_Z. +Qed. + +Lemma powerRZ_inv x alpha : (x <> 0)%R -> powerRZ (/ x) alpha = Rinv (powerRZ x alpha). +Proof. + intros; destruct (intP alpha). + - now simpl; rewrite Rinv_1. + - now rewrite <-!pow_powerRZ, ?Rinv_pow, ?pow_powerRZ. + - unfold powerRZ. + destruct (- n). + + now rewrite Rinv_1. + + now rewrite Rinv_pow. + + now rewrite <-Rinv_pow. +Qed. + +Lemma powerRZ_neg x : forall alpha, x <> R0 -> powerRZ x (- alpha) = powerRZ (/ x) alpha. +Proof. + intros [|n|n] H ; simpl. + - easy. + - now rewrite Rinv_pow. + - rewrite Rinv_pow by now apply Rinv_neq_0_compat. + now rewrite Rinv_involutive. +Qed. + +Lemma powerRZ_mult_distr : + forall m x y, ((0 <= m)%Z \/ (x * y <> 0)%R) -> + (powerRZ (x*y) m = powerRZ x m * powerRZ y m)%R. +Proof. + intros m x0 y0 Hmxy. + destruct (intP m) as [ | | n Hm ]. + - now simpl; rewrite Rmult_1_l. + - now rewrite <- !pow_powerRZ, Rpow_mult_distr. + - destruct Hmxy as [H|H]. + + assert(m = 0) as -> by now omega. + now rewrite <- Hm, Rmult_1_l. + + assert(x0 <> 0)%R by now intros ->; apply H; rewrite Rmult_0_l. + assert(y0 <> 0)%R by now intros ->; apply H; rewrite Rmult_0_r. + rewrite !powerRZ_neg by assumption. + rewrite Rinv_mult_distr by assumption. + now rewrite <- !pow_powerRZ, Rpow_mult_distr. +Qed. + +End PowerRZ. + +Local Infix "^Z" := powerRZ (at level 30, right associativity) : R_scope. + (*******************************) (* For easy interface *) (*******************************) diff --git a/theories/Reals/Rpower.v b/theories/Reals/Rpower.v index b3ce6fa33..93462a1fc 100644 --- a/theories/Reals/Rpower.v +++ b/theories/Reals/Rpower.v @@ -490,6 +490,20 @@ Proof. apply exp_Ropp. Qed. +Lemma powerRZ_Rpower x z : (0 < x)%R -> powerRZ x z = Rpower x (IZR z). +Proof. + intros Hx. + assert (x <> 0)%R + by now intros Habs; rewrite Habs in Hx; apply (Rlt_irrefl 0). + destruct (intP z). + - now rewrite Rpower_O. + - rewrite <- pow_powerRZ, <- Rpower_pow by assumption. + now rewrite INR_IZR_INZ. + - rewrite opp_IZR, Rpower_Ropp. + rewrite powerRZ_neg, powerRZ_inv by assumption. + now rewrite <- pow_powerRZ, <- INR_IZR_INZ, Rpower_pow. +Qed. + Theorem Rle_Rpower : forall e n m:R, 1 < e -> 0 <= n -> n <= m -> e ^R n <= e ^R m. Proof. -- cgit v1.2.3 From 1a43fda0dc9bb8d100808426980446353f8f1ae3 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 11:21:23 +0200 Subject: Removing internal support for accepting "{struct x}" and co in "Theorem with". There were actually no syntax for it, and I'm still unsure what good syntax to give to it, even more that it would be useful to have one. --- ide/texmacspp.ml | 2 +- intf/vernacexpr.mli | 2 +- parsing/g_vernac.ml4 | 4 ++-- printing/ppvernac.ml | 3 +-- vernac/lemmas.ml | 25 ++++--------------------- vernac/vernacentries.ml | 4 ++-- 6 files changed, 11 insertions(+), 29 deletions(-) diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml index e787e48bf..74d0438dd 100644 --- a/ide/texmacspp.ml +++ b/ide/texmacspp.ml @@ -553,7 +553,7 @@ let rec tmpp v loc = let str_dk = Kindops.string_of_definition_kind (l, false, dk) in let str_id = Id.to_string id in (xmlDef str_dk str_id loc [pp_expr e]) - | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement, None) ], b) -> + | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement) ], b) -> let str_tk = Kindops.string_of_theorem_kind tk in let str_id = Id.to_string id in (xmlThm str_tk str_id loc [pp_expr statement]) diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.mli index 25d3c705f..d62c639c2 100644 --- a/intf/vernacexpr.mli +++ b/intf/vernacexpr.mli @@ -209,7 +209,7 @@ type one_inductive_expr = plident * local_binder_expr list * constr_expr option * constructor_expr list type proof_expr = - plident option * (local_binder_expr list * constr_expr * (lident option * recursion_order_expr) option) + plident option * (local_binder_expr list * constr_expr) type syntax_modifier = | SetItemLevel of string list * Extend.production_level diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 71b93439b..6ede0490e 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -146,8 +146,8 @@ GEXTEND Gram [ [ thm = thm_token; id = pidentref; bl = binders; ":"; c = lconstr; l = LIST0 [ "with"; id = pidentref; bl = binders; ":"; c = lconstr -> - (Some id,(bl,c,None)) ] -> - VernacStartTheoremProof (thm, (Some id,(bl,c,None))::l, false) + (Some id,(bl,c)) ] -> + VernacStartTheoremProof (thm, (Some id,(bl,c))::l, false) | stre = assumption_token; nl = inline; bl = assum_list -> VernacAssumption (stre, nl, bl) | (kwd,stre) = assumptions_token; nl = inline; bl = assum_list -> diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index cfc2e48d1..d3b4007fb 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -386,13 +386,12 @@ open Decl_kinds ++ pr_opt (fun def -> str":=" ++ brk(1,2) ++ pr_pure_lconstr def) def ++ prlist (pr_decl_notation pr_constr) ntn - let pr_statement head (idpl,(bl,c,guard)) = + let pr_statement head (idpl,(bl,c)) = assert (not (Option.is_empty idpl)); let id, pl = Option.get idpl in hov 2 (head ++ spc() ++ pr_lident id ++ pr_univs pl ++ spc() ++ (match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++ - pr_opt (pr_guard_annot pr_lconstr_expr bl) guard ++ str":" ++ pr_spc_lconstr c) let pr_priority = function diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 409676276..e7d1919ce 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -88,25 +88,9 @@ let adjust_guardness_conditions const = function let find_mutually_recursive_statements thms = let n = List.length thms in - let inds = List.map (fun (id,(t,impls,annot)) -> + let inds = List.map (fun (id,(t,impls)) -> let (hyps,ccl) = decompose_prod_assum t in let x = (id,(t,impls)) in - match annot with - (* Explicit fixpoint decreasing argument is given *) - | Some (Some (_,id),CStructRec) -> - let i,b,typ = lookup_rel_id id hyps in - (match kind_of_term t with - | Ind ((kn,_ as ind), u) when - let mind = Global.lookup_mind kn in - mind.mind_finite == Decl_kinds.Finite && Option.is_empty b -> - [ind,x,i],[] - | _ -> - error "Decreasing argument is not an inductive assumption.") - (* Unsupported cases *) - | Some (_,(CWfRec _|CMeasureRec _)) -> - error "Only structural decreasing is supported for mutual statements." - (* Cofixpoint or fixpoint w/o explicit decreasing argument *) - | None | Some (None, CStructRec) -> let whnf_hyp_hds = map_rel_context_in_env (fun env c -> fst (whd_all_stack env Evd.empty c)) (Global.env()) hyps in @@ -178,7 +162,7 @@ let find_mutually_recursive_statements thms = (finite,guard,None), ordered_inds let look_for_possibly_mutual_statements = function - | [id,(t,impls,None)] -> + | [id,(t,impls)] -> (* One non recursively proved theorem *) None,[id,(t,impls)],None | _::_ as thms -> @@ -458,7 +442,7 @@ let start_proof_com ?inference_hook kind thms hook = | None -> Evd.from_env env0 | Some l -> Evd.from_ctx (Evd.make_evar_universe_context env0 l)) in - let thms = List.map (fun (sopt,(bl,t,guard)) -> + let thms = List.map (fun (sopt,(bl,t)) -> let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in let t', imps' = interp_type_evars_impls ~impls env evdref t in let flags = all_and_fail_flags in @@ -467,8 +451,7 @@ let start_proof_com ?inference_hook kind thms hook = let ids = List.map RelDecl.get_name ctx in (compute_proof_name (pi1 kind) sopt, (nf_evar !evdref (it_mkProd_or_LetIn t' ctx), - (ids, imps @ lift_implicits (List.length ids) imps'), - guard))) + (ids, imps @ lift_implicits (List.length ids) imps')))) thms in let recguard,thms,snl = look_for_possibly_mutual_statements thms in let evd, nf = Evarutil.nf_evars_and_universes !evdref in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index ca03ba3f3..07066d983 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -489,7 +489,7 @@ let vernac_definition locality p (local,k) ((loc,id as lid),pl) def = (match def with | ProveBody (bl,t) -> (* local binders, typ *) start_proof_and_print (local,p,DefinitionBody k) - [Some (lid,pl), (bl,t,None)] hook + [Some (lid,pl), (bl,t)] hook | DefineBody (bl,red_option,c,typ_opt) -> let red_option = match red_option with | None -> None @@ -2077,7 +2077,7 @@ let interp ?proof ~loc locality poly c = | VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n") (* Proof management *) - | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t,None)] false + | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t)] false | VernacFocus n -> vernac_focus n | VernacUnfocus -> vernac_unfocus () | VernacUnfocused -> vernac_unfocused () -- cgit v1.2.3 From 70930579d8454a4ff50ee0ea1e97a55863bf01f6 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 15:53:14 +0200 Subject: More explicit message when a {struct x} argument refers to a local definition. --- interp/topconstr.ml | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/interp/topconstr.ml b/interp/topconstr.ml index 89e04b69d..513fa8f16 100644 --- a/interp/topconstr.ml +++ b/interp/topconstr.ml @@ -176,7 +176,12 @@ let split_at_annot bl na = in (List.rev ans, CLocalAssum (r, k, t) :: rest) end - | CLocalDef _ as x :: rest -> aux (x :: acc) rest + | CLocalDef ((_,na),_,_) as x :: rest -> + if Name.equal (Name id) na then + user_err ~loc + (Nameops.pr_id id ++ str" must be a proper parameter and not a local definition.") + else + aux (x :: acc) rest | CLocalPattern (loc,_,_) :: rest -> Loc.raise ~loc (Stream.Error "pattern with quote not allowed after fix") | [] -> -- cgit v1.2.3 From 8d69323dff6e1afd17a13dfa8a980e76acb44cdc Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 15:53:35 +0200 Subject: Documenting how the recursive indices of a fixpoint are computed. Also documenting how the implicit arguments by position are computed. --- intf/constrexpr.mli | 2 +- kernel/constr.ml | 9 +++++++++ 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index 49bafadc8..a4a6eb909 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -19,7 +19,7 @@ open Decl_kinds type notation = string type explicitation = - | ExplByPos of int * Id.t option + | ExplByPos of int * Id.t option (* a reference to the n-th product starting from left *) | ExplByName of Id.t type binder_kind = diff --git a/kernel/constr.ml b/kernel/constr.ml index ce20751ab..d0c580d13 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -107,7 +107,16 @@ type constr = t type existential = existential_key * constr array type rec_declaration = Name.t array * constr array * constr array type fixpoint = (int array * int) * rec_declaration + (* The array of [int]'s tells for each component of the array of + mutual fixpoints the number of lambdas to skip before finding the + recursive argument (e.g., value is 2 in "fix f (x:A) (y:=t) (z:B) + (v:=u) (w:I) {struct w}"), telling to skip x and z and that w is + the recursive argument); + The second component [int] tells which component of the block is + returned *) type cofixpoint = int * rec_declaration + (* The component [int] tells which component of the block of + cofixpoint is returned *) type types = constr -- cgit v1.2.3 From 8a35d93061c67dcdbb12337b78fcb35d72957f51 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 19:13:33 +0200 Subject: Minor cosmetic commit. --- vernac/command.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/vernac/command.ml b/vernac/command.ml index 6eb7037f8..f31fce885 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -1155,7 +1155,7 @@ let interp_fixpoint l ntns = let interp_cofixpoint l ntns = let (env,_,pl,evd),fix,info = interp_recursive false l ntns in - check_recursive false env evd fix; + check_recursive false env evd fix; (fix,pl,Evd.evar_universe_context evd,info) let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns = -- cgit v1.2.3 From 21d308fc083ee395e7b1ff78e75ec3a1305cb2f4 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 19:12:44 +0200 Subject: Fixing several wrong computations of implicit arguments by position in the presence of let-ins. --- test-suite/success/ImplicitArguments.v | 6 ++++++ vernac/command.ml | 6 +++--- vernac/lemmas.ml | 2 +- 3 files changed, 10 insertions(+), 4 deletions(-) diff --git a/test-suite/success/ImplicitArguments.v b/test-suite/success/ImplicitArguments.v index f702aa62f..f07773f8b 100644 --- a/test-suite/success/ImplicitArguments.v +++ b/test-suite/success/ImplicitArguments.v @@ -21,3 +21,9 @@ Fixpoint app {A : Type} {n m : nat} (v : vector A n) (w : vector A m) : vector A (* Test sharing information between different hypotheses *) Parameters (a:_) (b:a=0). + +(* These examples were failing due to a lifting wrongly taking let-in into account *) + +Definition foo6 (x:=1) : forall {n:nat}, n=n := fun n => eq_refl. + +Fixpoint foo7 (x:=1) (n:nat) {p:nat} {struct n} : nat. diff --git a/vernac/command.ml b/vernac/command.ml index f31fce885..c24dbdf7c 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -95,7 +95,7 @@ let interp_definition pl bl p red_option c ctypopt = let ctx = Evd.make_evar_universe_context env pl in let evdref = ref (Evd.from_ctx ctx) in let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in - let nb_args = List.length ctx in + let nb_args = Context.Rel.nhyps ctx in let imps,pl,ce = match ctypopt with None -> @@ -838,7 +838,7 @@ type structured_fixpoint_expr = { let interp_fix_context env evdref isfix fix = let before, after = if isfix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in let impl_env, ((env', ctx), imps) = interp_context_evars env evdref before in - let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(List.length before) env' evdref after in + let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(Context.Rel.nhyps ctx) env' evdref after in let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot) @@ -1100,7 +1100,7 @@ let interp_recursive isfix fixl notations = let fixtypes = List.map2 build_fix_type fixctxs fixccls in let fixtypes = List.map (nf_evar !evdref) fixtypes in let fiximps = List.map3 - (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps)) + (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (Context.Rel.nhyps ctx) cclimps)) fixctximps fixcclimps fixctxs in let rec_sign = List.fold_left2 diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index e7d1919ce..430d48bc7 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -451,7 +451,7 @@ let start_proof_com ?inference_hook kind thms hook = let ids = List.map RelDecl.get_name ctx in (compute_proof_name (pi1 kind) sopt, (nf_evar !evdref (it_mkProd_or_LetIn t' ctx), - (ids, imps @ lift_implicits (List.length ids) imps')))) + (ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps')))) thms in let recguard,thms,snl = look_for_possibly_mutual_statements thms in let evd, nf = Evarutil.nf_evars_and_universes !evdref in -- cgit v1.2.3 From 2ddc9d12bd4616f10245c40bc0c87ae548911809 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Apr 2017 19:13:58 +0200 Subject: Fixing #5420 as well as many related bugs due to miscounting let-ins. - Supporting let-ins in tactic "fix", and hence in interactive Fixpoint and mutual theorems. - Documenting more precisely the meaning of n in tactic "fix id n". - Fixing computation of recursive index at interpretation time in the presence of let-ins. --- doc/refman/RefMan-tac.tex | 6 +++--- tactics/tactics.ml | 3 +++ vernac/command.ml | 16 +++++++++------- vernac/command.mli | 8 ++++---- vernac/lemmas.ml | 4 ++-- vernac/lemmas.mli | 4 ++-- 6 files changed, 23 insertions(+), 18 deletions(-) diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex index 3f1241186..0edc66f83 100644 --- a/doc/refman/RefMan-tac.tex +++ b/doc/refman/RefMan-tac.tex @@ -2618,9 +2618,9 @@ as the ones described in Section~\ref{Tac-induction}. In the syntax of the tactic, the identifier {\ident} is the name given to the induction hypothesis. The natural number {\num} tells on which premise of the current goal the induction acts, starting -from 1 and counting both dependent and non dependent -products. Especially, the current lemma must be composed of at least -{\num} products. +from 1, counting both dependent and non dependent +products, but skipping local definitions. Especially, the current +lemma must be composed of at least {\num} products. Like in a {\tt fix} expression, the induction hypotheses have to be used on structurally smaller arguments. diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 0aab77314..43fc1d511 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -507,6 +507,9 @@ let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) w else let open Context.Rel.Declaration in check_mutind (push_rel (LocalAssum (na, c1)) env) sigma (pred k) b +| LetIn (na, c1, t, b) -> + let open Context.Rel.Declaration in + check_mutind (push_rel (LocalDef (na, c1, t)) env) sigma k b | _ -> error "Not enough products." (* Refine as a fixpoint *) diff --git a/vernac/command.ml b/vernac/command.ml index c24dbdf7c..a6343a5f5 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -867,8 +867,10 @@ let prepare_recursive_declaration fixnames fixtypes fixdefs = (* Jump over let-bindings. *) -let compute_possible_guardness_evidences (ids,_,na) = - match na with +let compute_possible_guardness_evidences (ctx,_,recindex) = + (* A recursive index is characterized by the number of lambdas to + skip before finding the relevant inductive argument *) + match recindex with | Some i -> [i] | None -> (* If recursive argument was not given by user, we try all args. @@ -876,7 +878,7 @@ let compute_possible_guardness_evidences (ids,_,na) = but doing it properly involves delta-reduction, and it finally doesn't seem to worth the effort (except for huge mutual fixpoints ?) *) - List.interval 0 (List.length ids - 1) + List.interval 0 (Context.Rel.nhyps ctx - 1) type recursive_preentry = Id.t list * constr option list * types list @@ -1136,10 +1138,10 @@ let interp_recursive isfix fixl notations = let evd, nf = nf_evars_and_universes evd in let fixdefs = List.map (Option.map nf) fixdefs in let fixtypes = List.map nf fixtypes in - let fixctxnames = List.map (fun (_,ctx) -> List.map RelDecl.get_name ctx) fixctxs in + let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in (* Build the fix declaration block *) - (env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots + (env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots let check_recursive isfix env evd (fixnames,fixdefs,_) = check_evars_are_solved env evd Evd.empty; @@ -1162,7 +1164,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind if List.exists Option.is_empty fixdefs then (* Some bodies to define by proof *) let thms = - List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) + List.map3 (fun id t (ctx,imps,_) -> ((id,pl),(t,(List.map RelDecl.get_name ctx,imps)))) fixnames fixtypes fiximps in let init_tac = Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC) @@ -1199,7 +1201,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n if List.exists Option.is_empty fixdefs then (* Some bodies to define by proof *) let thms = - List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) + List.map3 (fun id t (ctx,imps,_) -> ((id,pl),(t,(List.map RelDecl.get_name ctx,imps)))) fixnames fixtypes fiximps in let init_tac = Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC) diff --git a/vernac/command.mli b/vernac/command.mli index bccc22ae9..b62611e2e 100644 --- a/vernac/command.mli +++ b/vernac/command.mli @@ -138,24 +138,24 @@ type recursive_preentry = val interp_fixpoint : structured_fixpoint_expr list -> decl_notation list -> recursive_preentry * lident list option * Evd.evar_universe_context * - (Name.t list * Impargs.manual_implicits * int option) list + (Context.Rel.t * Impargs.manual_implicits * int option) list val interp_cofixpoint : structured_fixpoint_expr list -> decl_notation list -> recursive_preentry * lident list option * Evd.evar_universe_context * - (Name.t list * Impargs.manual_implicits * int option) list + (Context.Rel.t * Impargs.manual_implicits * int option) list (** Registering fixpoints and cofixpoints in the environment *) val declare_fixpoint : locality -> polymorphic -> recursive_preentry * lident list option * Evd.evar_universe_context * - (Name.t list * Impargs.manual_implicits * int option) list -> + (Context.Rel.t * Impargs.manual_implicits * int option) list -> lemma_possible_guards -> decl_notation list -> unit val declare_cofixpoint : locality -> polymorphic -> recursive_preentry * lident list option * Evd.evar_universe_context * - (Name.t list * Impargs.manual_implicits * int option) list -> + (Context.Rel.t * Impargs.manual_implicits * int option) list -> decl_notation list -> unit (** Entry points for the vernacular commands Fixpoint and CoFixpoint *) diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 430d48bc7..ca511b573 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -100,10 +100,10 @@ let find_mutually_recursive_statements thms = match kind_of_term t with | Ind ((kn,_ as ind),u) when let mind = Global.lookup_mind kn in - mind.mind_finite <> Decl_kinds.CoFinite && is_local_assum decl -> + mind.mind_finite <> Decl_kinds.CoFinite -> [ind,x,i] | _ -> - []) 0 (List.rev whnf_hyp_hds)) in + []) 0 (List.rev (List.filter RelDecl.is_local_assum whnf_hyp_hds))) in let ind_ccl = let cclenv = push_rel_context hyps (Global.env()) in let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli index 39c089be9..6b45ed933 100644 --- a/vernac/lemmas.mli +++ b/vernac/lemmas.mli @@ -41,8 +41,8 @@ val start_proof_com : val start_proof_with_initialization : goal_kind -> Evd.evar_map -> (bool * lemma_possible_guards * unit Proofview.tactic list option) option -> - ((Id.t * universe_binders option) * - (types * (Name.t list * Impargs.manual_explicitation list))) list + ((Id.t (* name of thm *) * universe_binders option) * + (types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list -> int list option -> unit declaration_hook -> unit val universe_proof_terminator : -- cgit v1.2.3 From 41b3486501f0217f0b0c552f21d6f0374b55a3ba Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Sat, 28 Jan 2017 14:02:02 +0100 Subject: Mention template polymorphism in the documentation. Since About mentions it the doc should as well. --- doc/refman/RefMan-cic.tex | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/doc/refman/RefMan-cic.tex b/doc/refman/RefMan-cic.tex index b9c17d814..12f390e44 100644 --- a/doc/refman/RefMan-cic.tex +++ b/doc/refman/RefMan-cic.tex @@ -1120,6 +1120,12 @@ Check (fun (A:Prop) (B:Set) => prod A B). Check (fun (A:Type) (B:Prop) => prod A B). \end{coq_example} +Internally, Coq calls sort polymorphism of inductive types +{\em template polymorphism}. For instance: +\begin{coq_example} +About prod. +\end{coq_example} + \subsection{Destructors} The specification of inductive definitions with arities and constructors is quite natural. But we still have to say how to use an -- cgit v1.2.3 From d3a2acc9fceff7476bc2d9eaadab8411365172a2 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 31 Mar 2017 15:03:34 +0200 Subject: Use "template polymorphism" in the documentation. --- doc/refman/RefMan-cic.tex | 14 ++++---------- 1 file changed, 4 insertions(+), 10 deletions(-) diff --git a/doc/refman/RefMan-cic.tex b/doc/refman/RefMan-cic.tex index 12f390e44..e26d4b18d 100644 --- a/doc/refman/RefMan-cic.tex +++ b/doc/refman/RefMan-cic.tex @@ -79,8 +79,8 @@ An algebraic universe $u$ is either a variable (a qualified identifier with a number) or a successor of an algebraic universe (an expression $u+1$), or an upper bound of algebraic universes (an expression $max(u_1,...,u_n)$), or the base universe (the expression -$0$) which corresponds, in the arity of sort-polymorphic inductive -types (see Section \ref{Sort-polymorphism-inductive}), +$0$) which corresponds, in the arity of template polymorphic inductive +types (see Section \ref{Template-polymorphism}), to the predicative sort {\Set}. A graph of constraints between the universe variables is maintained globally. To ensure the existence of a mapping of the universes to the positive integers, the graph of @@ -977,8 +977,8 @@ Inductive exType (P:Type->Prop) : Type := %is recursive or not. We shall write the type $(x:_R T)C$ if it is %a recursive argument and $(x:_P T)C$ if the argument is not recursive. -\paragraph[Sort-polymorphism of inductive types.]{Sort-polymorphism of inductive types.\index{Sort-polymorphism of inductive types}} -\label{Sort-polymorphism-inductive} +\paragraph[Template polymorphism.]{Template polymorphism.\index{Template polymorphism}} +\label{Template-polymorphism} Inductive types declared in {\Type} are polymorphic over their arguments in {\Type}. @@ -1120,12 +1120,6 @@ Check (fun (A:Prop) (B:Set) => prod A B). Check (fun (A:Type) (B:Prop) => prod A B). \end{coq_example} -Internally, Coq calls sort polymorphism of inductive types -{\em template polymorphism}. For instance: -\begin{coq_example} -About prod. -\end{coq_example} - \subsection{Destructors} The specification of inductive definitions with arities and constructors is quite natural. But we still have to say how to use an -- cgit v1.2.3 From fe99efdbe409e47f20776c62a76d4de7f0188afc Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Tue, 11 Apr 2017 12:46:23 +0200 Subject: Update various comments to use "template polymorphism" Also remove obvious comments. --- checker/subtyping.ml | 2 +- checker/typeops.ml | 4 +--- engine/universes.mli | 2 +- kernel/subtyping.ml | 2 +- kernel/typeops.ml | 4 +--- pretyping/coercion.ml | 4 ++-- pretyping/typing.ml | 3 +-- proofs/logic.ml | 2 +- test-suite/failure/proofirrelevance.v | 5 ++--- test-suite/success/Case19.v | 2 +- 10 files changed, 12 insertions(+), 18 deletions(-) diff --git a/checker/subtyping.ml b/checker/subtyping.ml index 7eae9b831..a290b240d 100644 --- a/checker/subtyping.ml +++ b/checker/subtyping.ml @@ -113,7 +113,7 @@ let check_inductive env mp1 l info1 mib2 spec2 subst1 subst2= in let check_inductive_type env t1 t2 = - (* Due to sort-polymorphism in inductive types, the conclusions of + (* Due to template polymorphism, the conclusions of t1 and t2, if in Type, are generated as the least upper bounds of the types of the constructors. diff --git a/checker/typeops.ml b/checker/typeops.ml index 173e19ce1..9bb76ba72 100644 --- a/checker/typeops.ml +++ b/checker/typeops.ml @@ -278,13 +278,11 @@ let rec execute env cstr = let j = match f with | Ind ind -> - (* Sort-polymorphism of inductive types *) judge_of_inductive_knowing_parameters env ind jl | Const cst -> - (* Sort-polymorphism of constant *) judge_of_constant_knowing_parameters env cst jl | _ -> - (* No sort-polymorphism *) + (* No template polymorphism *) execute env f in let jl = Array.map2 (fun c ty -> c,ty) args jl in diff --git a/engine/universes.mli b/engine/universes.mli index 932de941a..83ca1ea60 100644 --- a/engine/universes.mli +++ b/engine/universes.mli @@ -223,7 +223,7 @@ val refresh_constraints : UGraph.t -> universe_context_set -> universe_context_s val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds -(** {6 Support for old-style sort-polymorphism } *) +(** {6 Support for template polymorphism } *) val solve_constraints_system : universe option array -> universe array -> universe array -> universe array diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml index c8ceb064d..d0fdf9fda 100644 --- a/kernel/subtyping.ml +++ b/kernel/subtyping.ml @@ -117,7 +117,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2 let mib2 = Declareops.subst_mind_body subst2 mib2 in let check_inductive_type cst name env t1 t2 = - (* Due to sort-polymorphism in inductive types, the conclusions of + (* Due to template polymorphism, the conclusions of t1 and t2, if in Type, are generated as the least upper bounds of the types of the constructors. diff --git a/kernel/typeops.ml b/kernel/typeops.ml index 7d9a2aac0..dbc0dcb73 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -367,15 +367,13 @@ let rec execute env cstr = let ft = match kind_of_term f with | Ind ind when Environ.template_polymorphic_pind ind env -> - (* Template sort-polymorphism of inductive types *) let args = Array.map (fun t -> lazy t) argst in type_of_inductive_knowing_parameters env ind args | Const cst when Environ.template_polymorphic_pconstant cst env -> - (* Template sort-polymorphism of constants *) let args = Array.map (fun t -> lazy t) argst in type_of_constant_knowing_parameters env cst args | _ -> - (* Full or no sort-polymorphism *) + (* No template polymorphism *) execute env f in diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 542db7fdf..f4514d7c0 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -479,8 +479,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = (* We eta-expand (hence possibly modifying the original term!) *) (* and look for a coercion c:u1->t1 s.t. fun x:u1 => v' (c x)) *) (* has type forall (x:u1), u2 (with v' recursively obtained) *) - (* Note: we retype the term because sort-polymorphism may have *) - (* weaken its type *) + (* Note: we retype the term because template polymorphism may have *) + (* weakened its type *) let name = match name with | Anonymous -> Name Namegen.default_dependent_ident | _ -> name in diff --git a/pretyping/typing.ml b/pretyping/typing.ml index c2a030bcd..00535adb7 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -313,14 +313,13 @@ let rec execute env evdref cstr = let j = match EConstr.kind !evdref f with | Ind (ind, u) when EInstance.is_empty u && Environ.template_polymorphic_ind ind env -> - (* Sort-polymorphism of inductive types *) make_judge f (inductive_type_knowing_parameters env !evdref (ind, u) jl) | Const (cst, u) when EInstance.is_empty u && Environ.template_polymorphic_constant cst env -> - (* Sort-polymorphism of inductive types *) make_judge f (constant_type_knowing_parameters env !evdref (cst, u) jl) | _ -> + (* No template polymorphism *) execute env evdref f in e_judge_of_apply env evdref j jl diff --git a/proofs/logic.ml b/proofs/logic.ml index e6024785d..d8c1f1f3b 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -379,7 +379,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let (acc',hdty,sigma,applicand) = if is_template_polymorphic env sigma (EConstr.of_constr f) then let ty = - (* Template sort-polymorphism of definition and inductive types *) + (* Template polymorphism of definitions and inductive types *) let firstmeta = Array.findi (fun i x -> occur_meta sigma (EConstr.of_constr x)) l in let args, _ = Option.cata (fun i -> CArray.chop i l) (l, [||]) firstmeta in type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) (Array.map EConstr.of_constr args) diff --git a/test-suite/failure/proofirrelevance.v b/test-suite/failure/proofirrelevance.v index b62f9b686..bb9579d48 100644 --- a/test-suite/failure/proofirrelevance.v +++ b/test-suite/failure/proofirrelevance.v @@ -1,6 +1,5 @@ -(* This was working in version 8.1beta (bug in the Sort-polymorphism - of inductive types), but this is inconsistent with classical logic - in Prop *) +(* This was working in version 8.1beta (bug in template polymorphism), + but this is inconsistent with classical logic in Prop *) Inductive bool_in_prop : Type := hide : bool -> bool_in_prop with bool : Type := true : bool | false : bool. diff --git a/test-suite/success/Case19.v b/test-suite/success/Case19.v index c29e52978..e59828def 100644 --- a/test-suite/success/Case19.v +++ b/test-suite/success/Case19.v @@ -1,5 +1,5 @@ (* This used to fail in Coq version 8.1 beta due to a non variable - universe (issued by the inductive sort-polymorphism) being sent by + universe (issued by template polymorphism) being sent by pretyping to the kernel (bug #1182) *) Variable T : Type. -- cgit v1.2.3 From d70af8a387d1199be3327b3e4ef21dda9bb2155e Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Tue, 11 Apr 2017 12:47:29 +0200 Subject: Update RefMan-pre to mention template polymorphism. --- doc/refman/RefMan-pre.tex | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/refman/RefMan-pre.tex b/doc/refman/RefMan-pre.tex index f36969e82..0441f952d 100644 --- a/doc/refman/RefMan-pre.tex +++ b/doc/refman/RefMan-pre.tex @@ -529,7 +529,7 @@ intensive computations. Christine Paulin implemented an extension of inductive types allowing recursively non uniform parameters. Hugo Herbelin implemented -sort-polymorphism for inductive types. +sort-polymorphism for inductive types (now called template polymorphism). Claudio Sacerdoti Coen improved the tactics for rewriting on arbitrary compatible equivalence relations. He also generalized rewriting to -- cgit v1.2.3 From 4e70791036a1ab189579e109b428f46f45698b59 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 13 Apr 2017 12:13:04 +0200 Subject: Adding a fold_glob_constr_with_binders combinator. Binding generalizable_vars_of_glob_constr, occur_glob_constr, free_glob_vars, and bound_glob_vars on it. Most of the functions of which it factorizes the code were bugged with respect to bindings in the return clause of "match" and in either the types or the bodies of "fix/cofix". --- interp/implicit_quantifiers.ml | 56 +-------------- pretyping/glob_ops.ml | 153 ++++++++++++++--------------------------- pretyping/glob_ops.mli | 1 + test-suite/.csdp.cache | Bin 89077 -> 89077 bytes test-suite/success/boundvars.v | 5 ++ 5 files changed, 57 insertions(+), 158 deletions(-) create mode 100644 test-suite/success/boundvars.v diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index 7f11c0a3b..d6749e918 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -131,61 +131,7 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp if Id.List.mem_assoc id vs then vs else (id, loc) :: vs else vs - | GApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args) - | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) -> - let vs' = vars bound vs ty in - let bound' = add_name_to_ids bound na in - vars bound' vs' c - | GLetIn (loc,na,b,ty,c) -> - let vs' = vars bound vs b in - let vs'' = Option.fold_left (vars bound) vs' ty in - let bound' = add_name_to_ids bound na in - vars bound' vs'' c - | GCases (loc,sty,rtntypopt,tml,pl) -> - let vs1 = vars_option bound vs rtntypopt in - let vs2 = List.fold_left (fun vs (tm,_) -> vars bound vs tm) vs1 tml in - List.fold_left (vars_pattern bound) vs2 pl - | GLetTuple (loc,nal,rtntyp,b,c) -> - let vs1 = vars_return_type bound vs rtntyp in - let vs2 = vars bound vs1 b in - let bound' = List.fold_left add_name_to_ids bound nal in - vars bound' vs2 c - | GIf (loc,c,rtntyp,b1,b2) -> - let vs1 = vars_return_type bound vs rtntyp in - let vs2 = vars bound vs1 c in - let vs3 = vars bound vs2 b1 in - vars bound vs3 b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> - let bound' = Array.fold_right Id.Set.add idl bound in - let vars_fix i vs fid = - let vs1,bound1 = - List.fold_left - (fun (vs,bound) (na,k,bbd,bty) -> - let vs' = vars_option bound vs bbd in - let vs'' = vars bound vs' bty in - let bound' = add_name_to_ids bound na in - (vs'',bound') - ) - (vs,bound') - bl.(i) - in - let vs2 = vars bound1 vs1 tyl.(i) in - vars bound1 vs2 bv.(i) - in - Array.fold_left_i vars_fix vs idl - | GCast (loc,c,k) -> let v = vars bound vs c in - (match k with CastConv t | CastVM t -> vars bound v t | _ -> v) - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs - - and vars_pattern bound vs (loc,idl,p,c) = - let bound' = List.fold_right Id.Set.add idl bound in - vars bound' vs c - - and vars_option bound vs = function None -> vs | Some p -> vars bound vs p - - and vars_return_type bound vs (na,tyopt) = - let bound' = add_name_to_ids bound na in - vars_option bound' vs tyopt + | c -> Glob_ops.fold_glob_constr_with_binders Id.Set.add vars bound vs c in fun rt -> let vars = List.rev (vars bound [] rt) in List.iter (fun (id, loc) -> diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index ebbfa195f..aa296aace 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -214,55 +214,57 @@ let fold_glob_constr f acc = function f acc c | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc -let iter_glob_constr f = fold_glob_constr (fun () -> f) () +let fold_return_type_with_binders f g v acc (na,tyopt) = + Option.fold_left (f (name_fold g na v)) acc tyopt + +let fold_glob_constr_with_binders g f v acc = function + | GVar _ -> acc + | GApp (_,c,args) -> List.fold_left (f v) (f v acc c) args + | GLambda (_,na,_,b,c) | GProd (_,na,_,b,c) -> + f (name_fold g na v) (f v acc b) c + | GLetIn (_,na,b,t,c) -> + f (name_fold g na v) (Option.fold_left (f v) (f v acc b) t) c + | GCases (_,_,rtntypopt,tml,pl) -> + let fold_pattern acc (_,idl,p,c) = f (List.fold_right g idl v) acc c in + let fold_tomatch (v',acc) (tm,(na,onal)) = + (Option.fold_left (fun v'' (_,_,nal) -> List.fold_right (name_fold g) nal v'') + (name_fold g na v') onal, + f v acc tm) in + let (v',acc) = List.fold_left fold_tomatch (v,acc) tml in + let acc = Option.fold_left (f v') acc rtntypopt in + List.fold_left fold_pattern acc pl + | GLetTuple (_,nal,rtntyp,b,c) -> + f v (f v (fold_return_type_with_binders f g v acc rtntyp) b) c + | GIf (_,c,rtntyp,b1,b2) -> + f v (f v (f v (fold_return_type_with_binders f g v acc rtntyp) c) b1) b2 + | GRec (_,_,idl,bll,tyl,bv) -> + let f' i acc fid = + let v,acc = + List.fold_left + (fun (v,acc) (na,k,bbd,bty) -> + (name_fold g na v, f v (Option.fold_left (f v) acc bbd) bty)) + (v,acc) + bll.(i) in + f (Array.fold_right g idl v) (f v acc tyl.(i)) (bv.(i)) in + Array.fold_left_i f' acc idl + | GCast (_,c,k) -> + let acc = match k with + | CastConv t | CastVM t | CastNative t -> f v acc t | CastCoerce -> acc in + f v acc c + | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc -let same_id na id = match na with -| Anonymous -> false -| Name id' -> Id.equal id id' +let iter_glob_constr f = fold_glob_constr (fun () -> f) () let occur_glob_constr id = - let rec occur = function + let rec occur barred acc = function | GVar (loc,id') -> Id.equal id id' - | GApp (loc,f,args) -> (occur f) || (List.exists occur args) - | GLambda (loc,na,bk,ty,c) -> - (occur ty) || (not (same_id na id) && (occur c)) - | GProd (loc,na,bk,ty,c) -> - (occur ty) || (not (same_id na id) && (occur c)) - | GLetIn (loc,na,b,t,c) -> - (Option.fold_left (fun b t -> occur t || b) (occur b) t) || (not (same_id na id) && (occur c)) - | GCases (loc,sty,rtntypopt,tml,pl) -> - (occur_option rtntypopt) - || (List.exists (fun (tm,_) -> occur tm) tml) - || (List.exists occur_pattern pl) - | GLetTuple (loc,nal,rtntyp,b,c) -> - occur_return_type rtntyp id - || (occur b) || (not (List.mem_f Name.equal (Name id) nal) && (occur c)) - | GIf (loc,c,rtntyp,b1,b2) -> - occur_return_type rtntyp id || (occur c) || (occur b1) || (occur b2) - | GRec (loc,fk,idl,bl,tyl,bv) -> - not (Array.for_all4 (fun fid bl ty bd -> - let rec occur_fix = function - [] -> not (occur ty) && (Id.equal fid id || not(occur bd)) - | (na,k,bbd,bty)::bl -> - not (occur bty) && - (match bbd with - Some bd -> not (occur bd) - | _ -> true) && - (match na with Name id' -> Id.equal id id' | _ -> not (occur_fix bl)) in - occur_fix bl) - idl bl tyl bv) - | GCast (loc,c,k) -> (occur c) || (match k with CastConv t - | CastVM t | CastNative t -> occur t | CastCoerce -> false) - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false - - and occur_pattern (loc,idl,p,c) = not (Id.List.mem id idl) && (occur c) - - and occur_option = function None -> false | Some p -> occur p - - and occur_return_type (na,tyopt) id = not (same_id na id) && occur_option tyopt - - in occur - + | c -> + (* [g] looks if [id] appears in a binding position, in which + case, we don't have to look in the corresponding subterm *) + let g id' barred = barred || Id.equal id id' in + let f barred acc c = acc || not barred && occur false acc c in + fold_glob_constr_with_binders g f barred acc c in + occur false false let add_name_to_ids set na = match na with @@ -270,64 +272,9 @@ let add_name_to_ids set na = | Name id -> Id.Set.add id set let free_glob_vars = - let rec vars bounded vs = function - | GVar (loc,id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs - | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args) - | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) -> - let vs' = vars bounded vs ty in - let bounded' = add_name_to_ids bounded na in - vars bounded' vs' c - | GLetIn (loc,na,b,ty,c) -> - let vs' = vars bounded vs b in - let vs'' = Option.fold_left (vars bounded) vs' ty in - let bounded' = add_name_to_ids bounded na in - vars bounded' vs'' c - | GCases (loc,sty,rtntypopt,tml,pl) -> - let vs1 = vars_option bounded vs rtntypopt in - let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in - List.fold_left (vars_pattern bounded) vs2 pl - | GLetTuple (loc,nal,rtntyp,b,c) -> - let vs1 = vars_return_type bounded vs rtntyp in - let vs2 = vars bounded vs1 b in - let bounded' = List.fold_left add_name_to_ids bounded nal in - vars bounded' vs2 c - | GIf (loc,c,rtntyp,b1,b2) -> - let vs1 = vars_return_type bounded vs rtntyp in - let vs2 = vars bounded vs1 c in - let vs3 = vars bounded vs2 b1 in - vars bounded vs3 b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> - let bounded' = Array.fold_right Id.Set.add idl bounded in - let vars_fix i vs fid = - let vs1,bounded1 = - List.fold_left - (fun (vs,bounded) (na,k,bbd,bty) -> - let vs' = vars_option bounded vs bbd in - let vs'' = vars bounded vs' bty in - let bounded' = add_name_to_ids bounded na in - (vs'',bounded') - ) - (vs,bounded') - bl.(i) - in - let vs2 = vars bounded1 vs1 tyl.(i) in - vars bounded1 vs2 bv.(i) - in - Array.fold_left_i vars_fix vs idl - | GCast (loc,c,k) -> let v = vars bounded vs c in - (match k with CastConv t | CastVM t | CastNative t -> vars bounded v t | _ -> v) - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs - - and vars_pattern bounded vs (loc,idl,p,c) = - let bounded' = List.fold_right Id.Set.add idl bounded in - vars bounded' vs c - - and vars_option bounded vs = function None -> vs | Some p -> vars bounded vs p - - and vars_return_type bounded vs (na,tyopt) = - let bounded' = add_name_to_ids bounded na in - vars_option bounded' vs tyopt - in + let rec vars bound vs = function + | GVar (loc,id') -> if Id.Set.mem id' bound then vs else Id.Set.add id' vs + | c -> fold_glob_constr_with_binders Id.Set.add vars bound vs c in fun rt -> let vs = vars Id.Set.empty Id.Set.empty rt in Id.Set.elements vs diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index 55e6b6533..af2834e49 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -37,6 +37,7 @@ val map_glob_constr_left_to_right : val warn_variable_collision : ?loc:Loc.t -> Id.t -> unit val fold_glob_constr : ('a -> glob_constr -> 'a) -> 'a -> glob_constr -> 'a +val fold_glob_constr_with_binders : (Id.t -> 'a -> 'a) -> ('a -> 'b -> glob_constr -> 'b) -> 'a -> 'b -> glob_constr -> 'b val iter_glob_constr : (glob_constr -> unit) -> glob_constr -> unit val occur_glob_constr : Id.t -> glob_constr -> bool val free_glob_vars : glob_constr -> Id.t list diff --git a/test-suite/.csdp.cache b/test-suite/.csdp.cache index ba85286dd..b99d80e95 100644 Binary files a/test-suite/.csdp.cache and b/test-suite/.csdp.cache differ diff --git a/test-suite/success/boundvars.v b/test-suite/success/boundvars.v new file mode 100644 index 000000000..7b6696af8 --- /dev/null +++ b/test-suite/success/boundvars.v @@ -0,0 +1,5 @@ +(* An example showing a bug in the detection of free variables *) +(* "x" is not free in the common type of "x" and "y" *) + +Check forall (x z:unit) (x y : match z as x return x=x with tt => eq_refl end = eq_refl), x=x. + -- cgit v1.2.3 From b4936da085b19ad508346d8e07ce1e922ef79c2d Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 13 Apr 2017 15:05:16 +0200 Subject: Using fold_glob_constr_with_binders to code bound_glob_vars. To use the generic combinator, we introduce a side effect. I believe that we have more to gain from a short code than from being purely functional. This also fixes the expected semantics since the variables binding the return type in "match" were not taking into account. --- pretyping/glob_ops.ml | 57 ++++++----------------------------------- test-suite/.csdp.cache | Bin 89077 -> 89077 bytes test-suite/success/boundvars.v | 9 +++++++ 3 files changed, 17 insertions(+), 49 deletions(-) diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index aa296aace..080ec5ed1 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -300,57 +300,16 @@ let add_and_check_ident id set = Id.Set.add id set let bound_glob_vars = - let rec vars bound = function - | GLambda (_,na,_,_,_) | GProd (_,na,_,_,_) | GLetIn (_,na,_,_,_) as c -> - let bound = name_fold add_and_check_ident na bound in - fold_glob_constr vars bound c - | GCases (loc,sty,rtntypopt,tml,pl) -> - let bound = vars_option bound rtntypopt in - let bound = - List.fold_left (fun bound (tm,_) -> vars bound tm) bound tml in - List.fold_left vars_pattern bound pl - | GLetTuple (loc,nal,rtntyp,b,c) -> - let bound = vars_return_type bound rtntyp in - let bound = vars bound b in - let bound = List.fold_right (name_fold add_and_check_ident) nal bound in - vars bound c - | GIf (loc,c,rtntyp,b1,b2) -> - let bound = vars_return_type bound rtntyp in - let bound = vars bound c in - let bound = vars bound b1 in - vars bound b2 - | GRec (loc,fk,idl,bl,tyl,bv) -> - let bound = Array.fold_right Id.Set.add idl bound in - let vars_fix i bound fid = - let bound = - List.fold_left - (fun bound (na,k,bbd,bty) -> - let bound = vars_option bound bbd in - let bound = vars bound bty in - name_fold add_and_check_ident na bound - ) - bound - bl.(i) - in - let bound = vars bound tyl.(i) in - vars bound bv.(i) - in - Array.fold_left_i vars_fix bound idl - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound - | GApp _ | GCast _ as c -> fold_glob_constr vars bound c - - and vars_pattern bound (loc,idl,p,c) = - let bound = List.fold_right add_and_check_ident idl bound in - vars bound c - - and vars_option bound = function None -> bound | Some p -> vars bound p - - and vars_return_type bound (na,tyopt) = - let bound = name_fold add_and_check_ident na bound in - vars_option bound tyopt + let rec vars bound = + fold_glob_constr_with_binders + (fun id () -> bound := add_and_check_ident id !bound) + (fun () () -> vars bound) + () () in fun rt -> - vars Id.Set.empty rt + let bound = ref Id.Set.empty in + vars bound rt; + !bound (** Mapping of names in binders *) diff --git a/test-suite/.csdp.cache b/test-suite/.csdp.cache index b99d80e95..ba85286dd 100644 Binary files a/test-suite/.csdp.cache and b/test-suite/.csdp.cache differ diff --git a/test-suite/success/boundvars.v b/test-suite/success/boundvars.v index 7b6696af8..fafe27292 100644 --- a/test-suite/success/boundvars.v +++ b/test-suite/success/boundvars.v @@ -3,3 +3,12 @@ Check forall (x z:unit) (x y : match z as x return x=x with tt => eq_refl end = eq_refl), x=x. +(* An example showing a bug in the detection of bound variables *) + +Goal forall x, match x return x = x with 0 => eq_refl | _ => eq_refl end = eq_refl. +intro. +match goal with +|- (match x as y in nat return y = y with O => _ | S n => _ end) = _ => assert (forall y, y = 0) end. +intro. +Check x0. (* Check that "y" has been bound to "x0" while matching "match x as x0 return x0=x0 with ... end" *) +Abort. -- cgit v1.2.3 From 6795bc07f53a842bcec76ad0329d0b4444a625ab Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 2 Apr 2017 21:22:55 +0200 Subject: Replacing costly merges in UGraph. --- kernel/uGraph.ml | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml index 4884d0deb..6971c0a2b 100644 --- a/kernel/uGraph.ml +++ b/kernel/uGraph.ml @@ -354,13 +354,15 @@ let get_new_edges g to_merge = UMap.empty to_merge in let ltle = - UMap.fold (fun _ n acc -> - UMap.merge (fun _ strict1 strict2 -> - match strict1, strict2 with - | Some true, _ | _, Some true -> Some true - | _, _ -> Some false) - acc n.ltle) - to_merge_lvl UMap.empty + let fold _ n acc = + let fold u strict acc = + if strict then UMap.add u strict acc + else if UMap.mem u acc then acc + else UMap.add u false acc + in + UMap.fold fold n.ltle acc + in + UMap.fold fold to_merge_lvl UMap.empty in let ltle, _ = clean_ltle g ltle in let ltle = -- cgit v1.2.3 From 1fb4d1ef961f056e3575c5e034cace85f2340509 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 15 Apr 2017 22:51:20 +0200 Subject: Fix bug #5377: @? patterns broken. --- pretyping/constr_matching.ml | 78 ++++++++++++++++++++++++++++++++----------- test-suite/bugs/closed/5377.v | 54 ++++++++++++++++++++++++++++++ 2 files changed, 113 insertions(+), 19 deletions(-) create mode 100644 test-suite/bugs/closed/5377.v diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 886a98263..3c47cfdc4 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -80,31 +80,71 @@ let add_binders na1 na2 binding_vars (names, terms as subst) = let rec build_lambda vars ctx m = match vars with | [] -> - let len = List.length ctx in - lift (-1 * len) m + if Vars.closed0 m then m else raise PatternMatchingFailure | n :: vars -> (* change [ x1 ... xn y z1 ... zm |- t ] into [ x1 ... xn z1 ... zm |- lam y. t ] *) let len = List.length ctx in - let init i = - if i < pred n then mkRel (i + 2) - else if Int.equal i (pred n) then mkRel 1 - else mkRel (i + 1) - in - let m = substl (List.init len init) m in let pre, suf = List.chop (pred n) ctx in - match suf with + let (na, t, suf) = match suf with | [] -> assert false - | (_, na, t) :: suf -> - let map i = if i > n then pred i else i in - let vars = List.map map vars in - (** Check that the abstraction is legal *) - let frels = free_rels t in - let brels = List.fold_right Int.Set.add vars Int.Set.empty in - let () = if not (Int.Set.subset frels brels) then raise PatternMatchingFailure in - (** Create the abstraction *) - let m = mkLambda (na, t, m) in - build_lambda vars (pre @ suf) m + | (_, na, t) :: suf -> (na, t, suf) + in + (** Check that the abstraction is legal by generating a transitive closure of + its dependencies. *) + let is_nondep t clear = match clear with + | [] -> true + | _ -> + let rels = free_rels t in + let check i b = b || not (Int.Set.mem i rels) in + List.for_all_i check 1 clear + in + let fold (_, _, t) clear = is_nondep t clear :: clear in + (** Produce a list of booleans: true iff we keep the hypothesis *) + let clear = List.fold_right fold pre [false] in + let clear = List.drop_last clear in + (** If the conclusion depends on a variable we cleared, failure *) + let () = if not (is_nondep m clear) then raise PatternMatchingFailure in + (** Create the abstracted term *) + let fold (k, accu) keep = + if keep then + let k = succ k in + (k, Some k :: accu) + else (k, None :: accu) + in + let keep, shift = List.fold_left fold (0, []) clear in + let shift = List.rev shift in + let map = function + | None -> mkProp (** dummy term *) + | Some i -> mkRel (i + 1) + in + (** [x1 ... xn y z1 ... zm] -> [x1 ... xn f(z1) ... f(zm) y] *) + let subst = + List.map map shift @ + mkRel 1 :: + List.mapi (fun i _ -> mkRel (i + keep + 2)) suf + in + let map i (id, na, c) = + let i = succ i in + let subst = List.skipn i subst in + let subst = List.map (fun c -> Vars.lift (- i) c) subst in + (id, na, substl subst c) + in + let pre = List.mapi map pre in + let pre = List.filter_with clear pre in + let m = substl subst m in + let map i = + if i > n then i - n + keep + else match List.nth shift (i - 1) with + | None -> + (** We cleared a variable that we wanted to abstract! *) + raise PatternMatchingFailure + | Some k -> k + in + let vars = List.map map vars in + (** Create the abstraction *) + let m = mkLambda (na, Vars.lift keep t, m) in + build_lambda vars (pre @ suf) m let rec extract_bound_aux k accu frels ctx = match ctx with | [] -> accu diff --git a/test-suite/bugs/closed/5377.v b/test-suite/bugs/closed/5377.v new file mode 100644 index 000000000..130d9f9ab --- /dev/null +++ b/test-suite/bugs/closed/5377.v @@ -0,0 +1,54 @@ +Goal ((forall (t : Type) (x y : t), + True -> + x = y)) -> False. +Proof. + intro HG. + let P := lazymatch goal with + | [ H : forall t x y, True -> @?P t x y + |- _ ] + => P + end in + pose (f := P). + unify f (fun (t : Type) (x y : t) => x = y). +Abort. + +Goal True. +Proof. +let c := lazymatch constr:(fun (T : nat -> Type) (y : nat) (_ : T y) => y) with + | fun _ y _ => @?C y => C + end in +pose (f := c). +unify f (fun n : nat => n). +Abort. + +Goal (forall x : nat, x = x -> x = x \/ x = x) -> True. +Proof. +intro. +let P := lazymatch goal with +| [ H : forall y, @?P y -> @?P y \/ _ |- _ ] + => P +end in +pose (f := P). +unify f (fun x : nat => x = x). +Abort. + +Goal (forall x : nat, x = x -> x = x \/ x = x) -> True. +Proof. +intro. +lazymatch goal with +| [ H : forall y, @?P y -> @?Q y \/ _ |- _ ] + => idtac +end. +Abort. + +Axiom eq : forall {T} (_ : T), Prop. + +Goal forall _ : (forall t (_ : eq t) (x : t), eq x), Prop. +Proof. +intro. +let P := lazymatch goal with +| [ H : forall t _ x, @?P t x |- _ ] + => P +end in +pose (f := P). +Abort. -- cgit v1.2.3 From 9354722485c71ba6fbe6e6462eae98113aa830cc Mon Sep 17 00:00:00 2001 From: Guillaume Melquiond Date: Thu, 20 Apr 2017 17:50:10 +0200 Subject: Fix nsatz not recognizing real literals. --- plugins/nsatz/Nsatz.v | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/plugins/nsatz/Nsatz.v b/plugins/nsatz/Nsatz.v index b11d15e5c..403f664e2 100644 --- a/plugins/nsatz/Nsatz.v +++ b/plugins/nsatz/Nsatz.v @@ -462,6 +462,11 @@ try (try apply Rsth; exact Rplus_opp_r. Defined. +Class can_compute_Z (z : Z) := dummy_can_compute_Z : True. +Hint Extern 0 (can_compute_Z ?v) => + match isZcst v with true => exact I end : typeclass_instances. +Instance reify_IZR z lvar {_ : can_compute_Z z} : reify (PEc z) lvar (IZR z). + Lemma R_one_zero: 1%R <> 0%R. discrR. Qed. -- cgit v1.2.3 From b07f7a2ac4f23a29d1762b4f78fc1506925a5019 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 19 Apr 2017 17:41:31 -0400 Subject: Add bedrock targets src and facade --- .travis.yml | 2 ++ Makefile.ci | 3 +-- dev/ci/ci-basic-overlay.sh | 13 ++++++++++++- dev/ci/ci-bedrock-facade.sh | 10 ++++++++++ dev/ci/ci-bedrock-src.sh | 10 ++++++++++ 5 files changed, 35 insertions(+), 3 deletions(-) create mode 100755 dev/ci/ci-bedrock-facade.sh create mode 100755 dev/ci/ci-bedrock-src.sh diff --git a/.travis.yml b/.travis.yml index 81f50af0a..77aac23b7 100644 --- a/.travis.yml +++ b/.travis.yml @@ -27,6 +27,8 @@ env: - TEST_TARGET="test-suite" COMPILER="4.02.3+32bit" - TEST_TARGET="validate" TW="travis_wait" - TEST_TARGET="validate" COMPILER="4.02.3+32bit" TW="travis_wait" + - TEST_TARGET="ci-bedrock-src" + - TEST_TARGET="ci-bedrock-facade" - TEST_TARGET="ci-color" - TEST_TARGET="ci-compcert" - TEST_TARGET="ci-coquelicot" diff --git a/Makefile.ci b/Makefile.ci index b055ada8e..4c4606aff 100644 --- a/Makefile.ci +++ b/Makefile.ci @@ -1,11 +1,10 @@ CI_TARGETS=ci-all ci-hott ci-math-comp ci-compcert ci-sf ci-cpdt \ ci-color ci-math-classes ci-tlc ci-fiat-crypto ci-fiat-parsers \ ci-coquelicot ci-flocq ci-iris-coq ci-metacoq ci-geocoq \ - ci-unimath ci-vst + ci-unimath ci-vst ci-bedrock-src ci-bedrock-facade .PHONY: $(CI_TARGETS) # Generic rule, we use make to easy travis integraton with mixed rules $(CI_TARGETS): ci-%: ./dev/ci/ci-$*.sh - diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh index 336ce9d8f..e0851816c 100644 --- a/dev/ci/ci-basic-overlay.sh +++ b/dev/ci/ci-basic-overlay.sh @@ -94,6 +94,18 @@ : ${fiat_crypto_CI_BRANCH:=master} : ${fiat_crypto_CI_GITURL:=https://github.com/mit-plv/fiat-crypto.git} +######################################################################## +# bedrock_src +######################################################################## +: ${bedrock_src_CI_BRANCH:=master} +: ${bedrock_src_CI_GITURL:=https://github.com/JasonGross/bedrock.git} + +######################################################################## +# bedrock_facade +######################################################################## +: ${bedrock_facade_CI_BRANCH:=master} +: ${bedrock_facade_CI_GITURL:=https://github.com/JasonGross/bedrock.git} + ######################################################################## # CoLoR ######################################################################## @@ -109,4 +121,3 @@ ######################################################################## : ${tlc_CI_BRANCH:=master} : ${tlc_CI_GITURL:=https://gforge.inria.fr/git/tlc/tlc.git} - diff --git a/dev/ci/ci-bedrock-facade.sh b/dev/ci/ci-bedrock-facade.sh new file mode 100755 index 000000000..95cfa3073 --- /dev/null +++ b/dev/ci/ci-bedrock-facade.sh @@ -0,0 +1,10 @@ +#!/usr/bin/env bash + +ci_dir="$(dirname "$0")" +source ${ci_dir}/ci-common.sh + +bedrock_facade_CI_DIR=${CI_BUILD_DIR}/bedrock-facade + +git_checkout ${bedrock_facade_CI_BRANCH} ${bedrock_facade_CI_GITURL} ${bedrock_facade_CI_DIR} + +( cd ${bedrock_facade_CI_DIR} && make -j ${NJOBS} facade ) diff --git a/dev/ci/ci-bedrock-src.sh b/dev/ci/ci-bedrock-src.sh new file mode 100755 index 000000000..532611d4b --- /dev/null +++ b/dev/ci/ci-bedrock-src.sh @@ -0,0 +1,10 @@ +#!/usr/bin/env bash + +ci_dir="$(dirname "$0")" +source ${ci_dir}/ci-common.sh + +bedrock_src_CI_DIR=${CI_BUILD_DIR}/bedrock-src + +git_checkout ${bedrock_src_CI_BRANCH} ${bedrock_src_CI_GITURL} ${bedrock_src_CI_DIR} + +( cd ${bedrock_src_CI_DIR} && make -j ${NJOBS} src ) -- cgit v1.2.3 From 15d415729962eddde2cd1d58e03449c8526ba626 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Mon, 24 Apr 2017 16:39:25 +0200 Subject: refman: remember the old name of template polymorphism. --- doc/refman/RefMan-cic.tex | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/doc/refman/RefMan-cic.tex b/doc/refman/RefMan-cic.tex index e26d4b18d..fdd272581 100644 --- a/doc/refman/RefMan-cic.tex +++ b/doc/refman/RefMan-cic.tex @@ -1120,6 +1120,10 @@ Check (fun (A:Prop) (B:Set) => prod A B). Check (fun (A:Type) (B:Prop) => prod A B). \end{coq_example} +\Rem Template polymorphism used to be called ``sort-polymorphism of +inductive types'' before universe polymorphism (see +Chapter~\ref{Universes-full}) was introduced. + \subsection{Destructors} The specification of inductive definitions with arities and constructors is quite natural. But we still have to say how to use an -- cgit v1.2.3 From 552544a3d385a3a59def038bdb0a22a69fe4b0a9 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 7 Mar 2017 16:42:24 +0100 Subject: Removing the tclNOTSAMEGOAL primitive from the API. The only use in Equality is reimplemented in the new engine. --- proofs/refiner.ml | 13 ------------- proofs/refiner.mli | 1 - tactics/equality.ml | 22 +++++++++++++++++++--- tactics/tacticals.ml | 1 - tactics/tacticals.mli | 1 - 5 files changed, 19 insertions(+), 19 deletions(-) diff --git a/proofs/refiner.ml b/proofs/refiner.ml index 5c7659ac0..d086f0bbc 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -174,19 +174,6 @@ let tclPROGRESS tac ptree = if Goal.V82.progress rslt ptree then rslt else user_err ~hdr:"Refiner.PROGRESS" (str"Failed to progress.") -(* Same as tclWEAK_PROGRESS but fails also if tactics generates several goals, - one of them being identical to the original goal *) -let tclNOTSAMEGOAL (tac : tactic) goal = - let same_goal gls1 evd2 gl2 = - Goal.V82.same_goal gls1.sigma gls1.it evd2 gl2 - in - let rslt = tac goal in - let {it=gls;sigma=sigma} = rslt in - if List.exists (same_goal goal sigma) gls - then user_err ~hdr:"Refiner.tclNOTSAMEGOAL" - (str"Tactic generated a subgoal identical to the original goal.") - else rslt - (* Execute tac, show the names of new hypothesis names created by tac in the "as" format and then forget everything. From the logical point of view [tclSHOWHYPS tac] is therefore equivalent to idtac, diff --git a/proofs/refiner.mli b/proofs/refiner.mli index 56f5facf8..de6502dc3 100644 --- a/proofs/refiner.mli +++ b/proofs/refiner.mli @@ -122,7 +122,6 @@ val tclDO : int -> tactic -> tactic val tclWEAK_PROGRESS : tactic -> tactic val tclPROGRESS : tactic -> tactic val tclSHOWHYPS : tactic -> tactic -val tclNOTSAMEGOAL : tactic -> tactic (** [tclIFTHENELSE tac1 tac2 tac3 gls] first applies [tac1] to [gls] then, if it succeeds, applies [tac2] to the resulting subgoals, diff --git a/tactics/equality.ml b/tactics/equality.ml index 7ae7446c8..e68be91e9 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -97,9 +97,6 @@ let _ = (* Rewriting tactics *) -let tclNOTSAMEGOAL tac = - Proofview.V82.tactic (Tacticals.tclNOTSAMEGOAL (Proofview.V82.of_tactic tac)) - type dep_proof_flag = bool (* true = support rewriting dependent proofs *) type freeze_evars_flag = bool (* true = don't instantiate existing evars *) @@ -268,6 +265,25 @@ let rewrite_elim with_evars frzevars cls c e = general_elim_clause with_evars flags cls c e end } +let tclNOTSAMEGOAL tac = + let goal gl = Proofview.Goal.goal (Proofview.Goal.assume gl) in + Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = project gl in + let ev = goal gl in + tac >>= fun () -> + Proofview.Goal.goals >>= fun gls -> + let check accu gl' = + gl' >>= fun gl' -> + let accu = accu || Goal.V82.same_goal sigma ev (project gl') (goal gl') in + Proofview.tclUNIT accu + in + Proofview.Monad.List.fold_left check false gls >>= fun has_same -> + if has_same then + tclZEROMSG (str"Tactic generated a subgoal identical to the original goal.") + else + Proofview.tclUNIT () + end } + (* Ad hoc asymmetric general_elim_clause *) let general_elim_clause with_evars frzevars cls rew elim = let open Pretype_errors in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 90b7d6581..97922a4fa 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -54,7 +54,6 @@ let tclDO = Refiner.tclDO let tclWEAK_PROGRESS = Refiner.tclWEAK_PROGRESS let tclPROGRESS = Refiner.tclPROGRESS let tclSHOWHYPS = Refiner.tclSHOWHYPS -let tclNOTSAMEGOAL = Refiner.tclNOTSAMEGOAL let tclTHENTRY = Refiner.tclTHENTRY let tclIFTHENELSE = Refiner.tclIFTHENELSE let tclIFTHENSELSE = Refiner.tclIFTHENSELSE diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 3b90ec514..33933924a 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -47,7 +47,6 @@ val tclDO : int -> tactic -> tactic val tclWEAK_PROGRESS : tactic -> tactic val tclPROGRESS : tactic -> tactic val tclSHOWHYPS : tactic -> tactic -val tclNOTSAMEGOAL : tactic -> tactic val tclTHENTRY : tactic -> tactic -> tactic val tclMAP : ('a -> tactic) -> 'a list -> tactic -- cgit v1.2.3 From 21ee9a82807176acecf917e2a1e6074bed1c88ff Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 7 Mar 2017 18:16:27 +0100 Subject: Removing the tclWEAK_PROGRESS tactical. The only remaining use was applied on the unfold tactic, and the behaviours of tclPROGRESS and tclWEAK_PROGRESS coincide whenever only one goal is produced by their argument tactic. --- proofs/goal.mli | 3 --- proofs/refiner.ml | 7 ------- proofs/refiner.mli | 1 - tactics/class_tactics.ml | 3 +-- tactics/tacticals.ml | 1 - tactics/tacticals.mli | 1 - 6 files changed, 1 insertion(+), 15 deletions(-) diff --git a/proofs/goal.mli b/proofs/goal.mli index a2fa34c05..ee2e73612 100644 --- a/proofs/goal.mli +++ b/proofs/goal.mli @@ -59,9 +59,6 @@ module V82 : sig second goal *) val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map - (* Principal part of the weak-progress tactical *) - val weak_progress : goal list Evd.sigma -> goal Evd.sigma -> bool - (* Principal part of the progress tactical *) val progress : goal list Evd.sigma -> goal Evd.sigma -> bool diff --git a/proofs/refiner.ml b/proofs/refiner.ml index d086f0bbc..bc12b3ba7 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -160,13 +160,6 @@ let rec tclTHENLIST = function let tclMAP tacfun l = List.fold_right (fun x -> (tclTHEN (tacfun x))) l tclIDTAC -(* PROGRESS tac ptree applies tac to the goal ptree and fails if tac leaves -the goal unchanged *) -let tclWEAK_PROGRESS tac ptree = - let rslt = tac ptree in - if Goal.V82.weak_progress rslt ptree then rslt - else user_err ~hdr:"Refiner.WEAK_PROGRESS" (str"Failed to progress.") - (* PROGRESS tac ptree applies tac to the goal ptree and fails if tac leaves the goal unchanged *) let tclPROGRESS tac ptree = diff --git a/proofs/refiner.mli b/proofs/refiner.mli index de6502dc3..e179589df 100644 --- a/proofs/refiner.mli +++ b/proofs/refiner.mli @@ -119,7 +119,6 @@ val tclAT_LEAST_ONCE : tactic -> tactic val tclFAIL : int -> Pp.std_ppcmds -> tactic val tclFAIL_lazy : int -> Pp.std_ppcmds Lazy.t -> tactic val tclDO : int -> tactic -> tactic -val tclWEAK_PROGRESS : tactic -> tactic val tclPROGRESS : tactic -> tactic val tclSHOWHYPS : tactic -> tactic diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index df222eed8..d70275dee 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -462,8 +462,7 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co else e_trivial_fail_db only_classes db_list local_db secvars in Tacticals.New.tclTHEN fst snd | Unfold_nth c -> - let tac = Proofview.V82.of_tactic (unfold_in_concl [AllOccurrences,c]) in - Proofview.V82.tactic (tclWEAK_PROGRESS tac) + Proofview.tclPROGRESS (unfold_in_concl [AllOccurrences,c]) | Extern tacast -> conclPattern concl p tacast in let tac = run_hint t tac in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 97922a4fa..5c97f27ba 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -51,7 +51,6 @@ let tclAT_LEAST_ONCE = Refiner.tclAT_LEAST_ONCE let tclFAIL = Refiner.tclFAIL let tclFAIL_lazy = Refiner.tclFAIL_lazy let tclDO = Refiner.tclDO -let tclWEAK_PROGRESS = Refiner.tclWEAK_PROGRESS let tclPROGRESS = Refiner.tclPROGRESS let tclSHOWHYPS = Refiner.tclSHOWHYPS let tclTHENTRY = Refiner.tclTHENTRY diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 33933924a..01b9e5e93 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -44,7 +44,6 @@ val tclAT_LEAST_ONCE : tactic -> tactic val tclFAIL : int -> std_ppcmds -> tactic val tclFAIL_lazy : int -> std_ppcmds Lazy.t -> tactic val tclDO : int -> tactic -> tactic -val tclWEAK_PROGRESS : tactic -> tactic val tclPROGRESS : tactic -> tactic val tclSHOWHYPS : tactic -> tactic val tclTHENTRY : tactic -> tactic -> tactic -- cgit v1.2.3 From 81c964cfd9d7d3c4ea146889a59cc332987be03c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 7 Mar 2017 20:01:08 +0100 Subject: Porting generalize_dep to the new tactic engine. --- tactics/tactics.ml | 39 ++++++++++++++++++--------------------- 1 file changed, 18 insertions(+), 21 deletions(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e79258582..1615aec89 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2811,20 +2811,18 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = mkProd_or_LetIn decl cl', sigma' let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = - let env = Tacmach.pf_env gl in - let ids = Tacmach.pf_ids_of_hyps gl in - let sigma, t = Typing.type_of env sigma c in - generalize_goal_gen env sigma ids i o t cl - -let new_generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = - let env = Tacmach.New.pf_env gl in - let ids = Tacmach.New.pf_ids_of_hyps gl in + let open Tacmach.New in + let env = pf_env gl in + let ids = pf_ids_of_hyps gl in let sigma, t = Typing.type_of env sigma c in generalize_goal_gen env sigma ids i o t cl -let old_generalize_dep ?(with_let=false) c gl = +let generalize_dep ?(with_let=false) c = + let open Tacmach.New in + let open Tacticals.New in + Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = pf_env gl in - let sign = pf_hyps gl in + let sign = Proofview.Goal.hyps gl in let sigma = project gl in let init_ids = ids_of_named_context (Global.named_context()) in let seek (d:named_declaration) (toquant:named_context) = @@ -2843,11 +2841,11 @@ let old_generalize_dep ?(with_let=false) c gl = -> id::tothin | _ -> tothin in - let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in + let cl' = it_mkNamedProd_or_LetIn (pf_concl gl) to_quantify in let body = if with_let then match EConstr.kind sigma c with - | Var id -> id |> Tacmach.pf_get_hyp gl |> NamedDecl.get_value + | Var id -> id |> (fun id -> pf_get_hyp id gl) |> NamedDecl.get_value | _ -> None else None in @@ -2856,20 +2854,19 @@ let old_generalize_dep ?(with_let=false) c gl = (** Check that the generalization is indeed well-typed *) let (evd, _) = Typing.type_of env evd cl'' in let args = Context.Named.to_instance mkVar to_quantify_rev in - tclTHENLIST - [tclEVARS evd; - Proofview.V82.of_tactic (apply_type cl'' (if Option.is_empty body then c::args else args)); - Proofview.V82.of_tactic (clear (List.rev tothin'))] - gl - -let generalize_dep ?(with_let = false) c = - Proofview.V82.tactic (old_generalize_dep ~with_let c) + let tac = + tclTHEN + (apply_type cl'' (if Option.is_empty body then c::args else args)) + (clear (List.rev tothin')) + in + Sigma.Unsafe.of_pair (tac, evd) + end } (** *) let generalize_gen_let lconstr = Proofview.Goal.s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let newcl, evd = - List.fold_right_i (new_generalize_goal gl) 0 lconstr + List.fold_right_i (generalize_goal gl) 0 lconstr (Tacmach.New.pf_concl gl,Tacmach.New.project gl) in let (evd, _) = Typing.type_of env evd newcl in -- cgit v1.2.3 From 804526b13b123c07dd11e1f23f30e8b01c0c8610 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 8 Mar 2017 09:12:41 +0100 Subject: Removing trivial compatibility layer in omega. --- plugins/omega/coq_omega.ml | 262 +++++++++++++++++++++++---------------------- 1 file changed, 135 insertions(+), 127 deletions(-) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 7780de712..c57719ac4 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -19,7 +19,7 @@ open Names open Nameops open Term open EConstr -open Tacticals +open Tacticals.New open Tacmach open Tactics open Logic @@ -41,7 +41,9 @@ let elim_id id = Proofview.Goal.enter { enter = begin fun gl -> simplest_elim (Tacmach.New.pf_global id gl) end } -let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl +let resolve_id id = Proofview.Goal.enter { enter = begin fun gl -> + apply (Tacmach.New.pf_global id gl) +end } let timing timer_name f arg = f arg @@ -146,7 +148,7 @@ let intern_id,unintern_id,reset_intern_tables = Hashtbl.add table v idx; Hashtbl.add co_table idx v; v), (fun () -> cpt := 0; Hashtbl.clear table) -let mk_then = tclTHENLIST +let mk_then tacs = tclTHENLIST tacs let exists_tac c = constructor_tac false (Some 1) 1 (ImplicitBindings [c]) @@ -580,9 +582,11 @@ let abstract_path sigma typ path t = let abstract = context sigma (fun i t -> term_occur:= t; mkRel i) path t in mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur -let focused_simpl path gl = +let focused_simpl path = + Proofview.V82.tactic begin fun gl -> let newc = context (project gl) (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl + end let focused_simpl path = focused_simpl path @@ -642,7 +646,8 @@ let decompile af = let mkNewMeta () = mkMeta (Evarutil.new_meta()) -let clever_rewrite_base_poly typ p result theorem gl = +let clever_rewrite_base_poly typ p result theorem = + Proofview.V82.tactic begin fun gl -> let full = pf_concl gl in let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in let t = @@ -658,12 +663,13 @@ let clever_rewrite_base_poly typ p result theorem gl = [abstracted]) in exact (applist(t,[mkNewMeta()])) gl + end -let clever_rewrite_base p result theorem gl = - clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem gl +let clever_rewrite_base p result theorem = + clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem -let clever_rewrite_base_nat p result theorem gl = - clever_rewrite_base_poly (Lazy.force coq_nat) p result theorem gl +let clever_rewrite_base_nat p result theorem = + clever_rewrite_base_poly (Lazy.force coq_nat) p result theorem let clever_rewrite_gen p result (t,args) = let theorem = applist(t, args) in @@ -673,12 +679,14 @@ let clever_rewrite_gen_nat p result (t,args) = let theorem = applist(t, args) in clever_rewrite_base_nat p result theorem -let clever_rewrite p vpath t gl = +let clever_rewrite p vpath t = + Proofview.V82.tactic begin fun gl -> let full = pf_concl gl in let (abstracted,occ) = abstract_path (project gl) (Lazy.force coq_Z) (List.rev p) full in let vargs = List.map (fun p -> occurrence (project gl) p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in exact (applist(t',[mkNewMeta()])) gl + end let rec shuffle p (t1,t2) = match t1,t2 with @@ -942,15 +950,15 @@ let rec transform sigma p t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in - Proofview.V82.of_tactic (unfold sp_Zminus) :: tac,t + unfold sp_Zminus :: tac,t | Kapp(Zsucc,[t1]) -> let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer one |])) in - Proofview.V82.of_tactic (unfold sp_Zsucc) :: tac,t + unfold sp_Zsucc :: tac,t | Kapp(Zpred,[t1]) -> let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer negone |])) in - Proofview.V82.of_tactic (unfold sp_Zpred) :: tac,t + unfold sp_Zpred :: tac,t | Kapp(Zmult,[t1;t2]) -> let tac1,t1' = transform sigma (P_APP 1 :: p) t1 and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in @@ -1068,7 +1076,7 @@ let replay_history tactic_normalisation = | HYP e :: l -> begin try - Tacticals.New.tclTHEN + tclTHEN (Id.List.assoc (hyp_of_tag e.id) tactic_normalisation) (loop l) with Not_found -> loop l end @@ -1080,16 +1088,16 @@ let replay_history tactic_normalisation = let k = if b then negone else one in let p_initial = [P_APP 1;P_TYPE] in let tac= shuffle_mult_right p_initial e1.body k e2.body in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ generalize_tac [mkApp (Lazy.force coq_OMEGA17, [| val_of eq1; val_of eq2; mk_integer k; mkVar id1; mkVar id2 |])]; - Proofview.V82.tactic (mk_then tac); + mk_then tac; (intros_using [aux]); - Proofview.V82.tactic (resolve_id aux); + resolve_id aux; reflexivity ] | CONTRADICTION (e1,e2) :: l -> @@ -1104,8 +1112,8 @@ let replay_history tactic_normalisation = Lazy.force coq_Gt; Lazy.force coq_Gt |]) in - Tacticals.New.tclTHENS - (Tacticals.New.tclTHENLIST [ + tclTHENS + (tclTHENLIST [ unfold sp_Zle; simpl_in_concl; intro; @@ -1118,7 +1126,7 @@ let replay_history tactic_normalisation = mkVar (hyp_of_tag e1.id); mkVar (hyp_of_tag e2.id) |]) in - Proofview.tclTHEN (Proofview.V82.tactic (tclTHEN (Proofview.V82.of_tactic (generalize_tac [theorem])) (mk_then tac))) (solve_le) + Proofview.tclTHEN (tclTHEN (generalize_tac [theorem]) (mk_then tac)) solve_le | DIVIDE_AND_APPROX (e1,e2,k,d) :: l -> let id = hyp_of_tag e1.id in let eq1 = val_of(decompile e1) @@ -1128,10 +1136,10 @@ let replay_history tactic_normalisation = let rhs = mk_plus (mk_times eq2 kk) dd in let state_eg = mk_eq eq1 rhs in let tac = scalar_norm_add [P_APP 3] e2.body in - Tacticals.New.tclTHENS + tclTHENS (cut state_eg) - [ Tacticals.New.tclTHENS - (Tacticals.New.tclTHENLIST [ + [ tclTHENS + (tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA1, @@ -1139,9 +1147,9 @@ let replay_history tactic_normalisation = (clear [aux;id]); (intros_using [id]); (cut (mk_gt kk dd)) ]) - [ Tacticals.New.tclTHENS + [ tclTHENS (cut (mk_gt kk izero)) - [ Tacticals.New.tclTHENLIST [ + [ tclTHENLIST [ (intros_using [aux1; aux2]); (generalize_tac [mkApp (Lazy.force coq_Zmult_le_approx, @@ -1149,13 +1157,13 @@ let replay_history tactic_normalisation = (clear [aux1;aux2;id]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (unfold sp_Zgt); simpl_in_concl; reflexivity ] ]; - Tacticals.New.tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; - Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ] + tclTHEN (mk_then tac) reflexivity ] | NOT_EXACT_DIVIDE (e1,k) :: l -> let c = floor_div e1.constant k in @@ -1166,10 +1174,10 @@ let replay_history tactic_normalisation = let kk = mk_integer k and dd = mk_integer d in let tac = scalar_norm_add [P_APP 2] e2.body in - Tacticals.New.tclTHENS + tclTHENS (cut (mk_gt dd izero)) - [ Tacticals.New.tclTHENS (cut (mk_gt kk dd)) - [Tacticals.New.tclTHENLIST [ + [ tclTHENS (cut (mk_gt kk dd)) + [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA4, @@ -1177,14 +1185,14 @@ let replay_history tactic_normalisation = (clear [aux1;aux2]); unfold sp_not; (intros_using [aux]); - Proofview.V82.tactic (resolve_id aux); - Proofview.V82.tactic (mk_then tac); + resolve_id aux; + mk_then tac; assumption ] ; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ] @@ -1197,9 +1205,9 @@ let replay_history tactic_normalisation = let state_eq = mk_eq eq1 (mk_times eq2 kk) in if e1.kind == DISE then let tac = scalar_norm [P_APP 3] e2.body in - Tacticals.New.tclTHENS + tclTHENS (cut state_eq) - [Tacticals.New.tclTHENLIST [ + [tclTHENLIST [ (intros_using [aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA18, @@ -1207,14 +1215,14 @@ let replay_history tactic_normalisation = (clear [aux1;id]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ] + tclTHEN (mk_then tac) reflexivity ] else let tac = scalar_norm [P_APP 3] e2.body in - Tacticals.New.tclTHENS (cut state_eq) + tclTHENS (cut state_eq) [ - Tacticals.New.tclTHENS + tclTHENS (cut (mk_gt kk izero)) - [Tacticals.New.tclTHENLIST [ + [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA3, @@ -1222,11 +1230,11 @@ let replay_history tactic_normalisation = (clear [aux1;aux2;id]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; - Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ] + tclTHEN (mk_then tac) reflexivity ] | (MERGE_EQ(e3,e1,e2)) :: l -> let id = new_identifier () in tag_hypothesis id e3; @@ -1239,16 +1247,16 @@ let replay_history tactic_normalisation = (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: scalar_norm [P_APP 3] e1.body in - Tacticals.New.tclTHENS + tclTHENS (cut (mk_eq eq1 (mk_inv eq2))) - [Tacticals.New.tclTHENLIST [ + [tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA8, [| eq1;eq2;mkVar id1;mkVar id2; mkVar aux|])]); (clear [id1;id2;aux]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity] + tclTHEN (mk_then tac) reflexivity] | STATE {st_new_eq=e;st_def=def;st_orig=orig;st_coef=m;st_var=v} :: l -> let id = new_identifier () @@ -1272,9 +1280,9 @@ let replay_history tactic_normalisation = [[P_APP 1]] (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: shuffle_mult_right p_initial orig.body m ({c= negone;v= v}::def.body) in - Tacticals.New.tclTHENS + tclTHENS (cut theorem) - [Tacticals.New.tclTHENLIST [ + [tclTHENLIST [ (intros_using [aux]); (elim_id aux); (clear [aux]); @@ -1282,11 +1290,11 @@ let replay_history tactic_normalisation = (generalize_tac [mkApp (Lazy.force coq_OMEGA9, [| mkVar vid;eq2;eq1;mm; mkVar id2;mkVar aux |])]); - Proofview.V82.tactic (mk_then tac); + mk_then tac; (clear [aux]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHEN (exists_tac eq1) reflexivity ] + tclTHEN (exists_tac eq1) reflexivity ] | SPLIT_INEQ(e,(e1,act1),(e2,act2)) :: l -> let id1 = new_identifier () and id2 = new_identifier () in @@ -1295,10 +1303,10 @@ let replay_history tactic_normalisation = let tac1 = norm_add [P_APP 2;P_TYPE] e.body in let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in let eq = val_of(decompile e) in - Tacticals.New.tclTHENS + tclTHENS (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) - [Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac1); (intros_using [id1]); (loop act1) ]; - Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac2); (intros_using [id2]); (loop act2) ]] + [tclTHENLIST [ mk_then tac1; (intros_using [id1]); (loop act1) ]; + tclTHENLIST [ mk_then tac2; (intros_using [id2]); (loop act2) ]] | SUM(e3,(k1,e1),(k2,e2)) :: l -> let id = new_identifier () in tag_hypothesis id e3; @@ -1317,10 +1325,10 @@ let replay_history tactic_normalisation = let p_initial = if e1.kind == DISE then [P_APP 1; P_TYPE] else [P_APP 2; P_TYPE] in let tac = shuffle_mult_right p_initial e1.body k2 e2.body in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (tac_thm, [| eq1; eq2; kk; mkVar id1; mkVar id2 |])]); - Proofview.V82.tactic (mk_then tac); + mk_then tac; (intros_using [id]); (loop l) ] @@ -1329,10 +1337,10 @@ let replay_history tactic_normalisation = and kk2 = mk_integer k2 in let p_initial = [P_APP 2;P_TYPE] in let tac= shuffle_mult p_initial k1 e1.body k2 e2.body in - Tacticals.New.tclTHENS (cut (mk_gt kk1 izero)) - [Tacticals.New.tclTHENS + tclTHENS (cut (mk_gt kk1 izero)) + [tclTHENS (cut (mk_gt kk2 izero)) - [Tacticals.New.tclTHENLIST [ + [tclTHENLIST [ (intros_using [aux2;aux1]); (generalize_tac [mkApp (Lazy.force coq_OMEGA7, [| @@ -1340,29 +1348,29 @@ let replay_history tactic_normalisation = mkVar aux1;mkVar aux2; mkVar id1;mkVar id2 |])]); (clear [aux1;aux2]); - Proofview.V82.tactic (mk_then tac); + mk_then tac; (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ] | CONSTANT_NOT_NUL(e,k) :: l -> - Tacticals.New.tclTHEN ((generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl + tclTHEN ((generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl | CONSTANT_NUL(e) :: l -> - Tacticals.New.tclTHEN (Proofview.V82.tactic (resolve_id (hyp_of_tag e))) reflexivity + tclTHEN (resolve_id (hyp_of_tag e)) reflexivity | CONSTANT_NEG(e,k) :: l -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkVar (hyp_of_tag e)]); unfold sp_Zle; simpl_in_concl; unfold sp_not; (intros_using [aux]); - Proofview.V82.tactic (resolve_id aux); + resolve_id aux; reflexivity ] | _ -> Proofview.tclUNIT () @@ -1380,12 +1388,12 @@ let normalize_equation sigma id flag theorem pos t t1 t2 (tactic,defs) = let (tac,t') = normalize sigma p_initial t in let shift_left = tclTHEN - (Proofview.V82.of_tactic (generalize_tac [mkApp (theorem, [| t1; t2; mkVar id |]) ])) - (tclTRY (Proofview.V82.of_tactic (clear [id]))) + (generalize_tac [mkApp (theorem, [| t1; t2; mkVar id |]) ]) + (tclTRY (clear [id])) in if not (List.is_empty tac) then let id' = new_identifier () in - ((id',(Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (shift_left); Proofview.V82.tactic (mk_then tac); (intros_using [id']) ])) + ((id',(tclTHENLIST [ shift_left; mk_then tac; (intros_using [id']) ])) :: tactic, compile id' flag t' :: defs) else @@ -1430,7 +1438,7 @@ let destructure_omega gl tac_def (id,c) = let reintroduce id = (* [id] cannot be cleared if dependent: protect it by a try *) - Tacticals.New.tclTHEN (Tacticals.New.tclTRY (clear [id])) (intro_using id) + tclTHEN (tclTRY (clear [id])) (intro_using id) open Proofview.Notations @@ -1449,7 +1457,7 @@ let coq_omega = let id = new_identifier () in let i = new_id () in tag_hypothesis id i; - (Tacticals.New.tclTHENLIST [ + (tclTHENLIST [ (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); (intros_using [v; id]); (elim_id id); @@ -1460,7 +1468,7 @@ let coq_omega = body = [{v=intern_id v; c=one}]; constant = zero; id = i} :: sys else - (Tacticals.New.tclTHENLIST [ + (tclTHENLIST [ (simplest_elim (applist (Lazy.force coq_new_var, [t]))); (intros_using [v;th]); tac ]), @@ -1476,13 +1484,13 @@ let coq_omega = with UNSOLVABLE -> let _,path = depend [] [] (history ()) in if !display_action_flag then display_action display_var path; - (Tacticals.New.tclTHEN prelude (replay_history tactic_normalisation path)) + (tclTHEN prelude (replay_history tactic_normalisation path)) end else begin try let path = simplify_strong (new_id,new_var_num,display_var) system in if !display_action_flag then display_action display_var path; - Tacticals.New.tclTHEN prelude (replay_history tactic_normalisation path) - with NO_CONTRADICTION -> Tacticals.New.tclZEROMSG (Pp.str"Omega can't solve this system") + tclTHEN prelude (replay_history tactic_normalisation path) + with NO_CONTRADICTION -> tclZEROMSG (Pp.str"Omega can't solve this system") end end } @@ -1495,36 +1503,36 @@ let nat_inject = Proofview.tclEVARMAP >>= fun sigma -> try match destructurate_term sigma t with | Kapp(Plus,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ - Proofview.V82.tactic (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) + tclTHENLIST [ + (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_plus),[t1;t2])); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ] | Kapp(Mult,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ - Proofview.V82.tactic (clever_rewrite_gen p (mk_times (mk_inj t1) (mk_inj t2)) + tclTHENLIST [ + (clever_rewrite_gen p (mk_times (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_mult),[t1;t2])); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ] | Kapp(Minus,[t1;t2]) -> let id = new_identifier () in - Tacticals.New.tclTHENS - (Tacticals.New.tclTHEN + tclTHENS + (tclTHEN (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) (intros_using [id])) [ - Tacticals.New.tclTHENLIST [ - Proofview.V82.tactic (clever_rewrite_gen p + tclTHENLIST [ + (clever_rewrite_gen p (mk_minus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_minus1),[t1;t2;mkVar id])); (loop [id,mkApp (Lazy.force coq_le, [| t2;t1 |])]); (explore (P_APP 1 :: p) t1); (explore (P_APP 2 :: p) t2) ]; - (Tacticals.New.tclTHEN - (Proofview.V82.tactic (clever_rewrite_gen p (mk_integer zero) - ((Lazy.force coq_inj_minus2),[t1;t2;mkVar id]))) + (tclTHEN + (clever_rewrite_gen p (mk_integer zero) + ((Lazy.force coq_inj_minus2),[t1;t2;mkVar id])) (loop [id,mkApp (Lazy.force coq_gt, [| t2;t1 |])])) ] | Kapp(S,[t']) -> @@ -1538,24 +1546,24 @@ let nat_inject = let rec loop p t : unit Proofview.tactic = try match destructurate_term sigma t with Kapp(S,[t]) -> - (Tacticals.New.tclTHEN - (Proofview.V82.tactic (clever_rewrite_gen p + (tclTHEN + (clever_rewrite_gen p (mkApp (Lazy.force coq_Zsucc, [| mk_inj t |])) - ((Lazy.force coq_inj_S),[t]))) + ((Lazy.force coq_inj_S),[t])) (loop (P_APP 1 :: p) t)) | _ -> explore p t with e when catchable_exception e -> explore p t in - if is_number t' then Proofview.V82.tactic (focused_simpl p) else loop p t + if is_number t' then focused_simpl p else loop p t | Kapp(Pred,[t]) -> let t_minus_one = mkApp (Lazy.force coq_minus, [| t; mkApp (Lazy.force coq_S, [| Lazy.force coq_O |]) |]) in - Tacticals.New.tclTHEN - (Proofview.V82.tactic (clever_rewrite_gen_nat (P_APP 1 :: p) t_minus_one - ((Lazy.force coq_pred_of_minus),[t]))) + tclTHEN + (clever_rewrite_gen_nat (P_APP 1 :: p) t_minus_one + ((Lazy.force coq_pred_of_minus),[t])) (explore p t_minus_one) - | Kapp(O,[]) -> Proofview.V82.tactic (focused_simpl p) + | Kapp(O,[]) -> focused_simpl p | _ -> Proofview.tclUNIT () with e when catchable_exception e -> Proofview.tclUNIT () @@ -1565,7 +1573,7 @@ let nat_inject = Proofview.tclEVARMAP >>= fun sigma -> begin try match destructurate_prop sigma t with Kapp(Le,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_le, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); @@ -1574,7 +1582,7 @@ let nat_inject = (loop lit) ] | Kapp(Lt,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_lt, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); @@ -1583,7 +1591,7 @@ let nat_inject = (loop lit) ] | Kapp(Ge,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_ge, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); @@ -1592,7 +1600,7 @@ let nat_inject = (loop lit) ] | Kapp(Gt,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_gt, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); @@ -1601,7 +1609,7 @@ let nat_inject = (loop lit) ] | Kapp(Neq,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_neq, [| t1;t2;mkVar i |]) ]); (explore [P_APP 1; P_TYPE] t1); @@ -1611,7 +1619,7 @@ let nat_inject = ] | Kapp(Eq,[typ;t1;t2]) -> if is_conv typ (Lazy.force coq_nat) then - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); (explore [P_APP 2; P_TYPE] t1); @@ -1697,20 +1705,20 @@ let fresh_id avoid id gl = let onClearedName id tac = (* We cannot ensure that hyps can be cleared (because of dependencies), *) (* so renaming may be necessary *) - Tacticals.New.tclTHEN - (Tacticals.New.tclTRY (clear [id])) - (Proofview.Goal.enter { enter = begin fun gl -> + tclTHEN + (tclTRY (clear [id])) + (Proofview.Goal.nf_enter { enter = begin fun gl -> let id = fresh_id [] id gl in - Tacticals.New.tclTHEN (introduction id) (tac id) + tclTHEN (introduction id) (tac id) end }) let onClearedName2 id tac = - Tacticals.New.tclTHEN - (Tacticals.New.tclTRY (clear [id])) - (Proofview.Goal.enter { enter = begin fun gl -> + tclTHEN + (tclTRY (clear [id])) + (Proofview.Goal.nf_enter { enter = begin fun gl -> let id1 = fresh_id [] (add_suffix id "_left") gl in let id2 = fresh_id [] (add_suffix id "_right") gl in - Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] + tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end }) let rec is_Prop sigma c = match EConstr.kind sigma c with @@ -1724,7 +1732,7 @@ let destructure_hyps = let decidability = decidability gl in let pf_nf = pf_nf gl in let rec loop = function - | [] -> (Tacticals.New.tclTHEN nat_inject coq_omega) + | [] -> (tclTHEN nat_inject coq_omega) | decl::lit -> let i = NamedDecl.get_id decl in Proofview.tclEVARMAP >>= fun sigma -> @@ -1732,17 +1740,17 @@ let destructure_hyps = | Kapp(False,[]) -> elim_id i | Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit | Kapp(Or,[t1;t2]) -> - (Tacticals.New.tclTHENS + (tclTHENS (elim_id i) [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit))); onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> - Tacticals.New.tclTHEN + tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> - Tacticals.New.tclTHEN + tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit))) @@ -1752,7 +1760,7 @@ let destructure_hyps = if is_Prop sigma (type_of t2) then let d1 = decidability t1 in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); (onClearedName i (fun i -> @@ -1763,7 +1771,7 @@ let destructure_hyps = | Kapp(Not,[t]) -> begin match destructurate_prop sigma t with Kapp(Or,[t1;t2]) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); (onClearedName i (fun i -> @@ -1771,7 +1779,7 @@ let destructure_hyps = ] | Kapp(And,[t1;t2]) -> let d1 = decidability t1 in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_and, [| t1; t2; d1; mkVar i |])]); @@ -1781,7 +1789,7 @@ let destructure_hyps = | Kapp(Iff,[t1;t2]) -> let d1 = decidability t1 in let d2 = decidability t2 in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_iff, [| t1; t2; d1; d2; mkVar i |])]); @@ -1793,7 +1801,7 @@ let destructure_hyps = (* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok. For t1, being decidable implies being Prop. *) let d1 = decidability t1 in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_imp, [| t1; t2; d1; mkVar i |])]); @@ -1802,7 +1810,7 @@ let destructure_hyps = ] | Kapp(Not,[t]) -> let d = decidability t in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit)))) @@ -1810,7 +1818,7 @@ let destructure_hyps = | Kapp(op,[t1;t2]) -> (try let thm = not_binop op in - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (generalize_tac [mkApp (Lazy.force thm, [| t1;t2;mkVar i|])]); (onClearedName i (fun _ -> loop lit)) @@ -1820,14 +1828,14 @@ let destructure_hyps = if !old_style_flag then begin match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (simplest_elim (mkApp (Lazy.force coq_not_eq, [|t1;t2;mkVar i|]))); (onClearedName i (fun _ -> loop lit)) ] | Kapp(Z,_) -> - Tacticals.New.tclTHENLIST [ + tclTHENLIST [ (simplest_elim (mkApp (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); @@ -1837,12 +1845,12 @@ let destructure_hyps = end else begin match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> - (Tacticals.New.tclTHEN + (tclTHEN (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) decl)) (loop lit)) | Kapp(Z,_) -> - (Tacticals.New.tclTHEN + (tclTHEN (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) decl)) (loop lit)) @@ -1870,23 +1878,23 @@ let destructure_goal = Proofview.V82.wrap_exceptions prop >>= fun prop -> match prop with | Kapp(Not,[t]) -> - (Tacticals.New.tclTHEN - (Tacticals.New.tclTHEN (unfold sp_not) intro) + (tclTHEN + (tclTHEN (unfold sp_not) intro) destructure_hyps) - | Kimp(a,b) -> (Tacticals.New.tclTHEN intro (loop b)) + | Kimp(a,b) -> (tclTHEN intro (loop b)) | Kapp(False,[]) -> destructure_hyps | _ -> let goal_tac = try let dec = decidability t in - Tacticals.New.tclTHEN + tclTHEN (Proofview.V82.tactic (Tacmach.refine (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |])))) intro with Undecidable -> Tactics.elim_type (Lazy.force coq_False) | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in - Tacticals.New.tclTHEN goal_tac destructure_hyps + tclTHEN goal_tac destructure_hyps in (loop concl) end } -- cgit v1.2.3 From 0cc7ed04a6a6db666da08a724df3998c1e4888f9 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 8 Mar 2017 10:12:01 +0100 Subject: Porting omega to the new tactic API. --- plugins/omega/coq_omega.ml | 55 ++++++++++++++++++++++++++++++++++------------ 1 file changed, 41 insertions(+), 14 deletions(-) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index c57719ac4..92b092ffe 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -20,7 +20,7 @@ open Nameops open Term open EConstr open Tacticals.New -open Tacmach +open Tacmach.New open Tactics open Logic open Libnames @@ -154,8 +154,8 @@ let exists_tac c = constructor_tac false (Some 1) 1 (ImplicitBindings [c]) let generalize_tac t = generalize t let elim t = simplest_elim t -let exact t = Tacmach.refine t let unfold s = Tactics.unfold_in_concl [Locus.AllOccurrences, Lazy.force s] +let pf_nf gl c = pf_apply Tacred.simpl gl c let rev_assoc k = let rec loop = function @@ -583,10 +583,11 @@ let abstract_path sigma typ path t = mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur let focused_simpl path = - Proofview.V82.tactic begin fun gl -> + let open Tacmach.New in + Proofview.Goal.nf_enter { enter = begin fun gl -> let newc = context (project gl) (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in - Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl - end + convert_concl_no_check newc DEFAULTcast + end } let focused_simpl path = focused_simpl path @@ -644,12 +645,19 @@ let decompile af = in loop af.body -let mkNewMeta () = mkMeta (Evarutil.new_meta()) +(** Backward compat to emulate the old Refine: normalize the goal conclusion *) +let new_hole env sigma c = + let c = Reductionops.nf_betaiota (Sigma.to_evar_map sigma) c in + Evarutil.new_evar env sigma c let clever_rewrite_base_poly typ p result theorem = - Proofview.V82.tactic begin fun gl -> + let open Tacmach.New in + let open Sigma in + Proofview.Goal.nf_enter { enter = begin fun gl -> let full = pf_concl gl in + let env = pf_env gl in let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in + Refine.refine { run = begin fun sigma -> let t = applist (mkLambda @@ -662,8 +670,11 @@ let clever_rewrite_base_poly typ p result theorem = [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), [abstracted]) in - exact (applist(t,[mkNewMeta()])) gl - end + let argt = mkApp (abstracted, [|result|]) in + let Sigma (hole, sigma, p) = new_hole env sigma argt in + Sigma (applist (t, [hole]), sigma, p) + end } + end } let clever_rewrite_base p result theorem = clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem @@ -679,14 +690,29 @@ let clever_rewrite_gen_nat p result (t,args) = let theorem = applist(t, args) in clever_rewrite_base_nat p result theorem +(** Solve using the term the term [t _] *) +let refine_app gl t = + let open Tacmach.New in + let open Sigma in + Refine.refine { run = begin fun sigma -> + let env = pf_env gl in + let ht = match EConstr.kind (Sigma.to_evar_map sigma) (pf_get_type_of gl t) with + | Prod (_, t, _) -> t + | _ -> assert false + in + let Sigma (hole, sigma, p) = new_hole env sigma ht in + Sigma (applist (t, [hole]), sigma, p) + end } + let clever_rewrite p vpath t = - Proofview.V82.tactic begin fun gl -> + let open Tacmach.New in + Proofview.Goal.nf_enter { enter = begin fun gl -> let full = pf_concl gl in let (abstracted,occ) = abstract_path (project gl) (Lazy.force coq_Z) (List.rev p) full in let vargs = List.map (fun p -> occurrence (project gl) p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in - exact (applist(t',[mkNewMeta()])) gl - end + refine_app gl t' + end } let rec shuffle p (t1,t2) = match t1,t2 with @@ -1888,8 +1914,9 @@ let destructure_goal = try let dec = decidability t in tclTHEN - (Proofview.V82.tactic (Tacmach.refine - (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |])))) + (Proofview.Goal.nf_enter { enter = begin fun gl -> + refine_app gl (mkApp (Lazy.force coq_dec_not_not, [| t; dec |])) + end }) intro with Undecidable -> Tactics.elim_type (Lazy.force coq_False) | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e -- cgit v1.2.3 From 1ef92c718ece547826f4c7e5c1ce78a6965e1ca6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 29 Mar 2017 23:16:06 +0200 Subject: Removing trivial compatibility layer in refl_omega. --- plugins/romega/const_omega.ml | 6 ++++-- plugins/romega/const_omega.mli | 2 +- plugins/romega/g_romega.ml4 | 2 +- plugins/romega/refl_omega.ml | 29 +++++++++++++++-------------- 4 files changed, 21 insertions(+), 18 deletions(-) diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml index 5c68078d7..8d7ae51fc 100644 --- a/plugins/romega/const_omega.ml +++ b/plugins/romega/const_omega.ml @@ -285,7 +285,7 @@ module type Int = sig val mk : Bigint.bigint -> Term.constr val parse_term : Term.constr -> parse_term - val parse_rel : Proof_type.goal Tacmach.sigma -> Term.constr -> parse_rel + val parse_rel : ([ `NF ], 'r) Proofview.Goal.t -> Term.constr -> parse_rel (* check whether t is built only with numbers and + * - *) val is_scalar : Term.constr -> bool end @@ -350,10 +350,12 @@ let parse_term t = | _ -> Tother with e when Logic.catchable_exception e -> Tother +let pf_nf gl c = Tacmach.New.pf_apply Tacred.simpl gl c + let parse_rel gl t = try match destructurate t with | Kapp("eq",[typ;t1;t2]) - when destructurate (EConstr.Unsafe.to_constr (Tacmach.pf_nf gl (EConstr.of_constr typ))) = Kapp("Z",[]) -> Req (t1,t2) + when destructurate (EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr typ))) = Kapp("Z",[]) -> Req (t1,t2) | Kapp("Zne",[t1;t2]) -> Rne (t1,t2) | Kapp("Z.le",[t1;t2]) -> Rle (t1,t2) | Kapp("Z.lt",[t1;t2]) -> Rlt (t1,t2) diff --git a/plugins/romega/const_omega.mli b/plugins/romega/const_omega.mli index af50ea0ff..ee7ff451a 100644 --- a/plugins/romega/const_omega.mli +++ b/plugins/romega/const_omega.mli @@ -168,7 +168,7 @@ module type Int = (* parsing a term (one level, except if a number is found) *) val parse_term : Term.constr -> parse_term (* parsing a relation expression, including = < <= >= > *) - val parse_rel : Proof_type.goal Tacmach.sigma -> Term.constr -> parse_rel + val parse_rel : ([ `NF ], 'r) Proofview.Goal.t -> Term.constr -> parse_rel (* Is a particular term only made of numbers and + * - ? *) val is_scalar : Term.constr -> bool end diff --git a/plugins/romega/g_romega.ml4 b/plugins/romega/g_romega.ml4 index 9a54ad778..df7e5cb99 100644 --- a/plugins/romega/g_romega.ml4 +++ b/plugins/romega/g_romega.ml4 @@ -38,7 +38,7 @@ let romega_tactic l = we'd better leave as little as possible in the conclusion, for an easier decidability argument. *) (Tactics.intros) - (Proofview.V82.tactic total_reflexive_omega_tactic)) + (total_reflexive_omega_tactic)) TACTIC EXTEND romega diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml index cfe14b230..a20589fb4 100644 --- a/plugins/romega/refl_omega.ml +++ b/plugins/romega/refl_omega.ml @@ -8,6 +8,7 @@ open Pp open Util +open Proofview.Notations open Const_omega module OmegaSolver = Omega_plugin.Omega.MakeOmegaSolver (Bigint) open OmegaSolver @@ -16,13 +17,12 @@ open OmegaSolver (* Especially useful debugging functions *) let debug = ref false -let show_goal gl = - if !debug then (); Tacticals.tclIDTAC gl +let show_goal = Tacticals.New.tclIDTAC let pp i = print_int i; print_newline (); flush stdout (* More readable than the prefix notation *) -let (>>) = Tacticals.tclTHEN +let (>>) = Tacticals.New.tclTHEN let mkApp = Term.mkApp @@ -739,7 +739,7 @@ and oproposition_of_constr env ((negated,depends,origin,path) as ctxt) gl c = (* Destructuration des hypothèses et de la conclusion *) let reify_gl env gl = - let concl = Tacmach.pf_concl gl in + let concl = Tacmach.New.pf_concl gl in let concl = EConstr.Unsafe.to_constr concl in let t_concl = Pnot (oproposition_of_constr env (true,[],id_concl,[O_mono]) gl concl) in @@ -760,7 +760,7 @@ let reify_gl env gl = | [] -> if !debug then print_env_reification env; [] in - let t_lhyps = loop (Tacmach.pf_hyps_types gl) in + let t_lhyps = loop (Tacmach.New.pf_hyps_types gl) in (id_concl,t_concl) :: t_lhyps let rec destructurate_pos_hyp orig list_equations list_depends = function @@ -1283,21 +1283,22 @@ let resolution env full_reified_goal systems_list = CCHyp{o_hyp=id_concl;o_path=[]} :: hyp_stated_vars @ initial_context in let decompose_tactic = decompose_tree env context solution_tree in - Proofview.V82.of_tactic (Tactics.generalize - (l_generalize_arg @ List.map EConstr.mkVar (List.tl l_hyps))) >> - Proofview.V82.of_tactic (Tactics.change_concl reified) >> - Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (app coq_do_omega [|decompose_tactic; normalization_trace|]))) >> + Tactics.generalize + (l_generalize_arg @ List.map EConstr.mkVar (List.tl l_hyps)) >> + Tactics.change_concl reified >> + Tactics.apply (EConstr.of_constr (app coq_do_omega [|decompose_tactic; normalization_trace|])) >> show_goal >> - Proofview.V82.of_tactic (Tactics.normalise_vm_in_concl) >> + Tactics.normalise_vm_in_concl >> (*i Alternatives to the previous line: - Normalisation without VM: Tactics.normalise_in_concl - Skip the conversion check and rely directly on the QED: Tacmach.convert_concl_no_check (Lazy.force coq_True) Term.VMcast >> i*) - Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (Lazy.force coq_I))) + Tactics.apply (EConstr.of_constr (Lazy.force coq_I)) -let total_reflexive_omega_tactic gl = +let total_reflexive_omega_tactic = + Proofview.Goal.nf_enter { enter = begin fun gl -> Coqlib.check_required_library ["Coq";"romega";"ROmega"]; rst_omega_eq (); rst_omega_var (); @@ -1306,9 +1307,9 @@ let total_reflexive_omega_tactic gl = let full_reified_goal = reify_gl env gl in let systems_list = destructurate_hyps full_reified_goal in if !debug then display_systems systems_list; - resolution env full_reified_goal systems_list gl + resolution env full_reified_goal systems_list with NO_CONTRADICTION -> CErrors.error "ROmega can't solve this system" - + end } (*i let tester = Tacmach.hide_atomic_tactic "TestOmega" test_tactic i*) -- cgit v1.2.3 From d272cd02ef9ba2509c266f58ee39f51106ae53c2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 11 Apr 2017 14:27:24 +0200 Subject: Fix the API of the new pf_constr_of_global. The current implementation was still using continuation passing-style, and furthermore was triggering a focus on the goals. We take advantage of the tactic features instead. --- plugins/cc/cctac.ml | 6 +++--- plugins/firstorder/rules.ml | 9 +++++---- plugins/fourier/fourierR.ml | 4 ++-- tactics/equality.ml | 6 +++--- tactics/tacticals.ml | 14 ++++++-------- tactics/tacticals.mli | 2 +- tactics/tactics.ml | 6 +++--- 7 files changed, 23 insertions(+), 24 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 2d9dec095..b2c609dcb 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -245,7 +245,7 @@ let app_global f args k = Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) let new_app_global f args k = - Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) + Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> k (mkApp (fc, args)) let new_refine c = Proofview.V82.tactic (refine c) let refine c = refine c @@ -492,7 +492,7 @@ let congruence_tac depth l = *) let mk_eq f c1 c2 k = - Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> + Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in @@ -501,7 +501,7 @@ let mk_eq f c1 c2 k = let evm, _ = type_of (pf_env gl) evm term in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k term) - end }) + end } let f_equal = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index a60fd4d8f..96601f74a 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -14,6 +14,7 @@ open Vars open Tacmach open Tactics open Tacticals +open Proofview.Notations open Termops open Formula open Sequent @@ -96,7 +97,7 @@ let left_and_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE (tclTHENLIST - [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim); + [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim); clear_global id; tclDO n (Proofview.V82.of_tactic intro)]) (wrap n false continue seq) @@ -110,12 +111,12 @@ let left_or_tac ind backtrack id continue seq gls= tclDO n (Proofview.V82.of_tactic intro); wrap n false continue seq] in tclIFTHENSVELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim)) (Array.map f v) backtrack gls let left_false_tac id= - Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim) + Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim) (* left arrow connective rules *) @@ -183,7 +184,7 @@ let forall_tac backtrack continue seq= let left_exists_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim)) (tclTHENLIST [clear_global id; tclDO n (Proofview.V82.of_tactic intro); (wrap (n-1) false continue seq)]) diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index e11cbc279..25d8f8c83 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -617,9 +617,9 @@ let rec fourier () = [Tacticals.New.tclORELSE (* TODO : Ring.polynom []*) (Proofview.tclUNIT ()) (Proofview.tclUNIT ()); - Tacticals.New.pf_constr_of_global (cget coq_sym_eqT) (fun symeq -> + Tacticals.New.pf_constr_of_global (cget coq_sym_eqT) >>= fun symeq -> (Tacticals.New.tclTHEN (apply symeq) - (apply (get coq_Rinv_1))))] + (apply (get coq_Rinv_1)))] ) ])); diff --git a/tactics/equality.ml b/tactics/equality.ml index e68be91e9..25c28cf4a 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -658,8 +658,8 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = | Some evd -> let e = build_coq_eq () in let sym = build_coq_eq_sym () in - Tacticals.New.pf_constr_of_global sym (fun sym -> - Tacticals.New.pf_constr_of_global e (fun e -> + Tacticals.New.pf_constr_of_global sym >>= fun sym -> + Tacticals.New.pf_constr_of_global e >>= fun e -> let eq = applist (e, [t1;c1;c2]) in tclTHENLAST (replace_core clause l2r eq) @@ -667,7 +667,7 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = [assumption; tclTHEN (apply sym) assumption; try_prove_eq - ]))) + ]) end } let replace c1 c2 = diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 5c97f27ba..c8441a8cc 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -732,13 +732,11 @@ module New = struct let case_nodep_then_using = general_elim_then_using gl_make_case_nodep false - let pf_constr_of_global ref tac = - Proofview.Goal.enter { enter = begin fun gl -> - let env = Proofview.Goal.env gl in - let sigma = Tacmach.New.project gl in - let (sigma, c) = Evd.fresh_global env sigma ref in - let c = EConstr.of_constr c in - Proofview.Unsafe.tclEVARS sigma <*> (tac c) - end } + let pf_constr_of_global ref = + Proofview.tclEVARMAP >>= fun sigma -> + Proofview.tclENV >>= fun env -> + let (sigma, c) = Evd.fresh_global env sigma ref in + let c = EConstr.of_constr c in + Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT c end diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 01b9e5e93..5a4ecbac7 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -263,5 +263,5 @@ module New : sig val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic - val pf_constr_of_global : Globnames.global_reference -> (constr -> unit Proofview.tactic) -> unit Proofview.tactic + val pf_constr_of_global : Globnames.global_reference -> constr Proofview.tactic end diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 1615aec89..9c2a1d850 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4721,7 +4721,7 @@ let symmetry_red allowred = | Some eq_data,_,_ -> Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) - (Tacticals.New.pf_constr_of_global eq_data.sym apply) + (Tacticals.New.pf_constr_of_global eq_data.sym >>= apply) | None,eq,eq_kind -> prove_symmetry eq eq_kind end } @@ -4817,8 +4817,8 @@ let transitivity_red allowred t = Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) (match t with - | None -> Tacticals.New.pf_constr_of_global eq_data.trans eapply - | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans (fun trans -> apply_list [trans;t])) + | None -> Tacticals.New.pf_constr_of_global eq_data.trans >>= eapply + | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans >>= fun trans -> apply_list [trans; t]) | None,eq,eq_kind -> match t with | None -> Tacticals.New.tclZEROMSG (str"etransitivity not supported for this relation.") -- cgit v1.2.3 From 91ff75cf42ebc883e2cfcdc4928154315984beb8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 11 Apr 2017 14:51:56 +0200 Subject: Removing tactic compatibility layer in Micromega. --- plugins/micromega/coq_micromega.ml | 57 ++++++++++++++++++-------------------- 1 file changed, 27 insertions(+), 30 deletions(-) diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 4b87e6e2e..eb26c5431 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -901,16 +901,13 @@ struct coq_Qeq, Mc.OpEq ] - let has_typ gl t1 typ = - let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in - EConstr.eq_constr (Tacmach.project gl) ty typ - + type gl = { env : Environ.env; sigma : Evd.evar_map } let is_convertible gl t1 t2 = - Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2 + Reductionops.is_conv gl.env gl.sigma t1 t2 let parse_zop gl (op,args) = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in match EConstr.kind sigma op with | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) | Ind((n,0),_) -> @@ -920,7 +917,7 @@ struct | _ -> failwith "parse_zop" let parse_rop gl (op,args) = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in match EConstr.kind sigma op with | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) | Ind((n,0),_) -> @@ -930,7 +927,7 @@ struct | _ -> failwith "parse_zop" let parse_qop gl (op,args) = - (assoc_const (Tacmach.project gl) op qop_table, args.(0) , args.(1)) + (assoc_const gl.sigma op qop_table, args.(0) , args.(1)) let is_constant sigma t = (* This is an approx *) match EConstr.kind sigma t with @@ -1154,7 +1151,7 @@ struct rop_spec let parse_arith parse_op parse_expr env cstr gl = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in if debug then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr cstr ++ fnl ()); match EConstr.kind sigma cstr with @@ -1199,7 +1196,7 @@ struct *) let parse_formula gl parse_atom env tg term = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in let parse_atom env tg t = try @@ -1208,7 +1205,7 @@ struct with e when CErrors.noncritical e -> (X(t),env,tg) in let is_prop term = - let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in + let sort = Retyping.get_sort_of gl.env gl.sigma term in Sorts.is_prop sort in let rec xparse_formula env tg term = @@ -1720,7 +1717,6 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* let vm = dump_varmap (spec.typ) (vm_of_list env) in (* todo : directly generate the proof term - or generalize before conversion? *) Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in Tacticals.New.tclTHENLIST [ Tactics.change_concl @@ -1730,7 +1726,7 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* ("__varmap", vm, Term.mkApp(Lazy.force coq_VarMap, [|spec.typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl)) + (Tacmach.New.pf_concl gl)) ] end } @@ -1967,11 +1963,13 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 Some (ids,ff',res') - (** * Parse the proof environment, and call micromega_tauto *) +let fresh_id avoid id gl = + Tactics.fresh_id_in_env avoid id (Proofview.Goal.env gl) + let micromega_gen parse_arith (negate:'cst atom -> 'cst mc_cnf) @@ -1979,17 +1977,17 @@ let micromega_gen unsat deduce spec dumpexpr prover tac = Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in - let sigma = Tacmach.project gl in - let concl = Tacmach.pf_concl gl in - let hyps = Tacmach.pf_hyps_types gl in + let sigma = Tacmach.New.project gl in + let concl = Tacmach.New.pf_concl gl in + let hyps = Tacmach.New.pf_hyps_types gl in try - let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in + let gl0 = { env = Tacmach.New.pf_env gl; sigma } in + let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let dumpexpr = Lazy.force dumpexpr in - match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with + match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl0 with | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Some (ids,ff',res') -> let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma dumpexpr ff' in @@ -1998,7 +1996,7 @@ let micromega_gen let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in - let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in + let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; @@ -2057,7 +2055,6 @@ let micromega_order_changer cert env ff = let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in let vm = dump_varmap (typ) (vm_of_list env) in Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in Tacticals.New.tclTHENLIST [ (Tactics.change_concl @@ -2069,7 +2066,7 @@ let micromega_order_changer cert env ff = [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t"), [|typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl))); + (Tacmach.New.pf_concl gl))); (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*) ] end } @@ -2088,20 +2085,20 @@ let micromega_genr prover tac = dump_proof = dump_psatz coq_Q dump_q } in Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in - let sigma = Tacmach.project gl in - let concl = Tacmach.pf_concl gl in - let hyps = Tacmach.pf_hyps_types gl in + let sigma = Tacmach.New.project gl in + let concl = Tacmach.New.pf_concl gl in + let hyps = Tacmach.New.pf_hyps_types gl in try - let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in + let gl0 = { env = Tacmach.New.pf_env gl; sigma } in + let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in - match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with + match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl0 with | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Some (ids,ff',res') -> let (ff,ids) = formula_hyps_concl @@ -2114,7 +2111,7 @@ let micromega_genr prover tac = let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in - let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in + let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; -- cgit v1.2.3 From 02f8a5fc7d445ee093bc80663646cfea0a915e6d Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 11 Apr 2017 17:36:35 +0200 Subject: Removing tactic compatibility layer in congruence. --- plugins/cc/cctac.ml | 128 ++++++++++++++++++++++++++++++--------------------- plugins/cc/cctac.mli | 2 +- 2 files changed, 76 insertions(+), 54 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index b2c609dcb..00b31ccce 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -239,21 +239,43 @@ let build_projection intype (cstr:pconstructor) special default gls= (* generate an adhoc tactic following the proof tree *) -let _M =mkMeta - let app_global f args k = - Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) - -let new_app_global f args k = Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> k (mkApp (fc, args)) -let new_refine c = Proofview.V82.tactic (refine c) -let refine c = refine c +let rec gen_holes env sigma t n accu = + let open Sigma in + if Int.equal n 0 then (sigma, List.rev accu) + else match EConstr.kind sigma t with + | Prod (_, u, t) -> + let sigma = Sigma.Unsafe.of_evar_map sigma in + let Sigma (ev, sigma, _) = Evarutil.new_evar env sigma u in + let sigma = Sigma.to_evar_map sigma in + let t = EConstr.Vars.subst1 ev t in + gen_holes env sigma t (pred n) (ev :: accu) + | _ -> assert false + +let app_global_with_holes f args n = + Proofview.Goal.enter { enter = begin fun gl -> + Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> + let env = Proofview.Goal.env gl in + let concl = Proofview.Goal.concl gl in + Refine.refine { Sigma.run = begin fun sigma -> + let sigma = Sigma.to_evar_map sigma in + let t = Tacmach.New.pf_get_type_of gl fc in + let t = Termops.prod_applist sigma t (Array.to_list args) in + let ans = mkApp (fc, args) in + let (sigma, holes) = gen_holes env sigma t n [] in + let ans = applist (ans, holes) in + let evdref = ref sigma in + let () = Typing.e_check env evdref ans concl in + Sigma.Unsafe.of_pair (ans, !evdref) + end } + end } let assert_before n c = - Proofview.Goal.enter { enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let evm, _ = Tacmach.New.pf_apply type_of gl c in - Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (assert_before n c) + Sigma.Unsafe.of_pair (assert_before n c, evm) end } let refresh_type env evm ty = @@ -281,18 +303,18 @@ let rec proof_tac p : unit Proofview.tactic = let l=constr_of_term p.p_lhs and r=constr_of_term p.p_rhs in refresh_universes (type_of l) (fun typ -> - new_app_global _sym_eq [|typ;r;l;c|] exact_check) + app_global _sym_eq [|typ;r;l;c|] exact_check) | Refl t -> let lr = constr_of_term t in refresh_universes (type_of lr) (fun typ -> - new_app_global _refl_equal [|typ;constr_of_term t|] exact_check) + app_global _refl_equal [|typ;constr_of_term t|] exact_check) | Trans (p1,p2)-> let t1 = constr_of_term p1.p_lhs and t2 = constr_of_term p1.p_rhs and t3 = constr_of_term p2.p_rhs in refresh_universes (type_of t2) (fun typ -> - let prf = new_app_global _trans_eq [|typ;t1;t2;t3;_M 1;_M 2|] in - Tacticals.New.tclTHENS (prf new_refine) [(proof_tac p1);(proof_tac p2)]) + let prf = app_global_with_holes _trans_eq [|typ;t1;t2;t3;|] 2 in + Tacticals.New.tclTHENS prf [(proof_tac p1);(proof_tac p2)]) | Congr (p1,p2)-> let tf1=constr_of_term p1.p_lhs and tx1=constr_of_term p2.p_lhs @@ -303,18 +325,18 @@ let rec proof_tac p : unit Proofview.tactic = refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx -> let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in let appx1 = mkLambda(Name id,typf,mkApp(mkRel 1,[|tx1|])) in - let lemma1 = app_global _f_equal [|typf;typfx;appx1;tf1;tf2;_M 1|] in - let lemma2 = app_global _f_equal [|typx;typfx;tf2;tx1;tx2;_M 1|] in + let lemma1 = app_global_with_holes _f_equal [|typf;typfx;appx1;tf1;tf2|] 1 in + let lemma2 = app_global_with_holes _f_equal [|typx;typfx;tf2;tx1;tx2|] 1 in let prf = - app_global _trans_eq + app_global_with_holes _trans_eq [|typfx; mkApp(tf1,[|tx1|]); mkApp(tf2,[|tx1|]); - mkApp(tf2,[|tx2|]);_M 2;_M 3|] in - Tacticals.New.tclTHENS (Proofview.V82.tactic (prf refine)) - [Tacticals.New.tclTHEN (Proofview.V82.tactic (lemma1 refine)) (proof_tac p1); + mkApp(tf2,[|tx2|])|] 2 in + Tacticals.New.tclTHENS prf + [Tacticals.New.tclTHEN lemma1 (proof_tac p1); Tacticals.New.tclFIRST - [Tacticals.New.tclTHEN (Proofview.V82.tactic (lemma2 refine)) (proof_tac p2); + [Tacticals.New.tclTHEN lemma2 (proof_tac p2); reflexivity; Tacticals.New.tclZEROMSG (Pp.str @@ -330,8 +352,8 @@ let rec proof_tac p : unit Proofview.tactic = build_projection intype cstr special default gl in let injt= - app_global _f_equal [|intype;outtype;proj;ti;tj;_M 1|] in - Tacticals.New.tclTHEN (Proofview.V82.tactic (injt refine)) (proof_tac prf))) + app_global_with_holes _f_equal [|intype;outtype;proj;ti;tj|] 1 in + Tacticals.New.tclTHEN injt (proof_tac prf))) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e end } @@ -341,27 +363,29 @@ let refute_tac c t1 t2 p = let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let false_t=mkApp (c,[|mkVar hid|]) in let k intype = - let neweq= new_app_global _eq [|intype;tt1;tt2|] in + let neweq= app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; simplest_elim false_t] in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt1) k end } -let refine_exact_check c gl = - let evm, _ = pf_apply type_of gl c in - Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl +let refine_exact_check c = + Proofview.Goal.s_enter { s_enter = begin fun gl -> + let evm, _ = Tacmach.New.pf_apply type_of gl c in + Sigma.Unsafe.of_pair (exact_check c, evm) + end } let convert_to_goal_tac c t1 t2 p = Proofview.Goal.enter { enter = begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let k sort = - let neweq= new_app_global _eq [|sort;tt1;tt2|] in + let neweq= app_global _eq [|sort;tt1;tt2|] in let e = Tacmach.New.pf_get_new_id (Id.of_string "e") gl in let x = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in let identity=mkLambda (Name x,sort,mkRel 1) in - let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in + let endt = app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) - [proof_tac p; Proofview.V82.tactic (endt refine_exact_check)] + [proof_tac p; endt refine_exact_check] in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt2) k end } @@ -392,27 +416,25 @@ let discriminate_tac (cstr,u as cstru) p = let pred = mkLambda(Name xid,outtype,mkRel 1) in let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let proj = build_projection intype cstru trivial concl gl in - let injt=app_global _f_equal + let injt = app_global _f_equal [|intype;outtype;proj;t1;t2;mkVar hid|] in let endt k = injt (fun injt -> app_global _eq_rect [|outtype;trivial;pred;identity;concl;injt|] k) in - let neweq=new_app_global _eq [|intype;t1;t2|] in + let neweq = app_global _eq [|intype;t1;t2|] in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm) (Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) - [proof_tac p; Proofview.V82.tactic (endt refine_exact_check)]) + [proof_tac p; endt refine_exact_check]) end } (* wrap everything *) -let build_term_to_complete uf meta pac = +let build_term_to_complete uf pac = let cinfo = get_constructor_info uf pac.cnode in - let real_args = List.map (fun i -> constr_of_term (term uf i)) pac.args in - let dummy_args = List.rev (List.init pac.arity meta) in - let all_args = List.rev_append real_args dummy_args in + let real_args = List.rev_map (fun i -> constr_of_term (term uf i)) pac.args in let (kn, u) = cinfo.ci_constr in - applist (mkConstructU (kn, EInstance.make u), all_args) + (applist (mkConstructU (kn, EInstance.make u), real_args), pac.arity) let cc_tactic depth additionnal_terms = Proofview.Goal.enter { enter = begin fun gl -> @@ -434,16 +456,17 @@ let cc_tactic depth additionnal_terms = let cstr=(get_constructor_info uf ipac.cnode).ci_constr in discriminate_tac cstr p | Incomplete -> + let open Glob_term in let env = Proofview.Goal.env gl in - let metacnt = ref 0 in - let newmeta _ = incr metacnt; _M !metacnt in - let terms_to_complete = - List.map - (build_term_to_complete uf newmeta) - (epsilons uf) in + let terms_to_complete = List.map (build_term_to_complete uf) (epsilons uf) in + let hole = GHole (Loc.ghost, Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None) in + let pr_missing (c, missing) = + let c = Detyping.detype ~lax:true false [] env sigma c in + let holes = List.init missing (fun _ -> hole) in + Printer.pr_glob_constr_env env (GApp (Loc.ghost, c, holes)) + in Feedback.msg_info - (Pp.str "Goal is solvable by congruence but \ - some arguments are missing."); + (Pp.str "Goal is solvable by congruence but some arguments are missing."); Feedback.msg_info (Pp.str " Try " ++ hov 8 @@ -451,7 +474,7 @@ let cc_tactic depth additionnal_terms = str "\"congruence with (" ++ prlist_with_sep (fun () -> str ")" ++ spc () ++ str "(") - (Termops.print_constr_env env sigma) + pr_missing terms_to_complete ++ str ")\"," end ++ @@ -472,13 +495,13 @@ let cc_tactic depth additionnal_terms = convert_to_hyp_tac ida ta idb tb p end } -let cc_fail gls = - user_err ~hdr:"Congruence" (Pp.str "congruence failed.") +let cc_fail = + Tacticals.New.tclZEROMSG (Pp.str "congruence failed.") let congruence_tac depth l = Tacticals.New.tclORELSE (Tacticals.New.tclTHEN (Tacticals.New.tclREPEAT introf) (cc_tactic depth l)) - (Proofview.V82.tactic cc_fail) + cc_fail (* Beware: reflexivity = constructor 1 = apply refl_equal might be slow now, let's rather do something equivalent @@ -493,14 +516,13 @@ let congruence_tac depth l = let mk_eq f c1 c2 k = Tacticals.New.pf_constr_of_global (Lazy.force f) >>= fun fc -> - Proofview.Goal.enter { enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in let term = mkApp (fc, [| ty; c1; c2 |]) in let evm, _ = type_of (pf_env gl) evm term in - Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) - (k term) + Sigma.Unsafe.of_pair (k term, evm) end } let f_equal = @@ -511,7 +533,7 @@ let f_equal = try (* type_of can raise an exception *) Tacticals.New.tclTHENS (mk_eq _eq c1 c2 Tactics.cut) - [Proofview.tclUNIT ();Tacticals.New.tclTRY ((new_app_global _refl_equal [||]) apply)] + [Proofview.tclUNIT ();Tacticals.New.tclTRY ((app_global _refl_equal [||]) apply)] with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in Proofview.tclORELSE diff --git a/plugins/cc/cctac.mli b/plugins/cc/cctac.mli index de6eb982e..5099d847b 100644 --- a/plugins/cc/cctac.mli +++ b/plugins/cc/cctac.mli @@ -14,7 +14,7 @@ val proof_tac: Ccproof.proof -> unit Proofview.tactic val cc_tactic : int -> constr list -> unit Proofview.tactic -val cc_fail : tactic +val cc_fail : unit Proofview.tactic val congruence_tac : int -> constr list -> unit Proofview.tactic -- cgit v1.2.3 From 7d26940665ccce2e4ee1ba6fc157e42f7a639861 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 12 Apr 2017 12:15:40 +0200 Subject: Removing tactic compatibility layer in Command. --- vernac/command.ml | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) diff --git a/vernac/command.ml b/vernac/command.ml index 45ff57955..b27d8a0a3 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -1198,11 +1198,8 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) fixnames fixtypes fiximps in let init_tac = - Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC) + Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC) fixdefs) in - let init_tac = - Option.map (List.map Proofview.V82.tactic) init_tac - in let evd = Evd.from_ctx ctx in Lemmas.start_proof_with_initialization (Global,poly,DefinitionBody Fixpoint) evd (Some(false,indexes,init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ())) @@ -1235,11 +1232,8 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) fixnames fixtypes fiximps in let init_tac = - Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC) + Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC) fixdefs) in - let init_tac = - Option.map (List.map Proofview.V82.tactic) init_tac - in let evd = Evd.from_ctx ctx in Lemmas.start_proof_with_initialization (Global,poly, DefinitionBody CoFixpoint) evd (Some(true,[],init_tac)) thms None (Lemmas.mk_hook (fun _ _ -> ())) -- cgit v1.2.3 From 95b669a96f143d930b228a8b04041457040dd984 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 12 Apr 2017 13:48:36 +0200 Subject: Removing compatibility layer in Leminv. --- tactics/leminv.ml | 15 ++++++++------- 1 file changed, 8 insertions(+), 7 deletions(-) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index daa962f1d..83f3da30a 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -260,22 +260,23 @@ let add_inversion_lemma_exn na com comsort bool tac = (* Applying a given inversion lemma *) (* ================================= *) -let lemInv id c gls = +let lemInv id c = + Proofview.Goal.enter { enter = begin fun gls -> try - let open Tacmach in let clause = mk_clenv_from_env (pf_env gls) (project gls) None (c, pf_unsafe_type_of gls c) in let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in - Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls + Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false with | NoSuchBinding -> user_err - (hov 0 (pr_econstr_env (Refiner.pf_env gls) (Refiner.project gls) c ++ spc () ++ str "does not refer to an inversion lemma.")) + (hov 0 (pr_econstr_env (pf_env gls) (project gls) c ++ spc () ++ str "does not refer to an inversion lemma.")) | UserError (a,b) -> user_err ~hdr:"LemInv" (str "Cannot refine current goal with the lemma " ++ - pr_leconstr_env (Refiner.pf_env gls) (Refiner.project gls) c) + pr_leconstr_env (pf_env gls) (project gls) c) + end } -let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id +let lemInv_gen id c = try_intros_until (fun id -> lemInv id c) id let lemInvIn id c ids = Proofview.Goal.enter { enter = begin fun gl -> @@ -289,7 +290,7 @@ let lemInvIn id c ids = else (tclTHEN (tclDO nb_of_new_hyp intro) (intros_replacing ids)) in - ((tclTHEN (tclTHEN (bring_hyps hyps) (Proofview.V82.tactic (lemInv id c))) + ((tclTHEN (tclTHEN (bring_hyps hyps) (lemInv id c)) (intros_replace_ids))) end } -- cgit v1.2.3 From a1fd5fb489237a1300adb242e2c9b6c74c82981b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 12 Apr 2017 14:16:04 +0200 Subject: Porting the firstorder plugin to the new tactic API. --- plugins/firstorder/formula.ml | 105 ++++++++++++------------ plugins/firstorder/formula.mli | 13 +-- plugins/firstorder/g_ground.ml4 | 39 +++++---- plugins/firstorder/ground.ml | 38 +++++---- plugins/firstorder/ground.mli | 4 +- plugins/firstorder/instances.ml | 112 ++++++++++++++------------ plugins/firstorder/instances.mli | 4 +- plugins/firstorder/rules.ml | 169 ++++++++++++++++++++++----------------- plugins/firstorder/rules.mli | 3 + plugins/firstorder/sequent.ml | 61 +++++++------- plugins/firstorder/sequent.mli | 30 +++---- plugins/firstorder/unify.ml | 58 +++++++------- plugins/firstorder/unify.mli | 7 +- 13 files changed, 346 insertions(+), 297 deletions(-) diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 7773f6a2f..ade94e98e 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -9,6 +9,7 @@ open Hipattern open Names open Term +open EConstr open Vars open Termops open Tacmach @@ -44,28 +45,27 @@ let rec nb_prod_after n c= 1+(nb_prod_after 0 b) | _ -> 0 -let construct_nhyps ind gls = +let construct_nhyps env ind = let nparams = (fst (Global.lookup_inductive (fst ind))).mind_nparams in - let constr_types = Inductiveops.arities_of_constructors (pf_env gls) ind in + let constr_types = Inductiveops.arities_of_constructors env ind in let hyp = nb_prod_after nparams in Array.map hyp constr_types (* indhyps builds the array of arrays of constructor hyps for (ind largs)*) -let ind_hyps nevar ind largs gls= - let types= Inductiveops.arities_of_constructors (pf_env gls) ind in +let ind_hyps env sigma nevar ind largs = + let types= Inductiveops.arities_of_constructors env ind in let myhyps t = - let t1=Term.prod_applist t largs in - let t2=snd (decompose_prod_n_assum nevar t1) in - fst (decompose_prod_assum t2) in + let t = EConstr.of_constr t in + let t1=Termops.prod_applist sigma t largs in + let t2=snd (decompose_prod_n_assum sigma nevar t1) in + fst (decompose_prod_assum sigma t2) in Array.map myhyps types -let special_nf gl= - let infos=CClosure.create_clos_infos !red_flags (pf_env gl) in - (fun t -> CClosure.norm_val infos (CClosure.inject t)) +let special_nf env sigma t = + Reductionops.clos_norm_flags !red_flags env sigma t -let special_whd gl= - let infos=CClosure.create_clos_infos !red_flags (pf_env gl) in - (fun t -> CClosure.whd_val infos (CClosure.inject t)) +let special_whd env sigma t = + Reductionops.clos_whd_flags !red_flags env sigma t type kind_of_formula= Arrow of constr*constr @@ -78,20 +78,19 @@ type kind_of_formula= let pop t = Vars.lift (-1) t -let kind_of_formula gl term = - let normalize=special_nf gl in - let cciterm=special_whd gl term in - match match_with_imp_term (project gl) (EConstr.of_constr cciterm) with - Some (a,b)-> Arrow(EConstr.Unsafe.to_constr a,(pop (EConstr.Unsafe.to_constr b))) +let kind_of_formula env sigma term = + let normalize = special_nf env sigma in + let cciterm = special_whd env sigma term in + match match_with_imp_term sigma cciterm with + Some (a,b)-> Arrow (a, pop b) |_-> - match match_with_forall_term (project gl) (EConstr.of_constr cciterm) with - Some (_,a,b)-> Forall(EConstr.Unsafe.to_constr a,EConstr.Unsafe.to_constr b) + match match_with_forall_term sigma cciterm with + Some (_,a,b)-> Forall (a, b) |_-> - match match_with_nodep_ind (project gl) (EConstr.of_constr cciterm) with + match match_with_nodep_ind sigma cciterm with Some (i,l,n)-> - let l = List.map EConstr.Unsafe.to_constr l in - let ind,u=EConstr.destInd (project gl) i in - let u = EConstr.EInstance.kind (project gl) u in + let ind,u=EConstr.destInd sigma i in + let u = EConstr.EInstance.kind sigma u in let (mib,mip) = Global.lookup_inductive ind in let nconstr=Array.length mip.mind_consnames in if Int.equal nconstr 0 then @@ -100,7 +99,7 @@ let kind_of_formula gl term = let has_realargs=(n>0) in let is_trivial= let is_constant c = - Int.equal (nb_prod (project gl) (EConstr.of_constr c)) mib.mind_nparams in + Int.equal (nb_prod sigma (EConstr.of_constr c)) mib.mind_nparams in Array.exists is_constant mip.mind_nf_lc in if Inductiveops.mis_is_recursive (ind,mib,mip) || (has_realargs && not is_trivial) @@ -112,11 +111,11 @@ let kind_of_formula gl term = else Or((ind,u),l,is_trivial) | _ -> - match match_with_sigma_type (project gl) (EConstr.of_constr cciterm) with + match match_with_sigma_type sigma cciterm with Some (i,l)-> - let (ind, u) = EConstr.destInd (project gl) i in - let u = EConstr.EInstance.kind (project gl) u in - Exists((ind, u), List.map EConstr.Unsafe.to_constr l) + let (ind, u) = EConstr.destInd sigma i in + let u = EConstr.EInstance.kind sigma u in + Exists((ind, u), l) |_-> Atom (normalize cciterm) type atoms = {positive:constr list;negative:constr list} @@ -127,29 +126,29 @@ let no_atoms = (false,{positive=[];negative=[]}) let dummy_id=VarRef (Id.of_string "_") (* "_" cannot be parsed *) -let build_atoms gl metagen side cciterm = +let build_atoms env sigma metagen side cciterm = let trivial =ref false and positive=ref [] and negative=ref [] in - let normalize=special_nf gl in - let rec build_rec env polarity cciterm= - match kind_of_formula gl cciterm with + let normalize=special_nf env sigma in + let rec build_rec subst polarity cciterm= + match kind_of_formula env sigma cciterm with False(_,_)->if not polarity then trivial:=true | Arrow (a,b)-> - build_rec env (not polarity) a; - build_rec env polarity b + build_rec subst (not polarity) a; + build_rec subst polarity b | And(i,l,b) | Or(i,l,b)-> if b then begin - let unsigned=normalize (substnl env 0 cciterm) in + let unsigned=normalize (substnl subst 0 cciterm) in if polarity then positive:= unsigned :: !positive else negative:= unsigned :: !negative end; - let v = ind_hyps 0 i l gl in + let v = ind_hyps env sigma 0 i l in let g i _ decl = - build_rec env polarity (lift i (RelDecl.get_type decl)) in + build_rec subst polarity (lift i (RelDecl.get_type decl)) in let f l = List.fold_left_i g (1-(List.length l)) () l in if polarity && (* we have a constant constructor *) @@ -158,16 +157,16 @@ let build_atoms gl metagen side cciterm = Array.iter f v | Exists(i,l)-> let var=mkMeta (metagen true) in - let v =(ind_hyps 1 i l gl).(0) in + let v =(ind_hyps env sigma 1 i l).(0) in let g i _ decl = - build_rec (var::env) polarity (lift i (RelDecl.get_type decl)) in + build_rec (var::subst) polarity (lift i (RelDecl.get_type decl)) in List.fold_left_i g (2-(List.length l)) () v | Forall(_,b)-> let var=mkMeta (metagen true) in - build_rec (var::env) polarity b + build_rec (var::subst) polarity b | Atom t-> - let unsigned=substnl env 0 t in - if not (isMeta unsigned) then (* discarding wildcard atoms *) + let unsigned=substnl subst 0 t in + if not (isMeta sigma unsigned) then (* discarding wildcard atoms *) if polarity then positive:= unsigned :: !positive else @@ -177,9 +176,9 @@ let build_atoms gl metagen side cciterm = Concl -> build_rec [] true cciterm | Hyp -> build_rec [] false cciterm | Hint -> - let rels,head=decompose_prod cciterm in - let env=List.rev_map (fun _->mkMeta (metagen true)) rels in - build_rec env false head;trivial:=false (* special for hints *) + let rels,head=decompose_prod sigma cciterm in + let subst=List.rev_map (fun _->mkMeta (metagen true)) rels in + build_rec subst false head;trivial:=false (* special for hints *) end; (!trivial, {positive= !positive; @@ -215,32 +214,32 @@ type t={id:global_reference; pat:(left_pattern,right_pattern) sum; atoms:atoms} -let build_formula side nam typ gl metagen= - let normalize = special_nf gl in +let build_formula env sigma side nam typ metagen= + let normalize = special_nf env sigma in try let m=meta_succ(metagen false) in let trivial,atoms= if !qflag then - build_atoms gl metagen side typ + build_atoms env sigma metagen side typ else no_atoms in let pattern= match side with Concl -> let pat= - match kind_of_formula gl typ with + match kind_of_formula env sigma typ with False(_,_) -> Rfalse | Atom a -> raise (Is_atom a) | And(_,_,_) -> Rand | Or(_,_,_) -> Ror | Exists (i,l) -> - let d = RelDecl.get_type (List.last (ind_hyps 0 i l gl).(0)) in + let d = RelDecl.get_type (List.last (ind_hyps env sigma 0 i l).(0)) in Rexists(m,d,trivial) | Forall (_,a) -> Rforall | Arrow (a,b) -> Rarrow in Right pat | _ -> let pat= - match kind_of_formula gl typ with + match kind_of_formula env sigma typ with False(i,_) -> Lfalse | Atom a -> raise (Is_atom a) | And(i,_,b) -> @@ -257,7 +256,7 @@ let build_formula side nam typ gl metagen= | Arrow (a,b) -> let nfa=normalize a in LA (nfa, - match kind_of_formula gl a with + match kind_of_formula env sigma a with False(i,l)-> LLfalse(i,l) | Atom t-> LLatom | And(i,l,_)-> LLand(i,l) diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli index 5db8ff59a..3f438c04a 100644 --- a/plugins/firstorder/formula.mli +++ b/plugins/firstorder/formula.mli @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Globnames val qflag : bool ref @@ -23,10 +24,10 @@ type ('a,'b) sum = Left of 'a | Right of 'b type counter = bool -> metavariable -val construct_nhyps : pinductive -> Proof_type.goal Tacmach.sigma -> int array +val construct_nhyps : Environ.env -> pinductive -> int array -val ind_hyps : int -> pinductive -> constr list -> - Proof_type.goal Tacmach.sigma -> Context.Rel.t array +val ind_hyps : Environ.env -> Evd.evar_map -> int -> pinductive -> + constr list -> EConstr.rel_context array type atoms = {positive:constr list;negative:constr list} @@ -34,7 +35,7 @@ type side = Hyp | Concl | Hint val dummy_id: global_reference -val build_atoms : Proof_type.goal Tacmach.sigma -> counter -> +val build_atoms : Environ.env -> Evd.evar_map -> counter -> side -> constr -> bool * atoms type right_pattern = @@ -69,6 +70,6 @@ type t={id: global_reference; (*exception Is_atom of constr*) -val build_formula : side -> global_reference -> types -> - Proof_type.goal Tacmach.sigma -> counter -> (t,types) sum +val build_formula : Environ.env -> Evd.evar_map -> side -> global_reference -> types -> + counter -> (t,types) sum diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 index 3c0319319..8ef6a09d0 100644 --- a/plugins/firstorder/g_ground.ml4 +++ b/plugins/firstorder/g_ground.ml4 @@ -13,7 +13,9 @@ open Formula open Sequent open Ground open Goptions -open Tacticals +open Tacmach.New +open Tacticals.New +open Proofview.Notations open Tacinterp open Libnames open Stdarg @@ -81,21 +83,29 @@ END let fail_solver=tclFAIL 0 (Pp.str "GTauto failed") -let gen_ground_tac flag taco ids bases gl= +let gen_ground_tac flag taco ids bases = let backup= !qflag in - try + Proofview.tclOR begin + Proofview.Goal.enter { enter = begin fun gl -> qflag:=flag; let solver= match taco with Some tac-> tac | None-> snd (default_solver ()) in - let startseq gl= + let startseq k = + Proofview.Goal.s_enter { s_enter = begin fun gl -> let seq=empty_seq !ground_depth in - let seq,gl = extend_with_ref_list ids seq gl in - extend_with_auto_hints bases seq gl in - let result=ground_tac (Proofview.V82.of_tactic solver) startseq gl in - qflag:=backup;result - with reraise -> qflag:=backup;raise reraise + let seq, sigma = extend_with_ref_list (pf_env gl) (project gl) ids seq in + let seq, sigma = extend_with_auto_hints (pf_env gl) (project gl) bases seq in + Sigma.Unsafe.of_pair (k seq, sigma) + end } + in + let result=ground_tac solver startseq in + qflag := backup; + result + end } + end + (fun (e, info) -> qflag := backup; Proofview.tclZERO ~info e) (* special for compatibility with Intuition @@ -144,18 +154,15 @@ END TACTIC EXTEND firstorder [ "firstorder" tactic_opt(t) firstorder_using(l) ] -> - [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) l []) ] + [ gen_ground_tac true (Option.map (tactic_of_value ist) t) l [] ] | [ "firstorder" tactic_opt(t) "with" ne_preident_list(l) ] -> - [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l) ] + [ gen_ground_tac true (Option.map (tactic_of_value ist) t) [] l ] | [ "firstorder" tactic_opt(t) firstorder_using(l) "with" ne_preident_list(l') ] -> - [ Proofview.V82.tactic (gen_ground_tac true (Option.map (tactic_of_value ist) t) l l') ] + [ gen_ground_tac true (Option.map (tactic_of_value ist) t) l l' ] END TACTIC EXTEND gintuition [ "gintuition" tactic_opt(t) ] -> - [ Proofview.V82.tactic (gen_ground_tac false (Option.map (tactic_of_value ist) t) [] []) ] + [ gen_ground_tac false (Option.map (tactic_of_value ist) t) [] [] ] END - -open Proofview.Notations -open Cc_plugin diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml index d6cd7e2a0..ab1dd07c1 100644 --- a/plugins/firstorder/ground.ml +++ b/plugins/firstorder/ground.ml @@ -12,8 +12,9 @@ open Sequent open Rules open Instances open Term -open Tacmach -open Tacticals +open Tacmach.New +open Tacticals.New +open Proofview.Notations let update_flags ()= let predref=ref Names.Cpred.empty in @@ -29,18 +30,24 @@ let update_flags ()= CClosure.betaiotazeta (Names.Id.Pred.full,Names.Cpred.complement !predref) -let ground_tac solver startseq gl= +let ground_tac solver startseq = + Proofview.Goal.enter { enter = begin fun gl -> update_flags (); - let rec toptac skipped seq gl= - if Tacinterp.get_debug()=Tactic_debug.DebugOn 0 - then Feedback.msg_debug (Printer.pr_goal gl); + let rec toptac skipped seq = + Proofview.Goal.enter { enter = begin fun gl -> + let () = + if Tacinterp.get_debug()=Tactic_debug.DebugOn 0 + then + let gl = { Evd.it = Proofview.Goal.goal (Proofview.Goal.assume gl); sigma = project gl } in + Feedback.msg_debug (Printer.pr_goal gl) + in tclORELSE (axiom_tac seq.gl seq) begin try - let (hd,seq1)=take_formula seq - and re_add s=re_add_formula_list skipped s in + let (hd,seq1)=take_formula (project gl) seq + and re_add s=re_add_formula_list (project gl) skipped s in let continue=toptac [] - and backtrack gl=toptac (hd::skipped) seq1 gl in + and backtrack =toptac (hd::skipped) seq1 in match hd.pat with Right rpat-> begin @@ -60,7 +67,7 @@ let ground_tac solver startseq gl= or_tac backtrack continue (re_add seq1) | Rfalse->backtrack | Rexists(i,dom,triv)-> - let (lfp,seq2)=collect_quantified seq in + let (lfp,seq2)=collect_quantified (project gl) seq in let backtrack2=toptac (lfp@skipped) seq2 in if !qflag && seq.depth>0 then quantified_tac lfp backtrack2 @@ -80,7 +87,7 @@ let ground_tac solver startseq gl= left_or_tac ind backtrack hd.id continue (re_add seq1) | Lforall (_,_,_)-> - let (lfp,seq2)=collect_quantified seq in + let (lfp,seq2)=collect_quantified (project gl) seq in let backtrack2=toptac (lfp@skipped) seq2 in if !qflag && seq.depth>0 then quantified_tac lfp backtrack2 @@ -119,7 +126,8 @@ let ground_tac solver startseq gl= ll_atom_tac typ la_tac hd.id continue (re_add seq1) end with Heap.EmptyHeap->solver - end gl in - let seq, gl' = startseq gl in - wrap (List.length (pf_hyps gl)) true (toptac []) seq gl' - + end + end } in + let n = List.length (Proofview.Goal.hyps gl) in + startseq (fun seq -> wrap n true (toptac []) seq) + end } diff --git a/plugins/firstorder/ground.mli b/plugins/firstorder/ground.mli index b5669463c..4fd1e38a2 100644 --- a/plugins/firstorder/ground.mli +++ b/plugins/firstorder/ground.mli @@ -6,6 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -val ground_tac: Tacmach.tactic -> - (Proof_type.goal Tacmach.sigma -> Sequent.t * Proof_type.goal Tacmach.sigma) -> Tacmach.tactic +val ground_tac: unit Proofview.tactic -> + ((Sequent.t -> unit Proofview.tactic) -> unit Proofview.tactic) -> unit Proofview.tactic diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 9dc2a51a6..62f811546 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -11,10 +11,12 @@ open Rules open CErrors open Util open Term +open EConstr open Vars -open Tacmach +open Tacmach.New open Tactics -open Tacticals +open Tacticals.New +open Proofview.Notations open Termops open Reductionops open Formula @@ -25,11 +27,12 @@ open Sigma.Notations open Context.Rel.Declaration let compare_instance inst1 inst2= + let cmp c1 c2 = OrderedConstr.compare (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in match inst1,inst2 with Phantom(d1),Phantom(d2)-> - (OrderedConstr.compare d1 d2) + (cmp d1 d2) | Real((m1,c1),n1),Real((m2,c2),n2)-> - ((-) =? (-) ==? OrderedConstr.compare) m2 m1 n1 n2 c1 c2 + ((-) =? (-) ==? cmp) m2 m1 n1 n2 c1 c2 | Phantom(_),Real((m,_),_)-> if Int.equal m 0 then -1 else 1 | Real((m,_),_),Phantom(_)-> if Int.equal m 0 then 1 else -1 @@ -56,12 +59,12 @@ let make_simple_atoms seq= | None->[] in {negative=seq.latoms;positive=ratoms} -let do_sequent setref triv id seq i dom atoms= +let do_sequent sigma setref triv id seq i dom atoms= let flag=ref true in let phref=ref triv in let do_atoms a1 a2 = let do_pair t1 t2 = - match unif_atoms i dom t1 t2 with + match unif_atoms sigma i dom t1 t2 with None->() | Some (Phantom _) ->phref:=true | Some c ->flag:=false;setref:=IS.add (c,id) !setref in @@ -71,26 +74,26 @@ let do_sequent setref triv id seq i dom atoms= do_atoms atoms (make_simple_atoms seq); !flag && !phref -let match_one_quantified_hyp setref seq lf= +let match_one_quantified_hyp sigma setref seq lf= match lf.pat with Left(Lforall(i,dom,triv))|Right(Rexists(i,dom,triv))-> - if do_sequent setref triv lf.id seq i dom lf.atoms then + if do_sequent sigma setref triv lf.id seq i dom lf.atoms then setref:=IS.add ((Phantom dom),lf.id) !setref | _ -> anomaly (Pp.str "can't happen") -let give_instances lf seq= +let give_instances sigma lf seq= let setref=ref IS.empty in - List.iter (match_one_quantified_hyp setref seq) lf; + List.iter (match_one_quantified_hyp sigma setref seq) lf; IS.elements !setref (* collector for the engine *) -let rec collect_quantified seq= +let rec collect_quantified sigma seq= try - let hd,seq1=take_formula seq in + let hd,seq1=take_formula sigma seq in (match hd.pat with Left(Lforall(_,_,_)) | Right(Rexists(_,_,_)) -> - let (q,seq2)=collect_quantified seq1 in + let (q,seq2)=collect_quantified sigma seq1 in ((hd::q),seq2) | _->[],seq) with Heap.EmptyHeap -> [],seq @@ -99,60 +102,61 @@ let rec collect_quantified seq= let dummy_bvid=Id.of_string "x" -let mk_open_instance id idc gl m t= - let env=pf_env gl in - let evmap=Refiner.project gl in +let mk_open_instance env evmap id idc m t = let var_id= if id==dummy_id then dummy_bvid else - let typ=pf_unsafe_type_of gl idc in + let typ=Typing.unsafe_type_of env evmap idc in (* since we know we will get a product, reduction is not too expensive *) - let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap typ)) in + let (nam,_,_)=destProd evmap (whd_all env evmap typ) in match nam with Name id -> id | Anonymous -> dummy_bvid in let revt=substl (List.init m (fun i->mkRel (m-i))) t in let rec aux n avoid env evmap decls = if Int.equal n 0 then evmap, decls else - let nid=(fresh_id avoid var_id gl) in + let nid=(fresh_id_in_env avoid var_id env) in let evmap = Sigma.Unsafe.of_evar_map evmap in let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in let evmap = Sigma.to_evar_map evmap in let decl = LocalAssum (Name nid, c) in aux (n-1) (nid::avoid) (EConstr.push_rel decl env) evmap (decl::decls) in let evmap, decls = aux m [] env evmap [] in - evmap, decls, revt + (evmap, decls, revt) (* tactics *) let left_instance_tac (inst,id) continue seq= let open EConstr in + Proofview.Goal.enter { enter = begin fun gl -> + let sigma = project gl in match inst with Phantom dom-> - if lookup (id,None) seq then + if lookup sigma (id,None) seq then tclFAIL 0 (Pp.str "already done") else - tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom))) + tclTHENS (cut dom) [tclTHENLIST - [Proofview.V82.of_tactic introf; - pf_constr_of_global id (fun idc -> - (fun gls-> Proofview.V82.of_tactic (generalize - [mkApp(idc, - [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|])]) gls)); - Proofview.V82.of_tactic introf; + [introf; + (pf_constr_of_global id >>= fun idc -> + Proofview.Goal.enter { enter = begin fun gl -> + let id0 = List.nth (pf_ids_of_hyps gl) 0 in + generalize [mkApp(idc, [|mkVar id0|])] + end }); + introf; tclSOLVE [wrap 1 false continue (deepen (record (id,None) seq))]]; - tclTRY (Proofview.V82.of_tactic assumption)] - | Real((m,t) as c,_)-> - if lookup (id,Some c) seq then + tclTRY assumption] + | Real((m,t),_)-> + let c = (m, EConstr.to_constr sigma t) in + if lookup sigma (id,Some c) seq then tclFAIL 0 (Pp.str "already done") else let special_generalize= if m>0 then - pf_constr_of_global id (fun idc -> - fun gl-> - let evmap,rc,ot = mk_open_instance id idc gl m t in - let ot = EConstr.of_constr ot in + (pf_constr_of_global id >>= fun idc -> + Proofview.Goal.s_enter { s_enter = begin fun gl-> + let (evmap, rc, ot) = mk_open_instance (pf_env gl) (project gl) id idc m t in let gt= it_mkLambda_or_LetIn (mkApp(idc,[|ot|])) rc in @@ -160,34 +164,38 @@ let left_instance_tac (inst,id) continue seq= try Typing.type_of (pf_env gl) evmap gt with e when CErrors.noncritical e -> error "Untypable instance, maybe higher-order non-prenex quantification" in - tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl) + Sigma.Unsafe.of_pair (generalize [gt], evmap) + end }) else - let t = EConstr.of_constr t in - pf_constr_of_global id (fun idc -> - Proofview.V82.of_tactic (generalize [mkApp(idc,[|t|])])) + pf_constr_of_global id >>= fun idc -> generalize [mkApp(idc,[|t|])] in tclTHENLIST [special_generalize; - Proofview.V82.of_tactic introf; + introf; tclSOLVE [wrap 1 false continue (deepen (record (id,Some c) seq))]] + end } let right_instance_tac inst continue seq= + let open EConstr in + Proofview.Goal.enter { enter = begin fun gl -> match inst with Phantom dom -> - tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom))) + tclTHENS (cut dom) [tclTHENLIST - [Proofview.V82.of_tactic introf; - (fun gls-> - Proofview.V82.of_tactic (split (ImplicitBindings - [EConstr.mkVar (Tacmach.pf_nth_hyp_id gls 1)])) gls); + [introf; + Proofview.Goal.enter { enter = begin fun gl -> + let id0 = List.nth (pf_ids_of_hyps gl) 0 in + split (ImplicitBindings [mkVar id0]) + end }; tclSOLVE [wrap 0 true continue (deepen seq)]]; - tclTRY (Proofview.V82.of_tactic assumption)] + tclTRY assumption] | Real ((0,t),_) -> - (tclTHEN (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr t]))) + (tclTHEN (split (ImplicitBindings [t])) (tclSOLVE [wrap 0 true continue (deepen seq)])) | Real ((m,t),_) -> tclFAIL 0 (Pp.str "not implemented ... yet") + end } let instance_tac inst= if (snd inst)==dummy_id then @@ -195,10 +203,10 @@ let instance_tac inst= else left_instance_tac inst -let quantified_tac lf backtrack continue seq gl= - let insts=give_instances lf seq in +let quantified_tac lf backtrack continue seq = + Proofview.Goal.enter { enter = begin fun gl -> + let insts=give_instances (project gl) lf seq in tclORELSE (tclFIRST (List.map (fun inst->instance_tac inst continue seq) insts)) - backtrack gl - - + backtrack + end } diff --git a/plugins/firstorder/instances.mli b/plugins/firstorder/instances.mli index ce711f3f9..47550f314 100644 --- a/plugins/firstorder/instances.mli +++ b/plugins/firstorder/instances.mli @@ -9,9 +9,9 @@ open Globnames open Rules -val collect_quantified : Sequent.t -> Formula.t list * Sequent.t +val collect_quantified : Evd.evar_map -> Sequent.t -> Formula.t list * Sequent.t -val give_instances : Formula.t list -> Sequent.t -> +val give_instances : Evd.evar_map -> Formula.t list -> Sequent.t -> (Unify.instance * global_reference) list val quantified_tac : Formula.t list -> seqtac with_backtracking diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 96601f74a..e0d2c38a7 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -10,10 +10,11 @@ open CErrors open Util open Names open Term +open EConstr open Vars -open Tacmach +open Tacmach.New open Tactics -open Tacticals +open Tacticals.New open Proofview.Notations open Termops open Formula @@ -23,148 +24,165 @@ open Locus module NamedDecl = Context.Named.Declaration +type tactic = unit Proofview.tactic + type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic type lseqtac= global_reference -> seqtac type 'a with_backtracking = tactic -> 'a -let wrap n b continue seq gls= +let wrap n b continue seq = + Proofview.Goal.nf_enter { enter = begin fun gls -> Control.check_for_interrupt (); - let nc=pf_hyps gls in + let nc = Proofview.Goal.hyps gls in let env=pf_env gls in + let sigma = project gls in let rec aux i nc ctx= if i<=0 then seq else match nc with []->anomaly (Pp.str "Not the expected number of hyps") | nd::q-> let id = NamedDecl.get_id nd in - if occur_var env (project gls) id (pf_concl gls) || - List.exists (occur_var_in_decl env (project gls) id) ctx then + if occur_var env sigma id (pf_concl gls) || + List.exists (occur_var_in_decl env sigma id) ctx then (aux (i-1) q (nd::ctx)) else - add_formula Hyp (VarRef id) (EConstr.Unsafe.to_constr (NamedDecl.get_type nd)) (aux (i-1) q (nd::ctx)) gls in + add_formula env sigma Hyp (VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) in let seq1=aux n nc [] in let seq2=if b then - add_formula Concl dummy_id (EConstr.Unsafe.to_constr (pf_concl gls)) seq1 gls else seq1 in - continue seq2 gls + add_formula env sigma Concl dummy_id (pf_concl gls) seq1 else seq1 in + continue seq2 + end } let basename_of_global=function VarRef id->id | _->assert false let clear_global=function - VarRef id-> Proofview.V82.of_tactic (clear [id]) + VarRef id-> clear [id] | _->tclIDTAC (* connection rules *) -let axiom_tac t seq= - try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check c)) - with Not_found->tclFAIL 0 (Pp.str "No axiom link") +let axiom_tac t seq = + Proofview.Goal.enter { enter = begin fun gl -> + try + pf_constr_of_global (find_left (project gl) t seq) >>= fun c -> + exact_no_check c + with Not_found -> tclFAIL 0 (Pp.str "No axiom link") + end } -let ll_atom_tac a backtrack id continue seq= +let ll_atom_tac a backtrack id continue seq = let open EConstr in tclIFTHENELSE - (try - tclTHENLIST - [pf_constr_of_global (find_left a seq) (fun left -> - pf_constr_of_global id (fun id -> - Proofview.V82.of_tactic (generalize [(mkApp(id, [|left|]))]))); + (tclTHENLIST + [(Proofview.tclEVARMAP >>= fun sigma -> + let gr = + try Proofview.tclUNIT (find_left sigma a seq) + with Not_found -> tclFAIL 0 (Pp.str "No link") + in + gr >>= fun gr -> + pf_constr_of_global gr >>= fun left -> + pf_constr_of_global id >>= fun id -> + generalize [(mkApp(id, [|left|]))]); clear_global id; - Proofview.V82.of_tactic intro] - with Not_found->tclFAIL 0 (Pp.str "No link")) + intro]) (wrap 1 false continue seq) backtrack (* right connectives rules *) let and_tac backtrack continue seq= - tclIFTHENELSE (Proofview.V82.of_tactic simplest_split) (wrap 0 true continue seq) backtrack + tclIFTHENELSE simplest_split (wrap 0 true continue seq) backtrack let or_tac backtrack continue seq= tclORELSE - (Proofview.V82.of_tactic (any_constructor false (Some (Proofview.V82.tactic (tclCOMPLETE (wrap 0 true continue seq)))))) + (any_constructor false (Some (tclCOMPLETE (wrap 0 true continue seq)))) backtrack let arrow_tac backtrack continue seq= - tclIFTHENELSE (Proofview.V82.of_tactic intro) (wrap 1 true continue seq) + tclIFTHENELSE intro (wrap 1 true continue seq) (tclORELSE - (tclTHEN (Proofview.V82.of_tactic introf) (tclCOMPLETE (wrap 1 true continue seq))) + (tclTHEN introf (tclCOMPLETE (wrap 1 true continue seq))) backtrack) (* left connectives rules *) -let left_and_tac ind backtrack id continue seq gls= - let n=(construct_nhyps ind gls).(0) in +let left_and_tac ind backtrack id continue seq = + Proofview.Goal.enter { enter = begin fun gl -> + let n=(construct_nhyps (pf_env gl) ind).(0) in tclIFTHENELSE (tclTHENLIST - [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim); + [(pf_constr_of_global id >>= simplest_elim); clear_global id; - tclDO n (Proofview.V82.of_tactic intro)]) + tclDO n intro]) (wrap n false continue seq) - backtrack gls + backtrack + end } -let left_or_tac ind backtrack id continue seq gls= - let v=construct_nhyps ind gls in +let left_or_tac ind backtrack id continue seq = + Proofview.Goal.enter { enter = begin fun gl -> + let v=construct_nhyps (pf_env gl) ind in let f n= tclTHENLIST [clear_global id; - tclDO n (Proofview.V82.of_tactic intro); + tclDO n intro; wrap n false continue seq] in tclIFTHENSVELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim)) + (pf_constr_of_global id >>= simplest_elim) (Array.map f v) - backtrack gls + backtrack + end } let left_false_tac id= - Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim) + Tacticals.New.pf_constr_of_global id >>= simplest_elim (* left arrow connective rules *) (* We use this function for false, and, or, exists *) -let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl= - let rcs=ind_hyps 0 indu largs gl in +let ll_ind_tac (ind,u as indu) largs backtrack id continue seq = + Proofview.Goal.enter { enter = begin fun gl -> + let rcs=ind_hyps (pf_env gl) (project gl) 0 indu largs in let vargs=Array.of_list largs in (* construire le terme H->B, le generaliser etc *) let myterm idc i= let rc=rcs.(i) in let p=List.length rc in + let u = EInstance.make u in let cstr=mkApp ((mkConstructU ((ind,(i+1)),u)),vargs) in let vars=Array.init p (fun j->mkRel (p-j)) in let capply=mkApp ((lift p cstr),vars) in let head=mkApp ((lift p idc),[|capply|]) in - EConstr.of_constr (it_mkLambda_or_LetIn head rc) in + EConstr.it_mkLambda_or_LetIn head rc in let lp=Array.length rcs in - let newhyps idc =List.init lp (myterm (EConstr.Unsafe.to_constr idc)) in + let newhyps idc =List.init lp (myterm idc) in tclIFTHENELSE (tclTHENLIST - [pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize (newhyps idc))); + [(pf_constr_of_global id >>= fun idc -> generalize (newhyps idc)); clear_global id; - tclDO lp (Proofview.V82.of_tactic intro)]) - (wrap lp false continue seq) backtrack gl + tclDO lp intro]) + (wrap lp false continue seq) backtrack + end } let ll_arrow_tac a b c backtrack id continue seq= let open EConstr in let open Vars in - let a = EConstr.of_constr a in - let b = EConstr.of_constr b in - let c = EConstr.of_constr c in let cc=mkProd(Anonymous,a,(lift 1 b)) in let d idc = mkLambda (Anonymous,b, mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in tclORELSE - (tclTHENS (Proofview.V82.of_tactic (cut c)) + (tclTHENS (cut c) [tclTHENLIST - [Proofview.V82.of_tactic introf; + [introf; clear_global id; wrap 1 false continue seq]; - tclTHENS (Proofview.V82.of_tactic (cut cc)) - [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check c)); + tclTHENS (cut cc) + [(pf_constr_of_global id >>= fun c -> exact_no_check c); tclTHENLIST - [pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize [d idc])); + [(pf_constr_of_global id >>= fun idc -> generalize [d idc]); clear_global id; - Proofview.V82.of_tactic introf; - Proofview.V82.of_tactic introf; + introf; + introf; tclCOMPLETE (wrap 2 true continue seq)]]]) backtrack @@ -172,38 +190,40 @@ let ll_arrow_tac a b c backtrack id continue seq= let forall_tac backtrack continue seq= tclORELSE - (tclIFTHENELSE (Proofview.V82.of_tactic intro) (wrap 0 true continue seq) + (tclIFTHENELSE intro (wrap 0 true continue seq) (tclORELSE - (tclTHEN (Proofview.V82.of_tactic introf) (tclCOMPLETE (wrap 0 true continue seq))) + (tclTHEN introf (tclCOMPLETE (wrap 0 true continue seq))) backtrack)) (if !qflag then tclFAIL 0 (Pp.str "reversible in 1st order mode") else backtrack) -let left_exists_tac ind backtrack id continue seq gls= - let n=(construct_nhyps ind gls).(0) in +let left_exists_tac ind backtrack id continue seq = + Proofview.Goal.enter { enter = begin fun gl -> + let n=(construct_nhyps (pf_env gl) ind).(0) in tclIFTHENELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id >>= simplest_elim)) + (Tacticals.New.pf_constr_of_global id >>= simplest_elim) (tclTHENLIST [clear_global id; - tclDO n (Proofview.V82.of_tactic intro); + tclDO n intro; (wrap (n-1) false continue seq)]) backtrack - gls + end } let ll_forall_tac prod backtrack id continue seq= tclORELSE - (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr prod))) + (tclTHENS (cut prod) [tclTHENLIST - [Proofview.V82.of_tactic intro; - pf_constr_of_global id (fun idc -> - (fun gls-> + [intro; + (pf_constr_of_global id >>= fun idc -> + Proofview.Goal.enter { enter = begin fun gls-> let open EConstr in - let id0=pf_nth_hyp_id gls 1 in + let id0 = List.nth (pf_ids_of_hyps gls) 0 in let term=mkApp(idc,[|mkVar(id0)|]) in - tclTHEN (Proofview.V82.of_tactic (generalize [term])) (Proofview.V82.of_tactic (clear [id0])) gls)); + tclTHEN (generalize [term]) (clear [id0]) + end }); clear_global id; - Proofview.V82.of_tactic intro; + intro; tclCOMPLETE (wrap 1 false continue (deepen seq))]; tclCOMPLETE (wrap 0 true continue (deepen seq))]) backtrack @@ -215,12 +235,13 @@ let ll_forall_tac prod backtrack id continue seq= let constant str = Coqlib.gen_constant "User" ["Init";"Logic"] str let defined_connectives=lazy - [AllOccurrences,EvalConstRef (fst (destConst (constant "not"))); - AllOccurrences,EvalConstRef (fst (destConst (constant "iff")))] + [AllOccurrences,EvalConstRef (fst (Term.destConst (constant "not"))); + AllOccurrences,EvalConstRef (fst (Term.destConst (constant "iff")))] let normalize_evaluables= - onAllHypsAndConcl - (function - None-> Proofview.V82.of_tactic (unfold_in_concl (Lazy.force defined_connectives)) - | Some id -> - Proofview.V82.of_tactic (unfold_in_hyp (Lazy.force defined_connectives) (id,InHypTypeOnly))) + Proofview.Goal.enter { enter = begin fun gl -> + unfold_in_concl (Lazy.force defined_connectives) <*> + tclMAP + (fun id -> unfold_in_hyp (Lazy.force defined_connectives) (id,InHypTypeOnly)) + (pf_ids_of_hyps gl) + end } diff --git a/plugins/firstorder/rules.mli b/plugins/firstorder/rules.mli index 381b7cd87..80a7ae2c2 100644 --- a/plugins/firstorder/rules.mli +++ b/plugins/firstorder/rules.mli @@ -7,10 +7,13 @@ (************************************************************************) open Term +open EConstr open Tacmach open Names open Globnames +type tactic = unit Proofview.tactic + type seqtac= (Sequent.t -> tactic) -> Sequent.t -> tactic type lseqtac= global_reference -> seqtac diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index fb0c22c2b..59b842c82 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open CErrors open Util open Formula @@ -57,11 +58,11 @@ end module OrderedConstr= struct - type t=constr + type t=Constr.t let compare=constr_ord end -type h_item = global_reference * (int*constr) option +type h_item = global_reference * (int*Constr.t) option module Hitem= struct @@ -81,13 +82,15 @@ module CM=Map.Make(OrderedConstr) module History=Set.Make(Hitem) -let cm_add typ nam cm= +let cm_add sigma typ nam cm= + let typ = EConstr.to_constr sigma typ in try let l=CM.find typ cm in CM.add typ (nam::l) cm with Not_found->CM.add typ [nam] cm -let cm_remove typ nam cm= +let cm_remove sigma typ nam cm= + let typ = EConstr.to_constr sigma typ in try let l=CM.find typ cm in let l0=List.filter (fun id-> not (Globnames.eq_gr id nam)) l in @@ -112,19 +115,19 @@ let deepen seq={seq with depth=seq.depth-1} let record item seq={seq with history=History.add item seq.history} -let lookup item seq= +let lookup sigma item seq= History.mem item seq.history || match item with (_,None)->false - | (id,Some ((m,t) as c))-> + | (id,Some (m, t))-> let p (id2,o)= match o with None -> false - | Some ((m2,t2) as c2)-> Globnames.eq_gr id id2 && m2>m && more_general c2 c in + | Some (m2, t2)-> Globnames.eq_gr id id2 && m2>m && more_general sigma (m2, EConstr.of_constr t2) (m, EConstr.of_constr t) in History.exists p seq.history -let add_formula side nam t seq gl= - match build_formula side nam t gl seq.cnt with +let add_formula env sigma side nam t seq = + match build_formula env sigma side nam t seq.cnt with Left f-> begin match side with @@ -136,7 +139,7 @@ let add_formula side nam t seq gl= | _ -> {seq with redexes=HP.add f seq.redexes; - context=cm_add f.constr nam seq.context} + context=cm_add sigma f.constr nam seq.context} end | Right t-> match side with @@ -144,18 +147,18 @@ let add_formula side nam t seq gl= {seq with gl=t;glatom=Some t} | _ -> {seq with - context=cm_add t nam seq.context; + context=cm_add sigma t nam seq.context; latoms=t::seq.latoms} -let re_add_formula_list lf seq= +let re_add_formula_list sigma lf seq= let do_one f cm= if f.id == dummy_id then cm - else cm_add f.constr f.id cm in + else cm_add sigma f.constr f.id cm in {seq with redexes=List.fold_right HP.add lf seq.redexes; context=List.fold_right do_one lf seq.context} -let find_left t seq=List.hd (CM.find t seq.context) +let find_left sigma t seq=List.hd (CM.find (EConstr.to_constr sigma t) seq.context) (*let rev_left seq= try @@ -164,7 +167,7 @@ let find_left t seq=List.hd (CM.find t seq.context) with Heap.EmptyHeap -> false *) -let rec take_formula seq= +let rec take_formula sigma seq= let hd=HP.maximum seq.redexes and hp=HP.remove seq.redexes in if hd.id == dummy_id then @@ -172,11 +175,11 @@ let rec take_formula seq= if seq.gl==hd.constr then hd,nseq else - take_formula nseq (* discarding deprecated goal *) + take_formula sigma nseq (* discarding deprecated goal *) else hd,{seq with redexes=hp; - context=cm_remove hd.constr hd.id seq.context} + context=cm_remove sigma hd.constr hd.id seq.context} let empty_seq depth= {redexes=HP.empty; @@ -196,18 +199,17 @@ let expand_constructor_hints = | gr -> [gr]) -let extend_with_ref_list l seq gl = +let extend_with_ref_list env sigma l seq = let l = expand_constructor_hints l in - let f gr (seq,gl) = - let gl, c = pf_eapply Evd.fresh_global gl gr in - let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in - let typ = EConstr.Unsafe.to_constr typ in - (add_formula Hyp gr typ seq gl,gl) in - List.fold_right f l (seq,gl) + let f gr (seq, sigma) = + let sigma, c = Evd.fresh_global env sigma gr in + let sigma, typ= Typing.type_of env sigma (EConstr.of_constr c) in + (add_formula env sigma Hyp gr typ seq, sigma) in + List.fold_right f l (seq, sigma) open Hints -let extend_with_auto_hints l seq gl= +let extend_with_auto_hints env sigma l seq = let seqref=ref seq in let f p_a_t = match repr_hint p_a_t.code with @@ -215,10 +217,9 @@ let extend_with_auto_hints l seq gl= | Res_pf_THEN_trivial_fail (c,_) -> let (c, _, _) = c in (try - let (gr, _) = Termops.global_of_constr (project gl) c in - let typ=(pf_unsafe_type_of gl c) in - let typ = EConstr.Unsafe.to_constr typ in - seqref:=add_formula Hint gr typ !seqref gl + let (gr, _) = Termops.global_of_constr sigma c in + let typ=(Typing.unsafe_type_of env sigma c) in + seqref:=add_formula env sigma Hint gr typ !seqref with Not_found->()) | _-> () in let g _ _ l = List.iter f l in @@ -230,7 +231,7 @@ let extend_with_auto_hints l seq gl= error ("Firstorder: "^dbname^" : No such Hint database") in Hint_db.iter g hdb in List.iter h l; - !seqref, gl (*FIXME: forgetting about universes*) + !seqref, sigma (*FIXME: forgetting about universes*) let print_cmap map= let print_entry c l s= diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli index 06c9251e7..18d68f54f 100644 --- a/plugins/firstorder/sequent.mli +++ b/plugins/firstorder/sequent.mli @@ -7,22 +7,23 @@ (************************************************************************) open Term +open EConstr open Formula open Tacmach open Globnames -module OrderedConstr: Set.OrderedType with type t=constr +module OrderedConstr: Set.OrderedType with type t=Constr.t -module CM: CSig.MapS with type key=constr +module CM: CSig.MapS with type key=Constr.t -type h_item = global_reference * (int*constr) option +type h_item = global_reference * (int*Constr.t) option module History: Set.S with type elt = h_item -val cm_add : constr -> global_reference -> global_reference list CM.t -> +val cm_add : Evd.evar_map -> constr -> global_reference -> global_reference list CM.t -> global_reference list CM.t -val cm_remove : constr -> global_reference -> global_reference list CM.t -> +val cm_remove : Evd.evar_map -> constr -> global_reference -> global_reference list CM.t -> global_reference list CM.t module HP: Heap.S with type elt=Formula.t @@ -40,23 +41,22 @@ val deepen: t -> t val record: h_item -> t -> t -val lookup: h_item -> t -> bool +val lookup: Evd.evar_map -> h_item -> t -> bool -val add_formula : side -> global_reference -> constr -> t -> - Proof_type.goal sigma -> t +val add_formula : Environ.env -> Evd.evar_map -> side -> global_reference -> constr -> t -> t -val re_add_formula_list : Formula.t list -> t -> t +val re_add_formula_list : Evd.evar_map -> Formula.t list -> t -> t -val find_left : constr -> t -> global_reference +val find_left : Evd.evar_map -> constr -> t -> global_reference -val take_formula : t -> Formula.t * t +val take_formula : Evd.evar_map -> t -> Formula.t * t val empty_seq : int -> t -val extend_with_ref_list : global_reference list -> - t -> Proof_type.goal sigma -> t * Proof_type.goal sigma +val extend_with_ref_list : Environ.env -> Evd.evar_map -> global_reference list -> + t -> t * Evd.evar_map -val extend_with_auto_hints : Hints.hint_db_name list -> - t -> Proof_type.goal sigma -> t * Proof_type.goal sigma +val extend_with_auto_hints : Environ.env -> Evd.evar_map -> Hints.hint_db_name list -> + t -> t * Evd.evar_map val print_cmap: global_reference list CM.t -> Pp.std_ppcmds diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 7cbfb8e7d..49bf07155 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -8,6 +8,7 @@ open Util open Term +open EConstr open Vars open Termops open Reductionops @@ -21,13 +22,12 @@ exception UFAIL of constr*constr to the equation set. Raises UFAIL with a pair of terms *) -let strip_outer_cast t = - EConstr.Unsafe.to_constr (strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *) - let pop t = Vars.lift (-1) t +let subst_meta subst t = + let subst = List.map (fun (m, c) -> (m, EConstr.Unsafe.to_constr c)) subst in + EConstr.of_constr (subst_meta subst (EConstr.Unsafe.to_constr t)) -let unif t1 t2= - let evd = Evd.empty in (** FIXME *) +let unif evd t1 t2= let bige=Queue.create () and sigma=ref [] in let bind i t= @@ -35,7 +35,7 @@ let unif t1 t2= (List.map (function (n,tn)->(n,subst_meta [i,t] tn)) !sigma) in let rec head_reduce t= (* forbids non-sigma-normal meta in head position*) - match kind_of_term t with + match EConstr.kind evd t with Meta i-> (try head_reduce (Int.List.assoc i !sigma) @@ -44,25 +44,25 @@ let unif t1 t2= Queue.add (t1,t2) bige; try while true do let t1,t2=Queue.take bige in - let nt1=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t1))) - and nt2=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t2))) in - match (kind_of_term nt1),(kind_of_term nt2) with + let nt1=head_reduce (whd_betaiotazeta evd t1) + and nt2=head_reduce (whd_betaiotazeta evd t2) in + match (EConstr.kind evd nt1),(EConstr.kind evd nt2) with Meta i,Meta j-> if not (Int.equal i j) then if i let t=subst_meta !sigma nt2 in - if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) && - not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then + if Int.Set.is_empty (free_rels evd t) && + not (occur_term evd (EConstr.mkMeta i) t) then bind i t else raise (UFAIL(nt1,nt2)) | _,Meta i -> let t=subst_meta !sigma nt1 in - if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) && - not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then + if Int.Set.is_empty (free_rels evd t) && + not (occur_term evd (EConstr.mkMeta i) t) then bind i t else raise (UFAIL(nt1,nt2)) - | Cast(_,_,_),_->Queue.add (strip_outer_cast nt1,nt2) bige - | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast nt2) bige + | Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige + | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))-> Queue.add (a,c) bige;Queue.add (pop b,pop d) bige | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> @@ -84,19 +84,19 @@ let unif t1 t2= for i=0 to l-1 do Queue.add (va.(i),vb.(i)) bige done - | _->if not (eq_constr_nounivs nt1 nt2) then raise (UFAIL (nt1,nt2)) + | _->if not (eq_constr_nounivs evd nt1 nt2) then raise (UFAIL (nt1,nt2)) done; assert false (* this place is unreachable but needed for the sake of typing *) with Queue.Empty-> !sigma -let value i t= +let value evd i t= let add x y= if x<0 then y else if y<0 then x else x+y in let rec vaux term= - if isMeta term && Int.equal (destMeta term) i then 0 else + if isMeta evd term && Int.equal (destMeta evd term) i then 0 else let f v t=add v (vaux t) in - let vr=fold_constr f (-1) term in + let vr=EConstr.fold evd f (-1) term in if vr<0 then -1 else vr+1 in vaux t @@ -104,11 +104,11 @@ type instance= Real of (int*constr)*int | Phantom of constr -let mk_rel_inst t= +let mk_rel_inst evd t= let new_rel=ref 1 in let rel_env=ref [] in let rec renum_rec d t= - match kind_of_term t with + match EConstr.kind evd t with Meta n-> (try mkRel (d+(Int.List.assoc n !rel_env)) @@ -117,15 +117,15 @@ let mk_rel_inst t= incr new_rel; rel_env:=(n,m) :: !rel_env; mkRel (m+d)) - | _ -> map_constr_with_binders succ renum_rec d t + | _ -> EConstr.map_with_binders evd succ renum_rec d t in let nt=renum_rec 0 t in (!new_rel - 1,nt) -let unif_atoms i dom t1 t2= +let unif_atoms evd i dom t1 t2= try - let t=Int.List.assoc i (unif t1 t2) in - if isMeta t then Some (Phantom dom) - else Some (Real(mk_rel_inst t,value i t1)) + let t=Int.List.assoc i (unif evd t1 t2) in + if isMeta evd t then Some (Phantom dom) + else Some (Real(mk_rel_inst evd t,value evd i t1)) with UFAIL(_,_) ->None | Not_found ->Some (Phantom dom) @@ -134,11 +134,11 @@ let renum_metas_from k n t= (* requires n = max (free_rels t) *) let l=List.init n (fun i->mkMeta (k+i)) in substl l t -let more_general (m1,t1) (m2,t2)= +let more_general evd (m1,t1) (m2,t2)= let mt1=renum_metas_from 0 m1 t1 and mt2=renum_metas_from m1 m2 t2 in try - let sigma=unif mt1 mt2 in - let p (n,t)= nfalse diff --git a/plugins/firstorder/unify.mli b/plugins/firstorder/unify.mli index 4fe9ad38d..c9cca9bd8 100644 --- a/plugins/firstorder/unify.mli +++ b/plugins/firstorder/unify.mli @@ -7,15 +7,16 @@ (************************************************************************) open Term +open EConstr exception UFAIL of constr*constr -val unif : constr -> constr -> (int*constr) list +val unif : Evd.evar_map -> constr -> constr -> (int*constr) list type instance= Real of (int*constr)*int (* nb trous*terme*valeur heuristique *) | Phantom of constr (* domaine de quantification *) -val unif_atoms : metavariable -> constr -> constr -> constr -> instance option +val unif_atoms : Evd.evar_map -> metavariable -> constr -> constr -> constr -> instance option -val more_general : (int*constr) -> (int*constr) -> bool +val more_general : Evd.evar_map -> (int*constr) -> (int*constr) -> bool -- cgit v1.2.3 From 98da9fdce866728f93bc7cb690275f5559aa9bae Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 16 Apr 2017 20:37:15 +0200 Subject: Removing various tactic compatibility layers in core tactics. --- plugins/ltac/g_class.ml4 | 2 +- plugins/ltac/g_rewrite.ml4 | 13 +++++++++---- plugins/ltac/rewrite.ml | 6 +++--- tactics/auto.ml | 11 ++++++----- tactics/class_tactics.ml | 28 +++++++++++++++++----------- tactics/class_tactics.mli | 2 +- 6 files changed, 37 insertions(+), 25 deletions(-) diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4 index 40f30c794..ff5e7d5ff 100644 --- a/plugins/ltac/g_class.ml4 +++ b/plugins/ltac/g_class.ml4 @@ -85,7 +85,7 @@ TACTIC EXTEND not_evar END TACTIC EXTEND is_ground - [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground ty) ] + [ "is_ground" constr(ty) ] -> [ is_ground ty ] END TACTIC EXTEND autoapply diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 index c50100bf5..fdcaedab3 100644 --- a/plugins/ltac/g_rewrite.ml4 +++ b/plugins/ltac/g_rewrite.ml4 @@ -19,6 +19,7 @@ open Geninterp open Extraargs open Tacmach open Tacticals +open Proofview.Notations open Rewrite open Stdarg open Pcoq.Vernac_ @@ -123,15 +124,19 @@ TACTIC EXTEND rewrite_strat END let clsubstitute o c = + Proofview.Goal.enter { enter = begin fun gl -> let is_tac id = match fst (fst (snd c)) with GVar (_, id') when Id.equal id' id -> true | _ -> false in - Tacticals.onAllHypsAndConcl + let hyps = Tacmach.New.pf_ids_of_hyps gl in + Tacticals.New.tclMAP (fun cl -> match cl with - | Some id when is_tac id -> tclIDTAC - | _ -> Proofview.V82.of_tactic (cl_rewrite_clause c o AllOccurrences cl)) + | Some id when is_tac id -> Tacticals.New.tclIDTAC + | _ -> cl_rewrite_clause c o AllOccurrences cl) + (None :: List.map (fun id -> Some id) hyps) + end } TACTIC EXTEND substitute -| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> [ Proofview.V82.tactic (clsubstitute o c) ] +| [ "substitute" orient(o) glob_constr_with_bindings(c) ] -> [ clsubstitute o c ] END diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index b84be4600..12a1566e2 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -2197,7 +2197,8 @@ let setoid_transitivity c = (transitivity_red true c) let setoid_symmetry_in id = - Proofview.V82.tactic (fun gl -> + let open Tacmach.New in + Proofview.Goal.enter { enter = begin fun gl -> let sigma = project gl in let ctype = pf_unsafe_type_of gl (mkVar id) in let binders,concl = decompose_prod_assum sigma ctype in @@ -2211,11 +2212,10 @@ let setoid_symmetry_in id = let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in let new_hyp' = mkApp (he, [| c2 ; c1 |]) in let new_hyp = it_mkProd_or_LetIn new_hyp' binders in - Proofview.V82.of_tactic (tclTHENLAST (Tactics.assert_after_replacing id new_hyp) (tclTHENLIST [ intros; setoid_symmetry; apply (mkVar id); Tactics.assumption ])) - gl) + end } let _ = Hook.set Tactics.setoid_reflexivity setoid_reflexivity let _ = Hook.set Tactics.setoid_symmetry setoid_symmetry diff --git a/tactics/auto.ml b/tactics/auto.ml index 74cb7a364..42230dff1 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -380,7 +380,7 @@ and my_find_search_delta sigma db_list local_db secvars hdc concl = and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t;poly=poly;db=dbname})) = let tactic = function | Res_pf (c,cl) -> unify_resolve_gen poly flags (c,cl) - | ERes_pf _ -> Proofview.V82.tactic (fun gl -> error "eres_pf") + | ERes_pf _ -> Proofview.Goal.enter { enter = fun gl -> Tacticals.New.tclZEROMSG (str "eres_pf") } | Give_exact (c, cl) -> exact poly (c, cl) | Res_pf_THEN_trivial_fail (c,cl) -> Tacticals.New.tclTHEN @@ -389,10 +389,11 @@ and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t;poly=poly;db= with "debug auto" we don't display the details of inner trivial *) (trivial_fail_db (no_dbg ()) (not (Option.is_empty flags)) db_list local_db) | Unfold_nth c -> - Proofview.V82.tactic (fun gl -> - if exists_evaluable_reference (pf_env gl) c then - tclPROGRESS (Proofview.V82.of_tactic (reduce (Unfold [AllOccurrences,c]) Locusops.onConcl)) gl - else tclFAIL 0 (str"Unbound reference") gl) + Proofview.Goal.enter { enter = begin fun gl -> + if exists_evaluable_reference (Tacmach.New.pf_env gl) c then + Tacticals.New.tclPROGRESS (reduce (Unfold [AllOccurrences,c]) Locusops.onConcl) + else Tacticals.New.tclFAIL 0 (str"Unbound reference") + end } | Extern tacast -> conclPattern concl p tacast in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index d70275dee..54e4405d1 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -221,18 +221,22 @@ let auto_unif_flags freeze st = resolve_evars = false } -let e_give_exact flags poly (c,clenv) gl = +let e_give_exact flags poly (c,clenv) = + let open Tacmach.New in + Proofview.Goal.s_enter { s_enter = begin fun gl -> + let sigma = project gl in let (c, _, _) = c in - let c, gl = + let c, sigma = if poly then let clenv', subst = Clenv.refresh_undefined_univs clenv in - let evd = evars_reset_evd ~with_conv_pbs:true gl.sigma clenv'.evd in + let evd = evars_reset_evd ~with_conv_pbs:true sigma clenv'.evd in let c = Vars.subst_univs_level_constr subst c in - c, {gl with sigma = evd} - else c, gl + c, evd + else c, sigma in - let t1 = pf_unsafe_type_of gl c in - Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl + let (sigma, t1) = Typing.type_of (pf_env gl) sigma c in + Sigma.Unsafe.of_pair (Clenvtac.unify ~flags t1 <*> exact_no_check c, sigma) + end } let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> let clenv', c = connect_hint_clenv poly c clenv gls in @@ -455,7 +459,7 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co { enter = fun gl -> unify_resolve_refine poly flags gl (c,None,clenv) } in Tacticals.New.tclTHEN tac Proofview.shelve_unifiable else - Proofview.V82.tactic (e_give_exact flags poly (c,clenv)) + e_give_exact flags poly (c,clenv) | Res_pf_THEN_trivial_fail (term,cl) -> let fst = with_prods nprods poly (term,cl) (unify_e_resolve poly flags) in let snd = if complete then Tacticals.New.tclIDTAC @@ -1613,9 +1617,11 @@ let not_evar c = | Evar _ -> Tacticals.New.tclFAIL 0 (str"Evar") | _ -> Proofview.tclUNIT () -let is_ground c gl = - if Evarutil.is_ground_term (project gl) c then tclIDTAC gl - else tclFAIL 0 (str"Not ground") gl +let is_ground c = + let open Tacticals.New in + Proofview.tclEVARMAP >>= fun sigma -> + if Evarutil.is_ground_term sigma c then tclIDTAC + else tclFAIL 0 (str"Not ground") let autoapply c i = let open Proofview.Notations in diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index a38be5972..738cc0feb 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -33,7 +33,7 @@ val head_of_constr : Id.t -> constr -> unit Proofview.tactic val not_evar : constr -> unit Proofview.tactic -val is_ground : constr -> tactic +val is_ground : constr -> unit Proofview.tactic val autoapply : constr -> Hints.hint_db_name -> unit Proofview.tactic -- cgit v1.2.3 From b4ca718e7cb62933ea39b61ed3737f3479517bdf Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 17 Apr 2017 15:07:37 +0200 Subject: Adding a dedicated travis overlay for fiat-parsers. --- dev/ci/ci-user-overlay.sh | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/dev/ci/ci-user-overlay.sh b/dev/ci/ci-user-overlay.sh index bb193ebb5..fad647291 100644 --- a/dev/ci/ci-user-overlay.sh +++ b/dev/ci/ci-user-overlay.sh @@ -25,7 +25,7 @@ echo $TRAVIS_PULL_REQUEST echo $TRAVIS_BRANCH echo $TRAVIS_COMMIT -if [ $TRAVIS_PULL_REQUEST == "461" ] || [ $TRAVIS_BRANCH == "stm+remove_compat_parsing" ]; then - mathcomp_CI_BRANCH=no_camlp4_compat - mathcomp_CI_GITURL=https://github.com/ejgallego/math-comp.git +if [ $TRAVIS_PULL_REQUEST == "568" ] || [ $TRAVIS_BRANCH == "remove-tactic-compat" ]; then + fiat_parsers_CI_BRANCH=fix-ml + fiat_parsers_CI_GITURL=https://github.com/ppedrot/fiat.git fi -- cgit v1.2.3 From da0d6da035206778c1d99ef51f471b4b22016492 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 24 Apr 2017 18:42:03 +0200 Subject: [toplevel] Don't mask critical exceptions in vernac interpretation. Indeed we were not correctly backtracking in the case of StackOverflow. It makes sense to remove the inner handler which is a leftover of a previous attempt, and handle interpretation errors in load as non-recoverable. This fixes: https://coq.inria.fr/bugs/show_bug.cgi?id=5485 --- toplevel/vernac.ml | 40 ++++++++++++++++++---------------------- 1 file changed, 18 insertions(+), 22 deletions(-) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 3359a1672..19753b6f6 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -115,28 +115,23 @@ let rec interp_vernac sid po (loc,com) = let f = Loadpath.locate_file fname in load_vernac verbosely sid f | v -> - try - let nsid, ntip = Stm.add sid (not !Flags.quiet) (loc,v) in - - (* Main STM interaction *) - if ntip <> `NewTip then - anomaly (str "vernac.ml: We got an unfocus operation on the toplevel!"); - (* Due to bug #5363 we cannot use observe here as we should, - it otherwise reveals bugs *) - (* Stm.observe nsid; *) - Stm.finish (); - - (* We could use a more refined criteria that depends on the - vernac. For now we imitate the old approach. *) - let hide_goals = !Flags.batch_mode || is_query v || !Flags.quiet || - not (Proof_global.there_are_pending_proofs ()) in - - if not hide_goals then Feedback.msg_notice (pr_open_cur_subgoals ()); - nsid - - with exn when CErrors.noncritical exn -> - ignore(Stm.edit_at sid); - raise exn + let nsid, ntip = Stm.add sid (not !Flags.quiet) (loc,v) in + + (* Main STM interaction *) + if ntip <> `NewTip then + anomaly (str "vernac.ml: We got an unfocus operation on the toplevel!"); + (* Due to bug #5363 we cannot use observe here as we should, + it otherwise reveals bugs *) + (* Stm.observe nsid; *) + Stm.finish (); + + (* We could use a more refined criteria that depends on the + vernac. For now we imitate the old approach. *) + let hide_goals = !Flags.batch_mode || is_query v || !Flags.quiet || + not (Proof_global.there_are_pending_proofs ()) in + + if not hide_goals then Feedback.msg_notice (pr_open_cur_subgoals ()); + nsid in try (* The -time option is only supported from console-based @@ -145,6 +140,7 @@ let rec interp_vernac sid po (loc,com) = let com = if !Flags.time then VernacTime (loc,com) else com in interp com with reraise -> + ignore(Stm.edit_at sid); let (reraise, info) = CErrors.push reraise in let loc' = Option.default Loc.ghost (Loc.get_loc info) in if Loc.is_ghost loc' then iraise (reraise, Loc.add_loc info loc) -- cgit v1.2.3 From 0fd7d060e0872a6ae3662dfb7a1fb940b80ef9df Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 24 Apr 2017 19:10:19 +0200 Subject: [toplevel] Don't check proofs in -quick mode. This fixes a logical error introduced in ce2b2058587224ade9261cd4127ef4f6e94d356b Patch by @gares, closes https://coq.inria.fr/bugs/show_bug.cgi?id=5484 --- toplevel/vernac.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 19753b6f6..f81f77e6e 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -123,7 +123,9 @@ let rec interp_vernac sid po (loc,com) = (* Due to bug #5363 we cannot use observe here as we should, it otherwise reveals bugs *) (* Stm.observe nsid; *) - Stm.finish (); + + let check_proof = Flags.(!compilation_mode = BuildVo || not !batch_mode) in + if check_proof then Stm.finish (); (* We could use a more refined criteria that depends on the vernac. For now we imitate the old approach. *) -- cgit v1.2.3 From adc2035410a339cfa88dae527b631f5131adaa54 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 7 Mar 2017 18:13:07 +0100 Subject: Fix an optimization failure in tclPROGRESS. Due to code reworking, a fastpath got anihilated because the slow path was computed altogether. We now only compute the slow check whenever the quick one fails. --- engine/proofview.ml | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/engine/proofview.ml b/engine/proofview.ml index f054038e9..99bd4bc4f 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -858,14 +858,12 @@ let tclPROGRESS t = let quick_test = initial.solution == final.solution && initial.comb == final.comb in - let exhaustive_test = + let test = + quick_test || Util.List.for_all2eq begin fun i f -> Progress.goal_equal initial.solution i final.solution f end initial.comb final.comb in - let test = - quick_test || exhaustive_test - in if not test then tclUNIT res else -- cgit v1.2.3 From cb635eab423cde23757725593ffdfb98f4016881 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 7 Mar 2017 18:13:07 +0100 Subject: Fix an optimization failure in tclPROGRESS. Due to code reworking, a fastpath got anihilated because the slow path was computed altogether. We now only compute the slow check whenever the quick one fails. --- engine/proofview.ml | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/engine/proofview.ml b/engine/proofview.ml index c01879765..2e6266cf1 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -828,14 +828,12 @@ let tclPROGRESS t = let quick_test = initial.solution == final.solution && initial.comb == final.comb in - let exhaustive_test = + let test = + quick_test || Util.List.for_all2eq begin fun i f -> Progress.goal_equal initial.solution i final.solution f end initial.comb final.comb in - let test = - quick_test || exhaustive_test - in if not test then tclUNIT res else -- cgit v1.2.3 From 4abb41d008fc754f21916dcac9cce49f2d04dd6d Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 24 Apr 2017 19:25:48 +0200 Subject: [toplevel] Use exception information for error printing. This is a partial backtrack on 63cfc77ddf3586262d905dc351b58669d185a55e. In that commit, we disregarded exception and tried to print error messages just by listening to feedback. However, feedback error messages are not always emitted due to https://coq.inria.fr/bugs/show_bug.cgi?id=5479 Thus meanwhile it is safer to go back to printing the information present in exceptions until we tweak the STM. This fixes https://coq.inria.fr/bugs/show_bug.cgi?id=5467 and many other glitches not reported, such errors in nested proofs. --- toplevel/coqloop.ml | 33 +++++++++++++-------------------- toplevel/coqtop.ml | 25 ++++++++++--------------- toplevel/vernac.ml | 50 +++++++++++++++++++++++++++++++++----------------- vernac/topfmt.ml | 23 +++++++++++++++++++++++ vernac/topfmt.mli | 11 +++++++---- 5 files changed, 86 insertions(+), 56 deletions(-) diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml index 4641a2bc8..b608488c8 100644 --- a/toplevel/coqloop.ml +++ b/toplevel/coqloop.ml @@ -149,20 +149,6 @@ let valid_buffer_loc ib loc = not (Loc.is_ghost loc) && let (b,e) = Loc.unloc loc in b-ib.start >= 0 && e-ib.start < ib.len && b<=e -(* This is specific to the toplevel *) -let pr_loc loc = - if Loc.is_ghost loc then str"" - else - let fname = loc.Loc.fname in - if CString.equal fname "" then - Loc.(str"Toplevel input, characters " ++ int loc.bp ++ - str"-" ++ int loc.ep ++ str":") - else - Loc.(str"File " ++ str "\"" ++ str fname ++ str "\"" ++ - str", line " ++ int loc.line_nb ++ str", characters " ++ - int (loc.bp-loc.bol_pos) ++ str"-" ++ int (loc.ep-loc.bol_pos) ++ - str":") - (* Toplevel error explanation. *) let error_info_for_buffer ?loc buf = Option.map (fun loc -> @@ -177,7 +163,7 @@ let error_info_for_buffer ?loc buf = else (mt (), nloc) (* we are in batch mode, don't adjust location *) else (mt (), loc) - in pr_loc loc ++ hl + in Topfmt.pr_loc loc ++ hl ) loc (* Actual printing routine *) @@ -292,6 +278,9 @@ let coqloop_feed (fb : Feedback.feedback) = let open Feedback in | FileDependency (_,_) -> () | FileLoaded (_,_) -> () | Custom (_,_,_) -> () + (* Re-enable when we switch back to feedback-based error printing *) + | Message (Error,loc,msg) -> () + (* TopErr.print_error_for_buffer ?loc lvl msg top_buffer *) | Message (lvl,loc,msg) -> TopErr.print_error_for_buffer ?loc lvl msg top_buffer @@ -318,11 +307,15 @@ let do_vernac sid = | CErrors.Drop -> (* Last chance *) if Mltop.is_ocaml_top() then raise CErrors.Drop else (Feedback.msg_error (str "There is no ML toplevel."); sid) - (* Exception printing is done now by the feedback listener. *) - (* XXX: We need this hack due to the side effects of the exception - printer and the reliance of Stm.define on attaching crutial - state to exceptions *) - | any -> ignore (CErrors.(iprint (push any))); sid + (* Exception printing should be done by the feedback listener, + however this is not yet ready so we rely on the exception for + now. *) + | any -> + let (e, info) = CErrors.push any in + let loc = Loc.get_loc info in + let msg = CErrors.iprint (e, info) in + TopErr.print_error_for_buffer ?loc Feedback.Error msg top_buffer; + sid (** Main coq loop : read vernacular expressions until Drop is entered. Ctrl-C is handled internally as Sys.Break instead of aborting Coq. diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index f5f43ff66..9cf075065 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -430,10 +430,10 @@ let get_native_name s = (** Prints info which is either an error or an anomaly and then exits with the appropriate error code *) -let fatal_error info anomaly = - let msg = info ++ fnl () in - Format.fprintf !Topfmt.err_ft "@[%a@]%!" pp_with msg; - exit (if anomaly then 129 else 1) +let fatal_error ?extra exn = + Topfmt.print_err_exn ?extra exn; + let exit_code = if CErrors.(is_anomaly exn || not (handled exn)) then 129 else 1 in + exit exit_code let parse_args arglist = let args = ref arglist in @@ -596,11 +596,7 @@ let parse_args arglist = in try parse () - with - | UserError(_, s) as e -> - if ismt s then exit 1 - else fatal_error (CErrors.print e) false - | any -> fatal_error (CErrors.print any) (CErrors.is_anomaly any) + with any -> fatal_error any let init_toplevel arglist = init_gc (); @@ -646,14 +642,13 @@ let init_toplevel arglist = check_vio_tasks (); outputstate () with any -> - let any = CErrors.push any in flush_all(); - let msg = - if !batch_mode && not Stateid.(equal (Stm.get_current_state ()) dummy) then mt () - else str "Error during initialization: " ++ CErrors.iprint any ++ fnl () + let extra = + if !batch_mode && not Stateid.(equal (Stm.get_current_state ()) dummy) + then None + else Some (str "Error during initialization: ") in - let is_anomaly e = CErrors.is_anomaly e || not (CErrors.handled e) in - fatal_error msg (is_anomaly (fst any)) + fatal_error ?extra any end; if !batch_mode then begin flush_all(); diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index f81f77e6e..8bcf2114b 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -107,6 +107,16 @@ let pr_open_cur_subgoals () = try Printer.pr_open_subgoals () with Proof_global.NoCurrentProof -> Pp.str "" +let vernac_error msg = + Format.fprintf !Topfmt.err_ft "@[%a@]%!" Pp.pp_with msg; + flush_all (); + exit 1 + +(* Reenable when we get back to feedback printing *) +(* let is_end_of_input any = match any with *) +(* Stm.End_of_input -> true *) +(* | _ -> false *) + let rec interp_vernac sid po (loc,com) = let interp = function | VernacLoad (verbosely, fname) -> @@ -161,17 +171,25 @@ and load_vernac verbosely sid file = * raised, which means that we raised the end of the file being loaded *) while true do let loc, ast = + Stm.parse_sentence !rsid in_pa + (* If an error in parsing occurs, we propagate the exception + so the caller of load_vernac will take care of it. However, + in the future it could be possible that we want to handle + all the errors as feedback events, thus in this case we + should relay the exception here for convenience. A + possibility is shown below, however we may want to refactor + this code: + try Stm.parse_sentence !rsid in_pa with - | Stm.End_of_input -> raise Stm.End_of_input - | any -> + | any when not is_end_of_input any -> let (e, info) = CErrors.push any in let loc = Loc.get_loc info in let msg = CErrors.iprint (e, info) in Feedback.msg_error ?loc msg; iraise (e, info) + *) in - (* Printing of vernacs *) if !beautify then pr_new_syntax in_pa chan_beautify loc (Some ast); Option.iter (vernac_echo loc) in_echo; @@ -231,13 +249,10 @@ let chop_extension f = let ensure_bname src tgt = let src, tgt = Filename.basename src, Filename.basename tgt in let src, tgt = chop_extension src, chop_extension tgt in - if src <> tgt then begin - Feedback.msg_error (str "Source and target file names must coincide, directories can differ" ++ fnl () ++ - str "Source: " ++ str src ++ fnl () ++ - str "Target: " ++ str tgt); - flush_all (); - exit 1 - end + if src <> tgt then + vernac_error (str "Source and target file names must coincide, directories can differ" ++ fnl () ++ + str "Source: " ++ str src ++ fnl () ++ + str "Target: " ++ str tgt) let ensure ext src tgt = ensure_bname src tgt; ensure_ext ext tgt @@ -246,17 +261,15 @@ let ensure_vo v vo = ensure ".vo" v vo let ensure_vio v vio = ensure ".vio" v vio let ensure_exists f = - if not (Sys.file_exists f) then begin - Feedback.msg_error (hov 0 (str "Can't find file" ++ spc () ++ str f)); - exit 1 - end + if not (Sys.file_exists f) then + vernac_error (hov 0 (str "Can't find file" ++ spc () ++ str f)) (* Compile a vernac file *) let compile verbosely f = let check_pending_proofs () = let pfs = Pfedit.get_all_proof_names () in - if not (List.is_empty pfs) then - (Feedback.msg_error (str "There are pending proofs"); flush_all (); exit 1) in + if not (List.is_empty pfs) then vernac_error (str "There are pending proofs") + in match !Flags.compilation_mode with | BuildVo -> let long_f_dot_v = ensure_v f in @@ -311,5 +324,8 @@ let compile verbosely f = let compile v f = ignore(CoqworkmgrApi.get 1); - compile v f; + begin + try compile v f + with any -> Topfmt.print_err_exn any + end; CoqworkmgrApi.giveback 1 diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml index c25dd55fb..a1835959c 100644 --- a/vernac/topfmt.ml +++ b/vernac/topfmt.ml @@ -260,6 +260,29 @@ let init_color_output () = *) let emacs_logger = gen_logger Emacs.quote_info Emacs.quote_warning + +(* This is specific to the toplevel *) +let pr_loc loc = + if Loc.is_ghost loc then str"" + else + let fname = loc.Loc.fname in + if CString.equal fname "" then + Loc.(str"Toplevel input, characters " ++ int loc.bp ++ + str"-" ++ int loc.ep ++ str":") + else + Loc.(str"File " ++ str "\"" ++ str fname ++ str "\"" ++ + str", line " ++ int loc.line_nb ++ str", characters " ++ + int (loc.bp-loc.bol_pos) ++ str"-" ++ int (loc.ep-loc.bol_pos) ++ + str":") + +let print_err_exn ?extra any = + let (e, info) = CErrors.push any in + let loc = Loc.get_loc info in + let msg_loc = pr_loc (Option.default Loc.ghost loc) in + let pre_hdr = pr_opt_no_spc (fun x -> x) extra ++ msg_loc in + let msg = CErrors.iprint (e, info) ++ fnl () in + std_logger ~pre_hdr Feedback.Error msg + (* Output to file, used only in extraction so a candidate for removal *) let ft_logger old_logger ft ?loc level mesg = let id x = x in diff --git a/vernac/topfmt.mli b/vernac/topfmt.mli index 909dd7077..6c8e0ae2f 100644 --- a/vernac/topfmt.mli +++ b/vernac/topfmt.mli @@ -36,19 +36,22 @@ val get_depth_boxes : unit -> int option val set_margin : int option -> unit val get_margin : unit -> int option -(** Headers for tagging *) -val err_hdr : Pp.std_ppcmds -val ann_hdr : Pp.std_ppcmds - (** Console display of feedback, we may add some location information *) val std_logger : ?pre_hdr:Pp.std_ppcmds -> Feedback.level -> Pp.std_ppcmds -> unit val emacs_logger : ?pre_hdr:Pp.std_ppcmds -> Feedback.level -> Pp.std_ppcmds -> unit +(** Color output *) val init_color_output : unit -> unit val clear_styles : unit -> unit val parse_color_config : string -> unit val dump_tags : unit -> (string * Terminal.style) list +(** Error printing *) +(* To be deprecated when we can fully move to feedback-based error + printing. *) +val pr_loc : Loc.t -> Pp.std_ppcmds +val print_err_exn : ?extra:Pp.std_ppcmds -> exn -> unit + (** [with_output_to_file file f x] executes [f x] with logging redirected to a file [file] *) val with_output_to_file : string -> ('a -> 'b) -> 'a -> 'b -- cgit v1.2.3 From e699313c7a3829016c853b2a1541c2e9151d709c Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 25 Apr 2017 18:36:09 +0200 Subject: [toplevel] Remove unused parameter from `Vernac.process_expr`. --- toplevel/coqloop.ml | 2 +- toplevel/vernac.ml | 8 ++++---- toplevel/vernac.mli | 10 +++++++--- 3 files changed, 12 insertions(+), 8 deletions(-) diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml index b608488c8..9a4f476bd 100644 --- a/toplevel/coqloop.ml +++ b/toplevel/coqloop.ml @@ -300,7 +300,7 @@ let do_vernac sid = resynch_buffer top_buffer; try let input = (top_buffer.tokens, None) in - Vernac.process_expr sid top_buffer.tokens (read_sentence sid (fst input)) + Vernac.process_expr sid (read_sentence sid (fst input)) with | Stm.End_of_input | CErrors.Quit -> top_stderr (fnl ()); raise CErrors.Quit diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 8bcf2114b..9ff320ee8 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -117,7 +117,7 @@ let vernac_error msg = (* Stm.End_of_input -> true *) (* | _ -> false *) -let rec interp_vernac sid po (loc,com) = +let rec interp_vernac sid (loc,com) = let interp = function | VernacLoad (verbosely, fname) -> let fname = Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (str x)) fname in @@ -195,7 +195,7 @@ and load_vernac verbosely sid file = Option.iter (vernac_echo loc) in_echo; checknav_simple (loc, ast); - let nsid = Flags.silently (interp_vernac !rsid in_pa) (loc, ast) in + let nsid = Flags.silently (interp_vernac !rsid) (loc, ast) in rsid := nsid done; !rsid @@ -221,9 +221,9 @@ and load_vernac verbosely sid file = of a new state label). An example of state-preserving command is one coming from the query panel of Coqide. *) -let process_expr sid po loc_ast = +let process_expr sid loc_ast = checknav_deep loc_ast; - interp_vernac sid po loc_ast + interp_vernac sid loc_ast (* XML output hooks *) let (f_xml_start_library, xml_start_library) = Hook.make ~default:ignore () diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli index e75f8f9e8..bbc095c68 100644 --- a/toplevel/vernac.mli +++ b/toplevel/vernac.mli @@ -8,11 +8,15 @@ (** Parsing of vernacular. *) -(** Reads and executes vernac commands from a stream. *) -val process_expr : Stateid.t -> Pcoq.Gram.coq_parsable -> Vernacexpr.vernac_expr Loc.located -> Stateid.t +(** [process_expr sid cmd] Executes vernac command [cmd]. Callers are + expected to handle and print errors in form of exceptions, however + care is taken so the state machine is left in a consistent + state. *) +val process_expr : Stateid.t -> Vernacexpr.vernac_expr Loc.located -> Stateid.t (** [load_vernac echo sid file] Loads [file] on top of [sid], will - echo the commands if [echo] is set. *) + echo the commands if [echo] is set. Callers are expected to handle + and print errors in form of exceptions. *) val load_vernac : bool -> Stateid.t -> string -> Stateid.t (** Compile a vernac file, (f is assumed without .v suffix) *) -- cgit v1.2.3 From 16b8ed513cb4a0e68456fc968a9a0107d02acf5a Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 25 Apr 2017 13:50:54 -0400 Subject: Give andb_prop a simpler proof No need to use `discriminate`. This is the hopefully uncontroversial part of https://github.com/coq/coq/pull/401. --- theories/Init/Datatypes.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Init/Datatypes.v b/theories/Init/Datatypes.v index 11d80dbc3..41e1fea61 100644 --- a/theories/Init/Datatypes.v +++ b/theories/Init/Datatypes.v @@ -65,7 +65,7 @@ Infix "&&" := andb : bool_scope. Lemma andb_prop : forall a b:bool, andb a b = true -> a = true /\ b = true. Proof. - destruct a; destruct b; intros; split; try (reflexivity || discriminate). + destruct a, b; repeat split; assumption. Qed. Hint Resolve andb_prop: bool. -- cgit v1.2.3 From 11aaa1fd8230a347f1dca1a0f349ea7c7f2768c3 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 10 Jun 2016 18:58:24 -0400 Subject: Add support for transparent abstract (no syntax) This is a small change that allows a transparent version of tclABSTRACT. Additionally, it factors the machinery of [abstract] through a plugin-accessible function which allows alternate continuations (other than exact_no_check. It might be nice to factor it further, into a cache_term function that caches a term, and a separate bit that calls cache_term with the result of running the tactic. --- tactics/tactics.ml | 28 ++++++++++++++++++---------- tactics/tactics.mli | 4 +++- 2 files changed, 21 insertions(+), 11 deletions(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e79258582..19627eb53 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4907,7 +4907,7 @@ let shrink_entry sign const = } in (const, args) -let abstract_subproof id gk tac = +let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK = let open Tacticals.New in let open Tacmach.New in let open Proofview.Notations in @@ -4957,8 +4957,8 @@ let abstract_subproof id gk tac = else (const, List.rev (Context.Named.to_instance Constr.mkVar sign)) in let args = List.map EConstr.of_constr args in - let cd = Entries.DefinitionEntry const in - let decl = (cd, IsProof Lemma) in + let cd = Entries.DefinitionEntry { const with Entries.const_entry_opaque = opaque } in + let decl = (cd, if opaque then IsProof Lemma else IsDefinition Definition) in let cst () = (** do not compute the implicit arguments, it may be costly *) let () = Impargs.make_implicit_args false in @@ -4976,18 +4976,21 @@ let abstract_subproof id gk tac = Entries.(snd (Future.force const.const_entry_body)) in let solve = Proofview.tclEFFECTS effs <*> - exact_no_check (applist (lem, args)) + tacK lem args in let tac = if not safe then Proofview.mark_as_unsafe <*> solve else solve in Sigma.Unsafe.of_pair (tac, evd) end } +let abstract_subproof id gk tac ?(opaque=true) = + cache_term_by_tactic_then id gk ~opaque:opaque tac (fun lem args -> exact_no_check (applist (lem, args))) + let anon_id = Id.of_string "anonymous" -let tclABSTRACT name_op tac = +let name_op_to_name name_op object_kind suffix = let open Proof_global in - let default_gk = (Global, false, Proof Theorem) in - let s, gk = match name_op with + let default_gk = (Global, false, object_kind) in + match name_op with | Some s -> (try let _, gk, _ = current_proof_statement () in s, gk with NoCurrentProof -> s, default_gk) @@ -4995,9 +4998,14 @@ let tclABSTRACT name_op tac = let name, gk = try let name, gk, _ = current_proof_statement () in name, gk with NoCurrentProof -> anon_id, default_gk in - add_suffix name "_subproof", gk - in - abstract_subproof s gk tac + add_suffix name suffix, gk + +let tclABSTRACT ?(opaque=true) name_op tac = + let open Proof_global in + let s, gk = if opaque + then name_op_to_name name_op (Proof Theorem) "_subproof" + else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in + abstract_subproof s gk tac ~opaque:opaque let unify ?(state=full_transparent_state) x y = Proofview.Goal.s_enter { s_enter = begin fun gl -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index ba4a9706d..d206011ee 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -401,7 +401,9 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr - val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic -val tclABSTRACT : Id.t option -> unit Proofview.tactic -> unit Proofview.tactic +val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic + +val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Id.t -> unit Proofview.tactic val specialize_eqs : Id.t -> unit Proofview.tactic -- cgit v1.2.3 From 5f3d20dc53ffd0537a84c93acd761c3c69081342 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 10 Jun 2016 19:12:49 -0400 Subject: Add transparent_abstract tactic --- doc/refman/RefMan-ltac.tex | 14 +++++++++----- plugins/ltac/extratactics.ml4 | 13 +++++++++++++ test-suite/success/transparent_abstract.v | 21 +++++++++++++++++++++ theories/Init/Prelude.v | 2 +- 4 files changed, 44 insertions(+), 6 deletions(-) create mode 100644 test-suite/success/transparent_abstract.v diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex index 9378529cb..46274e12f 100644 --- a/doc/refman/RefMan-ltac.tex +++ b/doc/refman/RefMan-ltac.tex @@ -1087,8 +1087,8 @@ Fail all:let n:= numgoals in guard n=2. Reset Initial. \end{coq_eval} -\subsubsection[Proving a subgoal as a separate lemma]{Proving a subgoal as a separate lemma\tacindex{abstract}\comindex{Qed exporting} -\index{Tacticals!abstract@{\tt abstract}}} +\subsubsection[Proving a subgoal as a separate lemma]{Proving a subgoal as a separate lemma\tacindex{abstract}\tacindex{transparent\_abstract}\comindex{Qed exporting} +\index{Tacticals!abstract@{\tt abstract}}\index{Tacticals!transparent\_abstract@{\tt transparent\_abstract}}} From the outside ``\texttt{abstract \tacexpr}'' is the same as {\tt solve \tacexpr}. Internally it saves an auxiliary lemma called @@ -1114,13 +1114,17 @@ on. This can be obtained thanks to the option below. {\tt Set Shrink Abstract} \end{quote} -When set, all lemmas generated through \texttt{abstract {\tacexpr}} are -quantified only over the variables that appear in the term constructed by -\texttt{\tacexpr}. +When set, all lemmas generated through \texttt{abstract {\tacexpr}} +and \texttt{transparent\_abstract {\tacexpr}} are quantified only over the +variables that appear in the term constructed by \texttt{\tacexpr}. \begin{Variants} \item \texttt{abstract {\tacexpr} using {\ident}}.\\ Give explicitly the name of the auxiliary lemma. +\item \texttt{transparent\_abstract {\tacexpr}}.\\ + Save the subproof in a transparent lemma rather than an opaque one. +\item \texttt{transparent\_abstract {\tacexpr} using {\ident}}.\\ + Give explicitly the name of the auxiliary transparent lemma. \end{Variants} \ErrMsg \errindex{Proof is not complete} diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 38fdfb759..a96623a5f 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -815,6 +815,19 @@ TACTIC EXTEND destauto | [ "destauto" "in" hyp(id) ] -> [ destauto_in id ] END +(**********************************************************************) + +(**********************************************************************) +(* A version of abstract constructing transparent terms *) +(* Introduced by Jason Gross and Benjamin Delaware in June 2016 *) +(**********************************************************************) + +TACTIC EXTEND transparent_abstract +| [ "transparent_abstract" tactic3(t) ] -> [ Proofview.Goal.nf_enter { enter = fun gl -> + Tactics.tclABSTRACT ~opaque:false None (Tacinterp.tactic_of_value ist t) } ] +| [ "transparent_abstract" tactic3(t) "using" ident(id) ] -> [ Proofview.Goal.nf_enter { enter = fun gl -> + Tactics.tclABSTRACT ~opaque:false (Some id) (Tacinterp.tactic_of_value ist t) } ] +END (* ********************************************************************* *) diff --git a/test-suite/success/transparent_abstract.v b/test-suite/success/transparent_abstract.v new file mode 100644 index 000000000..ff4509c4a --- /dev/null +++ b/test-suite/success/transparent_abstract.v @@ -0,0 +1,21 @@ +Class by_transparent_abstract {T} (x : T) := make_by_transparent_abstract : T. +Hint Extern 0 (@by_transparent_abstract ?T ?x) => change T; transparent_abstract exact_no_check x : typeclass_instances. + +Goal True /\ True. +Proof. + split. + transparent_abstract exact I using foo. + let x := (eval hnf in foo) in constr_eq x I. + let x := constr:(ltac:(constructor) : True) in + let T := type of x in + let x := constr:(_ : by_transparent_abstract x) in + let x := (eval cbv delta [by_transparent_abstract] in (let y : T := x in y)) in + pose x as x'. + simpl in x'. + let v := eval cbv [x'] in x' in tryif constr_eq v I then fail 0 else idtac. + hnf in x'. + let v := eval cbv [x'] in x' in tryif constr_eq v I then idtac else fail 0. + exact x'. +Defined. +Check eq_refl : I = foo. +Eval compute in foo. diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v index c58d23dad..e71a8774e 100644 --- a/theories/Init/Prelude.v +++ b/theories/Init/Prelude.v @@ -23,4 +23,4 @@ Declare ML Module "cc_plugin". Declare ML Module "ground_plugin". Declare ML Module "recdef_plugin". (* Default substrings not considered by queries like SearchAbout *) -Add Search Blacklist "_subproof" "Private_". +Add Search Blacklist "_subproof" "_subterm" "Private_". -- cgit v1.2.3 From 84845f766d9b9d532f615352fbc8a0e78e1727e9 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Mon, 27 Mar 2017 13:28:17 -0400 Subject: Mark transparent_abstract as risky in docs As per Enrico's request. --- doc/refman/RefMan-ltac.tex | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex index 46274e12f..c2f52e23b 100644 --- a/doc/refman/RefMan-ltac.tex +++ b/doc/refman/RefMan-ltac.tex @@ -1121,10 +1121,17 @@ variables that appear in the term constructed by \texttt{\tacexpr}. \begin{Variants} \item \texttt{abstract {\tacexpr} using {\ident}}.\\ Give explicitly the name of the auxiliary lemma. + Use this feature at your own risk; explicitly named and reused subterms + don't play well with asynchronous proofs. \item \texttt{transparent\_abstract {\tacexpr}}.\\ Save the subproof in a transparent lemma rather than an opaque one. + Use this feature at your own risk; building computationally relevant terms + with tactics is fragile. \item \texttt{transparent\_abstract {\tacexpr} using {\ident}}.\\ Give explicitly the name of the auxiliary transparent lemma. + Use this feature at your own risk; building computationally relevant terms + with tactics is fragile, and explicitly named and reused subterms + don't play well with asynchronous proofs. \end{Variants} \ErrMsg \errindex{Proof is not complete} -- cgit v1.2.3 From 12f34b2ebfcbe958ba53b49399c3fcaf01f7a18c Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 11 Apr 2017 09:50:55 -0400 Subject: Generalize cache_term_by_tactic_then This will allow a cache_term tactic that doesn't suffer from the Not_found anomalies of abstract in typeclass resolution. --- tactics/tactics.ml | 7 +++++-- tactics/tactics.mli | 2 +- 2 files changed, 6 insertions(+), 3 deletions(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 19627eb53..20de56645 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4907,7 +4907,7 @@ let shrink_entry sign const = } in (const, args) -let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK = +let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK = let open Tacticals.New in let open Tacmach.New in let open Proofview.Notations in @@ -4927,7 +4927,10 @@ let cache_term_by_tactic_then id gk ?(opaque=true) tac tacK = else (Context.Named.add d s1,s2)) global_sign (Context.Named.empty, empty_named_context_val) in let id = next_global_ident_away id (pf_ids_of_hyps gl) in - let concl = it_mkNamedProd_or_LetIn (Proofview.Goal.concl gl) sign in + let concl = match goal_type with + | None -> Proofview.Goal.concl gl + | Some ty -> ty in + let concl = it_mkNamedProd_or_LetIn concl sign in let concl = try flush_and_check_evars !evdref concl with Uninstantiated_evar _ -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index d206011ee..082812c5a 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -401,7 +401,7 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr - val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic -val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic +val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> ?goal_type:(constr option) -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic -- cgit v1.2.3 From b348a11ccc4913598b72e4ecbb58811bcccd7bfc Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 11 Apr 2017 12:34:07 -0400 Subject: Make opaque optional only for tclABSTRACT Also move named arguments to the beginning of the functions. As per https://github.com/coq/coq/pull/201#discussion_r110928302 --- tactics/tactics.ml | 8 ++++---- tactics/tactics.mli | 2 +- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 20de56645..8f791cdcf 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4907,7 +4907,7 @@ let shrink_entry sign const = } in (const, args) -let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK = +let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK = let open Tacticals.New in let open Tacmach.New in let open Proofview.Notations in @@ -4985,8 +4985,8 @@ let cache_term_by_tactic_then id gk ?(opaque=true) ?(goal_type=None) tac tacK = Sigma.Unsafe.of_pair (tac, evd) end } -let abstract_subproof id gk tac ?(opaque=true) = - cache_term_by_tactic_then id gk ~opaque:opaque tac (fun lem args -> exact_no_check (applist (lem, args))) +let abstract_subproof ~opaque id gk tac = + cache_term_by_tactic_then ~opaque:opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args))) let anon_id = Id.of_string "anonymous" @@ -5008,7 +5008,7 @@ let tclABSTRACT ?(opaque=true) name_op tac = let s, gk = if opaque then name_op_to_name name_op (Proof Theorem) "_subproof" else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in - abstract_subproof s gk tac ~opaque:opaque + abstract_subproof ~opaque:opaque s gk tac let unify ?(state=full_transparent_state) x y = Proofview.Goal.s_enter { s_enter = begin fun gl -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 082812c5a..07a803542 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -401,7 +401,7 @@ val generalize_dep : ?with_let:bool (** Don't lose let bindings *) -> constr - val unify : ?state:Names.transparent_state -> constr -> constr -> unit Proofview.tactic -val cache_term_by_tactic_then : Id.t -> Decl_kinds.goal_kind -> ?opaque:bool -> ?goal_type:(constr option) -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic +val cache_term_by_tactic_then : opaque:bool -> ?goal_type:(constr option) -> Id.t -> Decl_kinds.goal_kind -> unit Proofview.tactic -> (constr -> constr list -> unit Proofview.tactic) -> unit Proofview.tactic val tclABSTRACT : ?opaque:bool -> Id.t option -> unit Proofview.tactic -> unit Proofview.tactic -- cgit v1.2.3 From 1e046726dc9352f7979ebdeba0d750e44016fea5 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 11 Apr 2017 12:48:23 -0400 Subject: transparent abstract: Respond to review comment https://github.com/coq/coq/pull/201#discussion_r110952601 --- tactics/tactics.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 8f791cdcf..d02fe8665 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4986,7 +4986,7 @@ let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK = end } let abstract_subproof ~opaque id gk tac = - cache_term_by_tactic_then ~opaque:opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args))) + cache_term_by_tactic_then ~opaque id gk tac (fun lem args -> exact_no_check (applist (lem, args))) let anon_id = Id.of_string "anonymous" -- cgit v1.2.3 From e4262a89d7bc3d9b985d9a4a939f34176581abcb Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 11 Apr 2017 13:05:54 -0400 Subject: transparent abstract: Respond to review comment https://github.com/coq/coq/pull/201#discussion_r110957570 --- tactics/tactics.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index d02fe8665..465481703 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5008,7 +5008,7 @@ let tclABSTRACT ?(opaque=true) name_op tac = let s, gk = if opaque then name_op_to_name name_op (Proof Theorem) "_subproof" else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in - abstract_subproof ~opaque:opaque s gk tac + abstract_subproof ~opaque s gk tac let unify ?(state=full_transparent_state) x y = Proofview.Goal.s_enter { s_enter = begin fun gl -> -- cgit v1.2.3 From 2fc235b0e00c028d2e91c696926ad339232d51e3 Mon Sep 17 00:00:00 2001 From: Vadim Zaliva Date: Wed, 26 Apr 2017 11:21:10 -0700 Subject: Small typo in comment --- theories/Init/Logic.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v index f659c31f9..9ae9dde28 100644 --- a/theories/Init/Logic.v +++ b/theories/Init/Logic.v @@ -223,7 +223,7 @@ Proof. Qed. (** [(IF_then_else P Q R)], written [IF P then Q else R] denotes - either [P] and [Q], or [~P] and [Q] *) + either [P] and [Q], or [~P] and [R] *) Definition IF_then_else (P Q R:Prop) := P /\ Q \/ ~ P /\ R. -- cgit v1.2.3 From fa27856d2d4ac0f55b99b9406b74301057deb0aa Mon Sep 17 00:00:00 2001 From: Matej Košík Date: Thu, 27 Apr 2017 09:57:09 +0200 Subject: contracting the type of "Pfedit.solve_by_implicit_tactic" --- proofs/pfedit.mli | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index f9fb0b76d..7622a8776 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -190,4 +190,4 @@ val declare_implicit_tactic : unit Proofview.tactic -> unit val clear_implicit_tactic : unit -> unit (* Raise Exit if cannot solve *) -val solve_by_implicit_tactic : unit -> (env -> Evd.evar_map -> Evd.evar -> Evd.evar_map * EConstr.constr) option +val solve_by_implicit_tactic : unit -> Pretyping.inference_hook option -- cgit v1.2.3 From bc6de5aa4e00dfc19c4866de4876f6213546fa5c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 24 Mar 2017 11:31:50 +0100 Subject: Code cleaning in unification algorithm for universes. This patch is only moving code around and expliciting statically the invariants of the functions, so it should be 1:1 equivalent to the other one. Amongst other goodies, the unification function is not recursive anymore, which ensures that it will terminate. --- engine/uState.ml | 125 ++++++++++++++++++++++++++++--------------------------- 1 file changed, 64 insertions(+), 61 deletions(-) diff --git a/engine/uState.ml b/engine/uState.ml index c66af02bb..353d8976d 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -131,84 +131,87 @@ let instantiate_variable l b v = exception UniversesDiffer let process_universe_constraints ctx cstrs = + let open Univ in let univs = ctx.uctx_universes in let vars = ref ctx.uctx_univ_variables in let normalize = Universes.normalize_universe_opt_subst vars in - let rec unify_universes fo l d r local = + let is_local l = Univ.LMap.mem l !vars in + let varinfo x = + match Univ.Universe.level x with + | None -> Inl x + | Some l -> Inr l + in + let equalize_variables fo l l' r r' local = + (** Assumes l = [l',0] and r = [r',0] *) + let () = + if is_local l' then + instantiate_variable l' r vars + else if is_local r' then + instantiate_variable r' l vars + else if not (UGraph.check_eq univs l r) then + (* Two rigid/global levels, none of them being local, + one of them being Prop/Set, disallow *) + if Univ.Level.is_small l' || Univ.Level.is_small r' then + raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None)) + else if fo then + raise UniversesDiffer + in + Univ.enforce_eq_level l' r' local + in + let equalize_universes l r local = match varinfo l, varinfo r with + | Inr l', Inr r' -> equalize_variables false l l' r r' local + | Inr l, Inl r | Inl r, Inr l -> + let alg = Univ.LSet.mem l ctx.uctx_univ_algebraic in + let inst = Univ.univ_level_rem l r r in + if alg then (instantiate_variable l inst vars; local) + else + let lu = Univ.Universe.make l in + if Univ.univ_level_mem l r then + Univ.enforce_leq inst lu local + else raise (Univ.UniverseInconsistency (Univ.Eq, lu, r, None)) + | Inl _, Inl _ (* both are algebraic *) -> + if UGraph.check_eq univs l r then local + else raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None)) + in + let unify_universes (l, d, r) local = let l = normalize l and r = normalize r in if Univ.Universe.equal l r then local else - let varinfo x = - match Univ.Universe.level x with - | None -> Inl x - | Some l -> Inr (l, Univ.LMap.mem l !vars, Univ.LSet.mem l ctx.uctx_univ_algebraic) - in - if d == Universes.ULe then + match d with + | Universes.ULe -> if UGraph.check_leq univs l r then (** Keep Prop/Set <= var around if var might be instantiated by prop or set later. *) - if Univ.Universe.is_level l then - match Univ.Universe.level r with - | Some r -> - Univ.Constraint.add (Option.get (Univ.Universe.level l),Univ.Le,r) local - | _ -> local - else local + match Univ.Universe.level l, Univ.Universe.level r with + | Some l, Some r -> + Univ.Constraint.add (l, Univ.Le, r) local + | _ -> local else - match Univ.Universe.level r with + begin match Univ.Universe.level r with | None -> error ("Algebraic universe on the right") - | Some rl -> - if Univ.Level.is_small rl then + | Some r' -> + if Univ.Level.is_small r' then let levels = Univ.Universe.levels l in - Univ.LSet.fold (fun l local -> - if Univ.Level.is_small l || Univ.LMap.mem l !vars then - unify_universes fo (Univ.Universe.make l) Universes.UEq r local - else raise (Univ.UniverseInconsistency (Univ.Le, Univ.Universe.make l, r, None))) - levels local + let fold l' local = + let l = Univ.Universe.make l' in + if Univ.Level.is_small l' || is_local l' then + equalize_variables false l l' r r' local + else raise (Univ.UniverseInconsistency (Univ.Le, l, r, None)) + in + Univ.LSet.fold fold levels local else Univ.enforce_leq l r local - else if d == Universes.ULub then - match varinfo l, varinfo r with - | (Inr (l, true, _), Inr (r, _, _)) - | (Inr (r, _, _), Inr (l, true, _)) -> - instantiate_variable l (Univ.Universe.make r) vars; - Univ.enforce_eq_level l r local - | Inr (_, _, _), Inr (_, _, _) -> - unify_universes true l Universes.UEq r local + end + | Universes.ULub -> + begin match Universe.level l, Universe.level r with + | Some l', Some r' -> + equalize_variables true l l' r r' local | _, _ -> assert false - else (* d = Universes.UEq *) - match varinfo l, varinfo r with - | Inr (l', lloc, _), Inr (r', rloc, _) -> - let () = - if lloc then - instantiate_variable l' r vars - else if rloc then - instantiate_variable r' l vars - else if not (UGraph.check_eq univs l r) then - (* Two rigid/global levels, none of them being local, - one of them being Prop/Set, disallow *) - if Univ.Level.is_small l' || Univ.Level.is_small r' then - raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None)) - else - if fo then - raise UniversesDiffer - in - Univ.enforce_eq_level l' r' local - | Inr (l, loc, alg), Inl r - | Inl r, Inr (l, loc, alg) -> - let inst = Univ.univ_level_rem l r r in - if alg then (instantiate_variable l inst vars; local) - else - let lu = Univ.Universe.make l in - if Univ.univ_level_mem l r then - Univ.enforce_leq inst lu local - else raise (Univ.UniverseInconsistency (Univ.Eq, lu, r, None)) - | _, _ (* One of the two is algebraic or global *) -> - if UGraph.check_eq univs l r then local - else raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None)) + end + | Universes.UEq -> equalize_universes l r local in let local = - Universes.Constraints.fold (fun (l,d,r) local -> unify_universes false l d r local) - cstrs Univ.Constraint.empty + Universes.Constraints.fold unify_universes cstrs Univ.Constraint.empty in !vars, local -- cgit v1.2.3 From 051f2c1be929a46a0713b47b072bb5be0a7558d0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 24 Mar 2017 15:50:21 +0100 Subject: Fast path when checking equality of universe levels in UState. We export the relevant level equality function in UGraph which is way faster than checking that each one is smaller than the other as universes. --- engine/uState.ml | 2 +- kernel/uGraph.mli | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/engine/uState.ml b/engine/uState.ml index 353d8976d..e27d0536d 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -148,7 +148,7 @@ let process_universe_constraints ctx cstrs = instantiate_variable l' r vars else if is_local r' then instantiate_variable r' l vars - else if not (UGraph.check_eq univs l r) then + else if not (UGraph.check_eq_level univs l' r') then (* Two rigid/global levels, none of them being local, one of them being Prop/Set, disallow *) if Univ.Level.is_small l' || Univ.Level.is_small r' then diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli index e95cf4d1c..ed52832fa 100644 --- a/kernel/uGraph.mli +++ b/kernel/uGraph.mli @@ -17,6 +17,7 @@ type universes = t type 'a check_function = universes -> 'a -> 'a -> bool val check_leq : universe check_function val check_eq : universe check_function +val check_eq_level : universe_level check_function (** The empty graph of universes *) val empty_universes : universes -- cgit v1.2.3 From 6d4f2222aa7ff6039f9f386cbc38201a0cd60c08 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Apr 2017 17:21:14 +0200 Subject: Document the API changes. --- dev/doc/changes.txt | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/dev/doc/changes.txt b/dev/doc/changes.txt index 7f915b781..8ea1638c9 100644 --- a/dev/doc/changes.txt +++ b/dev/doc/changes.txt @@ -51,6 +51,15 @@ In Constrexpr_ops: interpreting "(x y z:_)" as "(x:_) (y:_) (z:_)" while the second ones were preserving the original sharing of the type. +** Tactic API ** + +- pf_constr_of_global now returns a tactic instead of taking a continuation. + Thus it only generates one instance of the global reference, and it is the + caller's responsibility to perform a focus on the goal. + +- The tclWEAK_PROGRESS and tclNOTSAMEGOAL tacticals were removed. Their usecase + was very specific. Use tclPROGRESS instead. + ** Ltac API ** Many Ltac specific API has been moved in its own ltac/ folder. Amongst other -- cgit v1.2.3 From e2a8edaf595827af82be67a90c0c5b22c987abe5 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Tue, 22 Nov 2016 13:50:10 +0100 Subject: A refined solution to the beta-iota discrepancies between 8.4 and 8.5 "refine". There is a long story of commits trying to improve the compatibility between 8.4 and 8.5 refine, as discussed in https://github.com/coq/coq/pull/346. ac9c5986b77bf4a783f2bd0ad571645694c960e1 add beta-iota in hypotheses and conclusion 8afac4f87d9d7e3add1c19485f475bd2207bfde7 remove beta-iota in hypotheses 08e87eb96ab67ead60d92394eec6066d9b52e55e re-add beta-iota in hypotheses c9c54122d1d9493a965b483939e119d52121d5a6 re-remove beta-iota in hypotheses 9194180e2da0f7f9a2b2c7574bb7261cc69ead17 revert re-remove beta-iota in hypotheses 6bb352a6743c7332b9715ac15e95c806a58d101c re-re-remove beta-iota in hypotheses if <= 8.5 d8baa76d86eaa691a5386669596a6004bb44bb7a idem if = 8.5 The current commit tries to identify (one of?) the exact points of divergence between 8.4 and 8.5 refine, namely the types inferred for the variables of a pattern-matching problem. Note that for the conclusion of each new goal, there were a nf_betaiota in 8.4 done in function Evarutil.evars_to_metas, so the compatibility expects that such a nf_betaiota on the conclusion of each goal remains. --- pretyping/cases.ml | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 6bc2a4f94..8a49cd548 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1245,6 +1245,12 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn let typs = List.map2 RelDecl.set_name names cs_args in + (* Beta-iota-normalize types to better compatibility of refine with 8.4 behavior *) + (* This is a bit too strong I think, in the sense that what we would *) + (* really like is to have beta-iota reduction only at the positions where *) + (* parameters are substituted *) + let typs = List.map (map_type (nf_betaiota !(pb.evdref))) typs in + (* We build the matrix obtained by expanding the matching on *) (* "C x1..xn as x" followed by a residual matching on eqn into *) (* a matching on "x1 .. xn eqn" *) -- cgit v1.2.3 From 9839375100365ea3bd28bfc2efdb701db7d83adb Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 5 Dec 2016 12:35:01 +0100 Subject: Test surgical use of beta-iota in the type of variables coming from pattern-matching for refine. --- test-suite/bugs/closed/5219.v | 10 ++++++++++ 1 file changed, 10 insertions(+) create mode 100644 test-suite/bugs/closed/5219.v diff --git a/test-suite/bugs/closed/5219.v b/test-suite/bugs/closed/5219.v new file mode 100644 index 000000000..f7cec1a0c --- /dev/null +++ b/test-suite/bugs/closed/5219.v @@ -0,0 +1,10 @@ +(* Test surgical use of beta-iota in the type of variables coming from + pattern-matching for refine *) + +Goal forall x : sigT (fun x => x = 1), True. + intro x; refine match x with + | existT _ x' e' => _ + end. + lazymatch goal with + | [ H : _ = _ |- _ ] => idtac + end. -- cgit v1.2.3 From cc1212c3cfbd9c39cbe981210758c67cf9095be2 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 27 Apr 2017 17:25:14 +0200 Subject: Tentative note in CHANGES about now applying βι while typing "match" branches. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In practice, this is almost invisible except when using "refine". So, in some sense, it is aligning the behavior of pretyping on the one of logic.ml's "refine" so that the more natural behavior of 8.4's refine on this issue is restored. --- CHANGES | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/CHANGES b/CHANGES index 60b88ea8d..8cb5573b2 100644 --- a/CHANGES +++ b/CHANGES @@ -18,6 +18,11 @@ Tactics missed before because of a missing normalization step. Hopefully this should be fairly uncommon. - "auto with real" can now discharge comparisons of literals +- The types of variables in patterns of "match" are now + beta-iota-reduced after type-checking. This has an impact on the + type of the variables that the tactic "refine" introduces in the + context, producing types a priori closer to the expectations. + Standard Library -- cgit v1.2.3 From 746066172b8ed508886feb20cee239920ca7a4c7 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 27 Apr 2017 17:33:51 +0200 Subject: Test for bug #5193: Uncaught exception Class_tactics.Search.ReachedLimitEx. --- test-suite/bugs/closed/5193.v | 14 ++++++++++++++ 1 file changed, 14 insertions(+) create mode 100644 test-suite/bugs/closed/5193.v diff --git a/test-suite/bugs/closed/5193.v b/test-suite/bugs/closed/5193.v new file mode 100644 index 000000000..cc8739afe --- /dev/null +++ b/test-suite/bugs/closed/5193.v @@ -0,0 +1,14 @@ +Class Eqdec A := eqdec : forall a b : A, {a=b}+{a<>b}. + +Typeclasses eauto := debug. +Set Typeclasses Debug Verbosity 2. + +Inductive Finx(n : nat) : Set := +| Fx1(i : nat)(e : n = S i) +| FxS(i : nat)(f : Finx i)(e : n = S i). + +Context `{Finx_eqdec : forall n, Eqdec (Finx n)}. + +Goal {x : Type & Eqdec x}. + eexists. + try typeclasses eauto 1 with typeclass_instances. -- cgit v1.2.3 From e574b4bdd974daa7d2ceecf799762be92fadff44 Mon Sep 17 00:00:00 2001 From: Paul Steckler Date: Thu, 27 Apr 2017 12:05:51 -0400 Subject: fix order of command-line arguments mentioned in Add LoadPath --- doc/refman/RefMan-oth.tex | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex index 56ce753cd..a25da7d92 100644 --- a/doc/refman/RefMan-oth.tex +++ b/doc/refman/RefMan-oth.tex @@ -697,8 +697,8 @@ which can be any valid path. \subsection[\tt Add LoadPath {\str} as {\dirpath}.]{\tt Add LoadPath {\str} as {\dirpath}.\comindex{Add LoadPath}\label{AddLoadPath}} -This command is equivalent to the command line option {\tt -Q {\dirpath} - {\str}}. It adds the physical directory {\str} to the current {\Coq} +This command is equivalent to the command line option {\tt -Q {\str} + {\dirpath}}. It adds the physical directory {\str} to the current {\Coq} loadpath and maps it to the logical directory {\dirpath}. \begin{Variants} @@ -707,8 +707,8 @@ Performs as {\tt Add LoadPath {\str} as {\dirpath}} but for the empty directory \end{Variants} \subsection[\tt Add Rec LoadPath {\str} as {\dirpath}.]{\tt Add Rec LoadPath {\str} as {\dirpath}.\comindex{Add Rec LoadPath}\label{AddRecLoadPath}} -This command is equivalent to the command line option {\tt -R {\dirpath} - {\str}}. It adds the physical directory {\str} and all its +This command is equivalent to the command line option {\tt -R {\str} + {\dirpath}}. It adds the physical directory {\str} and all its subdirectories to the current {\Coq} loadpath. \begin{Variants} -- cgit v1.2.3 From 34d8de84ceb853c98bc80a0623f9afeae317d75f Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:25:49 +0200 Subject: Locally disable some warnings. --- checker/declarations.ml | 2 ++ dev/top_printers.ml | 1 + kernel/nativecode.ml | 2 ++ kernel/nativelambda.ml | 3 +++ kernel/nativevalues.ml | 3 ++- lib/backtrace.ml | 1 + 6 files changed, 11 insertions(+), 1 deletion(-) diff --git a/checker/declarations.ml b/checker/declarations.ml index 1fe02c8b6..56d437c16 100644 --- a/checker/declarations.ml +++ b/checker/declarations.ml @@ -6,6 +6,7 @@ open Term (** Substitutions, code imported from kernel/mod_subst *) module Deltamap = struct + [@@@ocaml.warning "-32-34"] type t = delta_resolver let empty = MPmap.empty, KNmap.empty let is_empty (mm, km) = MPmap.is_empty mm && KNmap.is_empty km @@ -25,6 +26,7 @@ end let empty_delta_resolver = Deltamap.empty module Umap = struct + [@@@ocaml.warning "-32-34"] type 'a t = 'a umap_t let empty = MPmap.empty, MBImap.empty let is_empty (m1,m2) = MPmap.is_empty m1 && MBImap.is_empty m2 diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 474cc85c1..918e98f77 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -7,6 +7,7 @@ (************************************************************************) (* Printers for the ocaml toplevel. *) +[@@@ocaml.warning "-32"] open Util open Pp diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index d9659d681..ba80ff590 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -16,6 +16,8 @@ open Nativeinstr open Nativelambda open Pre_env +[@@@ocaml.warning "-32-37"] + (** This file defines the mllambda code generation phase of the native compiler. mllambda represents a fragment of ML, and can easily be printed to OCaml code. *) diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml index 366f9a0a6..fcb75c661 100644 --- a/kernel/nativelambda.ml +++ b/kernel/nativelambda.ml @@ -16,6 +16,9 @@ open Nativeinstr module RelDecl = Context.Rel.Declaration +(* I'm not messing with this stuff. *) +[@@@ocaml.warning "-32"] + (** This file defines the lambda code generation phase of the native compiler *) exception NotClosed diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml index 965ed67b0..8d5f6388c 100644 --- a/kernel/nativevalues.ml +++ b/kernel/nativevalues.ml @@ -334,6 +334,7 @@ let l_or accu x y = if is_int x && is_int y then no_check_l_or x y else accu x y +[@@@ocaml.warning "-37"] type coq_carry = | Caccu of t | C0 of t @@ -430,7 +431,7 @@ let addmuldiv accu x y z = if is_int x && is_int y && is_int z then no_check_addmuldiv x y z else accu x y z - +[@@@ocaml.warning "-34"] type coq_bool = | Baccu of t | Btrue diff --git a/lib/backtrace.ml b/lib/backtrace.ml index b3b8bdea2..be9f40c1f 100644 --- a/lib/backtrace.ml +++ b/lib/backtrace.ml @@ -5,6 +5,7 @@ (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) +[@@@ocaml.warning "-37"] type raw_frame = | Known_location of bool (* is_raise *) -- cgit v1.2.3 From 8a3cd2fe699540f1ae5a56917d0f6b951f81d731 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:29:35 +0200 Subject: Remove unused [rec] keywords --- checker/checker.ml | 2 +- checker/inductive.ml | 2 +- lib/cWarnings.ml | 2 +- plugins/ltac/pptactic.ml | 4 ++-- plugins/nsatz/nsatz.ml | 2 +- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/checker/checker.ml b/checker/checker.ml index 95a9ea78b..57e1806cf 100644 --- a/checker/checker.ml +++ b/checker/checker.ml @@ -221,7 +221,7 @@ let where = function | Some s -> if !Flags.debug then (str"in " ++ str s ++ str":" ++ spc ()) else (mt ()) -let rec explain_exn = function +let explain_exn = function | Stream.Failure -> hov 0 (anomaly_string () ++ str "uncaught Stream.Failure.") | Stream.Error txt -> diff --git a/checker/inductive.ml b/checker/inductive.ml index c4ffc141f..ae2f7de8a 100644 --- a/checker/inductive.ml +++ b/checker/inductive.ml @@ -149,7 +149,7 @@ let remember_subst u subst = (* Bind expected levels of parameters to actual levels *) (* Propagate the new levels in the signature *) -let rec make_subst env = +let make_subst env = let rec make subst = function | LocalDef _ :: sign, exp, args -> make subst (sign, exp, args) diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml index 2f569d284..6a28d1e00 100644 --- a/lib/cWarnings.ml +++ b/lib/cWarnings.ml @@ -166,7 +166,7 @@ let normalize_flags_string s = let flags = normalize_flags ~silent:false flags in string_of_flags flags -let rec parse_warnings items = +let parse_warnings items = CList.iter (fun (status, name) -> set_status ~name status) items (* For compatibility, we accept "none" *) diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index 39ae1f41d..b73b66e56 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -250,7 +250,7 @@ type 'a extra_genarg_printer = let pr_alias_key key = try let prods = (KNmap.find key !prnotation_tab).pptac_prods in - let rec pr = function + let pr = function | TacTerm s -> primitive s | TacNonTerm (_, symb, _) -> str (Printf.sprintf "(%s)" (pr_user_symbol symb)) in @@ -314,7 +314,7 @@ type 'a extra_genarg_printer = | Extend.Uentry _ | Extend.Uentryl _ -> str "ltac:(" ++ prtac (1, Any) arg ++ str ")" - let rec pr_targ prtac symb arg = match symb with + let pr_targ prtac symb arg = match symb with | Extend.Uentry tag when is_genarg tag (ArgumentType wit_tactic) -> prtac (1, Any) arg | Extend.Uentryl (_, l) -> prtac (l, Any) arg diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index db8f3e4b2..a5b016611 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -437,7 +437,7 @@ open Ideal that has the same size than lp and where true indicates an element that has been removed *) -let rec clean_pol lp = +let clean_pol lp = let t = Hashpol.create 12 in let find p = try Hashpol.find t p with -- cgit v1.2.3 From 2826683746569b9d78aa01e319315ab554e1619b Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:36:45 +0200 Subject: Fix omitted labels in function calls --- grammar/tacextend.mlp | 2 +- ide/coqOps.ml | 16 ++++++++-------- ide/coqide.ml | 10 +++++----- ide/ide_slave.ml | 2 +- ide/ideutils.ml | 2 +- ide/preferences.ml | 20 ++++++++++---------- ide/session.ml | 12 ++++++------ ide/wg_Command.ml | 6 +++--- ide/wg_Completion.ml | 26 +++++++++++++------------- ide/wg_Find.ml | 6 +++--- ide/wg_ProofView.ml | 6 +++--- ide/wg_ScriptView.ml | 4 ++-- ide/wg_Segment.ml | 6 +++--- kernel/reduction.ml | 4 ++-- lib/cWarnings.ml | 2 +- parsing/cLexer.ml4 | 18 +++++++++--------- plugins/extraction/extract_env.ml | 2 +- plugins/ltac/tacentries.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 18 +++++++++--------- pretyping/pretyping.ml | 2 +- pretyping/reductionops.ml | 6 ++++-- proofs/redexpr.ml | 2 +- stm/stm.ml | 4 ++-- tactics/class_tactics.ml | 6 +++--- tactics/hints.ml | 2 +- tools/coqc.ml | 4 ++-- tools/coqdep.ml | 2 +- toplevel/coqtop.ml | 4 ++-- toplevel/vernac.ml | 2 +- vernac/class.ml | 4 ++-- vernac/command.ml | 4 ++-- vernac/ind_tables.ml | 6 +++--- vernac/metasyntax.ml | 2 +- 33 files changed, 108 insertions(+), 106 deletions(-) diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp index 1dd8da12a..b14fba975 100644 --- a/grammar/tacextend.mlp +++ b/grammar/tacextend.mlp @@ -129,7 +129,7 @@ let declare_tactic loc tacname ~level classification clause = match clause with let se = <:expr< { Tacexpr.mltac_tactic = $entry$; Tacexpr.mltac_plugin = $plugin_name$ } >> in let gl = mlexpr_of_clause clause in let level = mlexpr_of_int level in - let obj = <:expr< fun () -> Tacentries.add_ml_tactic_notation $se$ $level$ $gl$ >> in + let obj = <:expr< fun () -> Tacentries.add_ml_tactic_notation $se$ ~{ level = $level$ } $gl$ >> in declare_str_items loc [ <:str_item< do { Tacenv.register_ml_tactic $se$ (Array.of_list $make_fun_clauses loc tacname clause$); diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 222b9eed9..259557f40 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -117,7 +117,7 @@ end = struct (b#get_iter_at_mark s.start)#offset (b#get_iter_at_mark s.stop)#offset (ellipsize - ((b#get_iter_at_mark s.start)#get_slice (b#get_iter_at_mark s.stop))) + ((b#get_iter_at_mark s.start)#get_slice ~stop:(b#get_iter_at_mark s.stop))) (String.concat "," (List.map str_of_flag s.flags)) (ellipsize (String.concat "," @@ -207,7 +207,7 @@ object (self) in List.iter (fun s -> set_index s (s.index + 1)) after; set_index s (document_length - List.length after); - ignore ((SentenceId.connect s)#changed self#on_changed); + ignore ((SentenceId.connect s)#changed ~callback:self#on_changed); document_length <- document_length + 1; List.iter (fun f -> f `INSERT) cbs @@ -221,8 +221,8 @@ object (self) List.iter (fun f -> f `REMOVE) cbs initializer - let _ = (Doc.connect doc)#pushed self#on_push in - let _ = (Doc.connect doc)#popped self#on_pop in + let _ = (Doc.connect doc)#pushed ~callback:self#on_push in + let _ = (Doc.connect doc)#popped ~callback:self#on_pop in () end @@ -273,15 +273,15 @@ object(self) else iter in let iter = sentence_start iter in - script#buffer#place_cursor iter; + script#buffer#place_cursor ~where:iter; ignore (script#scroll_to_iter ~use_align:true ~yalign:0. iter) in - let _ = segment#connect#clicked on_click in + let _ = segment#connect#clicked ~callback:on_click in () method private tooltip_callback ~x ~y ~kbd tooltip = - let x, y = script#window_to_buffer_coords `WIDGET x y in - let iter = script#get_iter_at_location x y in + let x, y = script#window_to_buffer_coords ~tag:`WIDGET ~x ~y in + let iter = script#get_iter_at_location ~x ~y in if iter#has_tag Tags.Script.tooltip then begin let s = let rec aux iter = diff --git a/ide/coqide.ml b/ide/coqide.ml index 25858acce..0b7567a5a 100644 --- a/ide/coqide.ml +++ b/ide/coqide.ml @@ -792,11 +792,11 @@ let coqtop_arguments sn = sn.messages#push Feedback.Error (Pp.str msg) else dialog#destroy () in - let _ = entry#connect#activate ok_cb in - let _ = ok#connect#clicked ok_cb in + let _ = entry#connect#activate ~callback:ok_cb in + let _ = ok#connect#clicked ~callback:ok_cb in let cancel = GButton.button ~stock:`CANCEL ~packing:box#add () in let cancel_cb () = dialog#destroy () in - let _ = cancel#connect#clicked cancel_cb in + let _ = cancel#connect#clicked ~callback:cancel_cb in dialog#show () let coqtop_arguments = cb_on_current_term coqtop_arguments @@ -1274,8 +1274,8 @@ let build_ui () = if b then toolbar#misc#show () else toolbar#misc#hide () in stick show_toolbar toolbar refresh_toolbar; - let _ = source_style#connect#changed refresh_style in - let _ = source_language#connect#changed refresh_language in + let _ = source_style#connect#changed ~callback:refresh_style in + let _ = source_language#connect#changed ~callback:refresh_language in (* Color configuration *) Tags.Script.incomplete#set_property diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index bf86db08f..ca0ef38d3 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -82,7 +82,7 @@ let add ((s,eid),(sid,verbose)) = let loc_ast = Stm.parse_sentence sid pa in let newid, rc = Stm.add ~ontop:sid verbose loc_ast in let rc = match rc with `NewTip -> CSig.Inl () | `Unfocus id -> CSig.Inr id in - ide_cmd_checks newid loc_ast; + ide_cmd_checks ~id:newid loc_ast; (* TODO: the "" parameter is a leftover of the times the protocol * used to include stderr/stdout output. * diff --git a/ide/ideutils.ml b/ide/ideutils.ml index da867e689..d7d38a5ec 100644 --- a/ide/ideutils.ml +++ b/ide/ideutils.ml @@ -58,7 +58,7 @@ let insert_with_tags (buf : #GText.buffer_skel) mark rmark tags text = let () = buf#insert ~iter:(buf#get_iter_at_mark mark) text in let start = buf#get_iter_at_mark mark in let stop = buf#get_iter_at_mark rmark in - let iter tag = buf#apply_tag tag start stop in + let iter tag = buf#apply_tag tag ~start ~stop in List.iter iter tags let insert_xml ?(mark = `INSERT) ?(tags = []) (buf : #GText.buffer_skel) msg = diff --git a/ide/preferences.ml b/ide/preferences.ml index f0fd45d77..e1fc4005b 100644 --- a/ide/preferences.ml +++ b/ide/preferences.ml @@ -73,8 +73,8 @@ end let stick (pref : 'a preference) (obj : #GObj.widget as 'obj) (cb : 'a -> unit) = let _ = cb pref#get in - let p_id = pref#connect#changed (fun v -> cb v) in - let _ = obj#misc#connect#destroy (fun () -> pref#connect#disconnect p_id) in + let p_id = pref#connect#changed ~callback:(fun v -> cb v) in + let _ = obj#misc#connect#destroy ~callback:(fun () -> pref#connect#disconnect p_id) in () (** Useful marshallers *) @@ -314,7 +314,7 @@ let attach_modifiers (pref : string preference) prefix = in GtkData.AccelMap.foreach change in - pref#connect#changed cb + pref#connect#changed ~callback:cb let modifier_for_navigation = new preference ~name:["modifier_for_navigation"] ~init:"" ~repr:Repr.(string) @@ -408,10 +408,10 @@ let background_color = new preference ~name:["background_color"] ~init:"cornsilk" ~repr:Repr.(string) let attach_bg (pref : string preference) (tag : GText.tag) = - pref#connect#changed (fun c -> tag#set_property (`BACKGROUND c)) + pref#connect#changed ~callback:(fun c -> tag#set_property (`BACKGROUND c)) let attach_fg (pref : string preference) (tag : GText.tag) = - pref#connect#changed (fun c -> tag#set_property (`FOREGROUND c)) + pref#connect#changed ~callback:(fun c -> tag#set_property (`FOREGROUND c)) let processing_color = new preference ~name:["processing_color"] ~init:"light blue" ~repr:Repr.(string) @@ -468,7 +468,7 @@ let create_tag name default = let iter table = let tag = GText.tag ~name () in table#add tag#as_tag; - ignore (pref#connect#changed (fun _ -> set_tag tag)); + ignore (pref#connect#changed ~callback:(fun _ -> set_tag tag)); set_tag tag; in List.iter iter [Tags.Script.table; Tags.Proof.table; Tags.Message.table]; @@ -601,8 +601,8 @@ object (self) box#pack italic#coerce; box#pack underline#coerce; let cb but obj = obj#set_sensitive (not but#active) in - let _ = fg_unset#connect#toggled (fun () -> cb fg_unset fg_color#misc) in - let _ = bg_unset#connect#toggled (fun () -> cb bg_unset bg_color#misc) in + let _ = fg_unset#connect#toggled ~callback:(fun () -> cb fg_unset fg_color#misc) in + let _ = bg_unset#connect#toggled ~callback:(fun () -> cb bg_unset bg_color#misc) in () end @@ -692,7 +692,7 @@ let configure ?(apply=(fun () -> ())) () = ~color:(Tags.color_of_string pref#get) ~packing:(table#attach ~left:1 ~top:i) () in - let _ = button#connect#color_set begin fun () -> + let _ = button#connect#color_set ~callback:begin fun () -> pref#set (Tags.string_of_color button#color) end in let reset _ = @@ -754,7 +754,7 @@ let configure ?(apply=(fun () -> ())) () = let button text (pref : bool preference) = let active = pref#get in let but = GButton.check_button ~label:text ~active ~packing:box#pack () in - ignore (but#connect#toggled (fun () -> pref#set but#active)) + ignore (but#connect#toggled ~callback:(fun () -> pref#set but#active)) in let () = button "Dynamic word wrap" dynamic_word_wrap in let () = button "Show line number" show_line_number in diff --git a/ide/session.ml b/ide/session.ml index 6262820e7..7aea75ac8 100644 --- a/ide/session.ml +++ b/ide/session.ml @@ -249,8 +249,8 @@ let make_table_widget ?sort cd cb = let () = data#set_headers_visible true in let () = data#set_headers_clickable true in let refresh clr = data#misc#modify_base [`NORMAL, `NAME clr] in - let _ = background_color#connect#changed refresh in - let _ = data#misc#connect#realize (fun () -> refresh background_color#get) in + let _ = background_color#connect#changed ~callback:refresh in + let _ = data#misc#connect#realize ~callback:(fun () -> refresh background_color#get) in let mk_rend c = GTree.cell_renderer_text [], ["text",c] in let cols = List.map2 (fun (_,c) (_,n,v) -> @@ -308,8 +308,8 @@ let create_errpage (script : Wg_ScriptView.script_view) : errpage = !callback errs; List.iter (fun (lno, msg) -> access (fun columns store -> let line = store#append () in - store#set line (find_int_col "Line" columns) lno; - store#set line (find_string_col "Error message" columns) msg)) + store#set ~row:line ~column:(find_int_col "Line" columns) lno; + store#set ~row:line ~column:(find_string_col "Error message" columns) msg)) errs end method on_update ~callback:cb = callback := cb @@ -348,8 +348,8 @@ let create_jobpage coqtop coqops : jobpage = else false) else let line = store#append () in - store#set line column id; - store#set line (find_string_col "Job name" columns) job)) + store#set ~row:line ~column id; + store#set ~row:line ~column:(find_string_col "Job name" columns) job)) jobs end method on_update ~callback:cb = callback := cb diff --git a/ide/wg_Command.ml b/ide/wg_Command.ml index 3fcb7ce49..621c46b94 100644 --- a/ide/wg_Command.ml +++ b/ide/wg_Command.ml @@ -91,8 +91,8 @@ object(self) let result = GText.view ~packing:r_bin#add () in views <- (frame#coerce, result, combo#entry) :: views; let cb clr = result#misc#modify_base [`NORMAL, `NAME clr] in - let _ = background_color#connect#changed cb in - let _ = result#misc#connect#realize (fun () -> cb background_color#get) in + let _ = background_color#connect#changed ~callback:cb in + let _ = result#misc#connect#realize ~callback:(fun () -> cb background_color#get) in let cb ft = result#misc#modify_font (Pango.Font.from_string ft) in stick text_font result cb; result#misc#set_can_focus true; (* false causes problems for selection *) @@ -165,7 +165,7 @@ object(self) self#new_page_maker; self#new_query_aux ~grab_now:false (); frame#misc#hide (); - let _ = background_color#connect#changed self#refresh_color in + let _ = background_color#connect#changed ~callback:self#refresh_color in self#refresh_color background_color#get; ignore(notebook#event#connect#key_press ~callback:(fun ev -> if GdkEvent.Key.keyval ev = GdkKeysyms._Escape then (self#hide; true) diff --git a/ide/wg_Completion.ml b/ide/wg_Completion.ml index aeae3e1fd..3bb6b780e 100644 --- a/ide/wg_Completion.ml +++ b/ide/wg_Completion.ml @@ -154,7 +154,7 @@ object (self) let () = store#clear () in let iter prop = let iter = store#append () in - store#set iter column prop + store#set ~row:iter ~column prop in let () = current_completion <- (pref, props) in Proposals.iter iter props @@ -267,7 +267,7 @@ object (self) (** Position of view w.r.t. window *) let (ux, uy) = Gdk.Window.get_position view#misc#window in (** Relative buffer position to view *) - let (dx, dy) = view#window_to_buffer_coords `WIDGET 0 0 in + let (dx, dy) = view#window_to_buffer_coords ~tag:`WIDGET ~x:0 ~y:0 in (** Iter position *) let iter = view#buffer#get_iter pos in let coords = view#get_iter_location iter in @@ -397,11 +397,11 @@ object (self) let () = self#select_first () in let () = obj#misc#show () in let () = self#manage_scrollbar () in - obj#resize 1 1 + obj#resize ~width:1 ~height:1 method private start_callback off = let (x, y, w, h) = self#coordinates (`OFFSET off) in - let () = obj#move x (y + 3 * h / 2) in + let () = obj#move ~x ~y:(y + 3 * h / 2) in () method private update_callback (off, word, props) = @@ -433,21 +433,21 @@ object (self) else false in (** Style handling *) - let _ = view#misc#connect#style_set self#refresh_style in + let _ = view#misc#connect#style_set ~callback:self#refresh_style in let _ = self#refresh_style () in let _ = data#set_resize_mode `PARENT in let _ = frame#set_resize_mode `PARENT in (** Callback to model *) - let _ = model#connect#start_completion self#start_callback in - let _ = model#connect#update_completion self#update_callback in - let _ = model#connect#end_completion self#end_callback in + let _ = model#connect#start_completion ~callback:self#start_callback in + let _ = model#connect#update_completion ~callback:self#update_callback in + let _ = model#connect#end_completion ~callback:self#end_callback in (** Popup interaction *) - let _ = view#event#connect#key_press key_cb in + let _ = view#event#connect#key_press ~callback:key_cb in (** Hiding the popup when necessary*) - let _ = view#misc#connect#hide obj#misc#hide in - let _ = view#event#connect#button_press (fun _ -> self#hide (); false) in - let _ = view#connect#move_cursor move_cb in - let _ = view#event#connect#focus_out (fun _ -> self#hide (); false) in + let _ = view#misc#connect#hide ~callback:obj#misc#hide in + let _ = view#event#connect#button_press ~callback:(fun _ -> self#hide (); false) in + let _ = view#connect#move_cursor ~callback:move_cb in + let _ = view#event#connect#focus_out ~callback:(fun _ -> self#hide (); false) in () end diff --git a/ide/wg_Find.ml b/ide/wg_Find.ml index 3d847ddcc..f84b9063b 100644 --- a/ide/wg_Find.ml +++ b/ide/wg_Find.ml @@ -186,8 +186,8 @@ class finder name (view : GText.view) = in let find_cb = generic_cb self#hide self#find_forward in let replace_cb = generic_cb self#hide self#replace in - let _ = find_entry#event#connect#key_press find_cb in - let _ = replace_entry#event#connect#key_press replace_cb in + let _ = find_entry#event#connect#key_press ~callback:find_cb in + let _ = replace_entry#event#connect#key_press ~callback:replace_cb in (** TextView interaction *) let view_cb ev = @@ -197,7 +197,7 @@ class finder name (view : GText.view) = else false else false in - let _ = view#event#connect#key_press view_cb in + let _ = view#event#connect#key_press ~callback:view_cb in () end diff --git a/ide/wg_ProofView.ml b/ide/wg_ProofView.ml index 3cbe58388..f801091cf 100644 --- a/ide/wg_ProofView.ml +++ b/ide/wg_ProofView.ml @@ -188,8 +188,8 @@ let proof_view () = let default_clipboard = GData.clipboard Gdk.Atom.primary in let _ = buffer#add_selection_clipboard default_clipboard in let cb clr = view#misc#modify_base [`NORMAL, `NAME clr] in - let _ = background_color#connect#changed cb in - let _ = view#misc#connect#realize (fun () -> cb background_color#get) in + let _ = background_color#connect#changed ~callback:cb in + let _ = view#misc#connect#realize ~callback:(fun () -> cb background_color#get) in let cb ft = view#misc#modify_font (Pango.Font.from_string ft) in stick text_font view cb; @@ -226,5 +226,5 @@ let proof_view () = (* Is there a better way to connect the signal ? *) (* Can this be done in the object constructor? *) let w_cb _ = pf#refresh ~force:false in - ignore (view#misc#connect#size_allocate w_cb); + ignore (view#misc#connect#size_allocate ~callback:w_cb); pf diff --git a/ide/wg_ScriptView.ml b/ide/wg_ScriptView.ml index 218cedb36..7c058febd 100644 --- a/ide/wg_ScriptView.ml +++ b/ide/wg_ScriptView.ml @@ -460,8 +460,8 @@ object (self) let _ = GtkSignal.connect ~sgn:move_line_signal ~callback obj in (** Plug on preferences *) let cb clr = self#misc#modify_base [`NORMAL, `NAME clr] in - let _ = background_color#connect#changed cb in - let _ = self#misc#connect#realize (fun () -> cb background_color#get) in + let _ = background_color#connect#changed ~callback:cb in + let _ = self#misc#connect#realize ~callback:(fun () -> cb background_color#get) in let cb b = self#set_wrap_mode (if b then `WORD else `NONE) in stick dynamic_word_wrap self cb; diff --git a/ide/wg_Segment.ml b/ide/wg_Segment.ml index dbc1740ef..d527a0d13 100644 --- a/ide/wg_Segment.ml +++ b/ide/wg_Segment.ml @@ -75,7 +75,7 @@ object (self) self#redraw (); end in - let _ = box#misc#connect#size_allocate cb in + let _ = box#misc#connect#size_allocate ~callback:cb in let clicked_cb ev = match model with | None -> true | Some md -> @@ -86,7 +86,7 @@ object (self) let () = clicked#call idx in true in - let _ = eventbox#event#connect#button_press clicked_cb in + let _ = eventbox#event#connect#button_press ~callback:clicked_cb in let cb show = if show then self#misc#show () else self#misc#hide () in stick show_progress_bar self cb; (** Initial pixmap *) @@ -102,7 +102,7 @@ object (self) | `SET (i, color) -> if self#misc#visible then self#fill_range color i (i + 1) in - md#changed changed_cb + md#changed ~callback:changed_cb method private fill_range color i j = match model with | None -> () diff --git a/kernel/reduction.ml b/kernel/reduction.ml index cd975ee9a..ba714ada2 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -487,14 +487,14 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (FInd (ind1,u1), FInd (ind2,u2)) -> if eq_ind ind1 ind2 then - (let cuniv = convert_instances false u1 u2 cuniv in + (let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv) else raise NotConvertible | (FConstruct ((ind1,j1),u1), FConstruct ((ind2,j2),u2)) -> if Int.equal j1 j2 && eq_ind ind1 ind2 then - (let cuniv = convert_instances false u1 u2 cuniv in + (let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv) else raise NotConvertible diff --git a/lib/cWarnings.ml b/lib/cWarnings.ml index 6a28d1e00..71e02b3ba 100644 --- a/lib/cWarnings.ml +++ b/lib/cWarnings.ml @@ -82,7 +82,7 @@ let set_all_warnings_status status = let set_category_status ~name status = let names = Hashtbl.find categories name in - List.iter (fun name -> set_warning_status name status) names + List.iter (fun name -> set_warning_status ~name status) names let is_all_keyword name = CString.equal name "all" let is_none_keyword s = CString.equal s "none" diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4 index 6d259e1b1..f25b394f0 100644 --- a/parsing/cLexer.ml4 +++ b/parsing/cLexer.ml4 @@ -345,13 +345,13 @@ let rec string loc ~comm_level bp len = parser if esc then string loc ~comm_level bp (store len '"') s else (loc, len) | [< ''('; s >] -> (parser - | [< ''*'; s >] -> - string loc - (Option.map succ comm_level) + | [< ''*'; s >] -> + let comm_level = Option.map succ comm_level in + string loc ~comm_level bp (store (store len '(') '*') s | [< >] -> - string loc comm_level bp (store len '(') s) s + string loc ~comm_level bp (store len '(') s) s | [< ''*'; s >] -> (parser | [< '')'; s >] -> @@ -361,9 +361,9 @@ let rec string loc ~comm_level bp len = parser | _ -> () in let comm_level = Option.map pred comm_level in - string loc comm_level bp (store (store len '*') ')') s + string loc ~comm_level bp (store (store len '*') ')') s | [< >] -> - string loc comm_level bp (store len '*') s) s + string loc ~comm_level bp (store len '*') s) s | [< ''\n' as c; s >] ep -> (* If we are parsing a comment, the string if not part of a token so we update the first line of the location. Otherwise, we update the last @@ -372,8 +372,8 @@ let rec string loc ~comm_level bp len = parser if Option.has_some comm_level then bump_loc_line loc ep else bump_loc_line_last loc ep in - string loc comm_level bp (store len c) s - | [< 'c; s >] -> string loc comm_level bp (store len c) s + string loc ~comm_level bp (store len c) s + | [< 'c; s >] -> string loc ~comm_level bp (store len c) s | [< _ = Stream.empty >] ep -> let loc = set_loc_pos loc bp ep in err loc Unterminated_string @@ -613,7 +613,7 @@ let rec next_token loc = parser bp | [< ' ('0'..'9' as c); len = number (store 0 c) >] ep -> comment_stop bp; (INT (get_buff len), set_loc_pos loc bp ep) - | [< ''\"'; (loc,len) = string loc None bp 0 >] ep -> + | [< ''\"'; (loc,len) = string loc ~comm_level:None bp 0 >] ep -> comment_stop bp; (STRING (get_buff len), set_loc_pos loc bp ep) | [< ' ('(' as c); diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml index 2b12462ad..322fbcea7 100644 --- a/plugins/extraction/extract_env.ml +++ b/plugins/extraction/extract_env.ml @@ -657,7 +657,7 @@ let extraction_library is_rec m = let l = List.rev (environment_until (Some dir_m)) in let select l (mp,struc) = if Visit.needed_mp mp - then (mp, extract_structure env mp no_delta true struc) :: l + then (mp, extract_structure env mp no_delta ~all:true struc) :: l else l in let struc = List.fold_left select [] l in diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index cd8c9e471..8cda73b4b 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -320,7 +320,7 @@ let add_tactic_notation local n prods e = let ids = List.map_filter cons_production_parameter prods in let prods = List.map interp_prod_item prods in let tac = Tacintern.glob_tactic_env ids (Global.env()) e in - add_glob_tactic_notation local n prods false ids tac + add_glob_tactic_notation local ~level:n prods false ids tac (**********************************************************************) (* ML Tactic entries *) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f3555ebc4..fc7963b2d 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1261,7 +1261,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in let rp = mk_upat_for env0 sigma0 (ise, rp) all_ok in let find_T, end_T = mk_tpattern_matcher ?raise_NoMatch sigma0 occ rp in - let concl = find_T env0 concl0 1 do_subst in + let concl = find_T env0 concl0 1 ~k:do_subst in let _ = end_T () in concl | Some (sigma, (X_In_T (hole, p) | In_X_In_T (hole, p))) -> @@ -1273,11 +1273,11 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = (* we start from sigma, so hole is considered a rigid head *) let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in - let concl = find_T env0 concl0 1 (fun env c _ h -> + let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h - (fun env _ -> do_subst env e_body))) in + ~k:(fun env _ -> do_subst env e_body))) in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_In_X_In_T (e, hole, p)) -> @@ -1289,11 +1289,11 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let find_X, end_X = mk_tpattern_matcher sigma noindex holep in let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in - let concl = find_T env0 concl0 1 (fun env c _ h -> + let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in - fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> - find_E env e_body h do_subst))) in + fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> + find_E env e_body h ~k:do_subst))) in let _ = end_E () in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_As_X_In_T (e, hole, p)) -> @@ -1306,10 +1306,10 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma occ holep in - let concl = find_TE env0 concl0 1 (fun env c _ h -> + let concl = find_TE env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in - fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> + fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> let e_sigma = unify_HO env sigma (EConstr.of_constr e_body) (EConstr.of_constr e) in let e_body = fs e_sigma e in do_subst env e_body e_body h))) in @@ -1352,7 +1352,7 @@ let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = let ise, u = mk_tpattern env sigma0 (ise,EConstr.Unsafe.to_constr t) ok L2R p in let find_U, end_U = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in - let concl = find_U env concl h (fun _ _ _ -> mkRel) in + let concl = find_U env concl h ~k:(fun _ _ _ -> mkRel) in let rdx, _, (sigma, uc, p) = end_U () in sigma, uc, EConstr.of_constr p, EConstr.of_constr concl, EConstr.of_constr rdx diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index ae87cd8c0..497bde340 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -195,7 +195,7 @@ let _ = (** Miscellaneous interpretation functions *) let interp_universe_level_name evd (loc,s) = let names, _ = Global.global_universe_names () in - if CString.string_contains s "." then + if CString.string_contains ~where:s ~what:"." then match List.rev (CString.split '.' s) with | [] -> anomaly (str"Invalid universe name " ++ str s) | n :: dp -> diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 270320538..da3add2e5 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1117,7 +1117,9 @@ let local_whd_state_gen flags sigma = whrec let raw_whd_state_gen flags env = - let f sigma s = fst (whd_state_gen (get_refolding_in_reduction ()) false flags env sigma s) in + let f sigma s = fst (whd_state_gen ~refold:(get_refolding_in_reduction ()) + ~tactic_mode:false + flags env sigma s) in f let stack_red_of_state_red f = @@ -1127,7 +1129,7 @@ let stack_red_of_state_red f = (* Drops the Cst_stack *) let iterate_whd_gen refold flags env sigma s = let rec aux t = - let (hd,sk),_ = whd_state_gen refold false flags env sigma (t,Stack.empty) in + let (hd,sk),_ = whd_state_gen ~refold ~tactic_mode:false flags env sigma (t,Stack.empty) in let whd_sk = Stack.map aux sk in Stack.zip sigma ~refold (hd,whd_sk) in aux s diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 0fe5c73f1..cb3538422 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -43,7 +43,7 @@ let cbv_native env sigma c = let whd_cbn flags env sigma t = let (state,_) = - (whd_state_gen true true flags env sigma (t,Reductionops.Stack.empty)) + (whd_state_gen ~refold:true ~tactic_mode:true flags env sigma (t,Reductionops.Stack.empty)) in Reductionops.Stack.zip ~refold:true sigma state diff --git a/stm/stm.ml b/stm/stm.ml index e823373f7..3738aa94a 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1380,7 +1380,7 @@ end = struct (* {{{ *) if not drop then begin let checked_proof = Future.chain ~pure:false future_proof (fun p -> let pobject, _ = - Proof_global.close_future_proof stop (Future.from_val ~fix_exn p) in + Proof_global.close_future_proof ~feedback_id:stop (Future.from_val ~fix_exn p) in let terminator = (* The one sent by master is an InvalidKey *) Lemmas.(standard_proof_terminator [] (mk_hook (fun _ _ -> ()))) in stm_vernac_interp stop @@ -2432,7 +2432,7 @@ let merge_proof_branch ~valid ?id qast keep brname = let id = VCS.new_node ?id () in VCS.merge id ~ours:(Qed (qed None)) brname; VCS.delete_branch brname; - VCS.propagate_sideff None; + VCS.propagate_sideff ~replay:None; `Ok | { VCS.kind = `Edit (mode, qed_id, master_id, _,_) } -> let ofp = diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 54e4405d1..c430e776a 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -357,12 +357,12 @@ let shelve_dependencies gls = let hintmap_of sigma hdc secvars concl = match hdc with - | None -> fun db -> Hint_db.map_none secvars db + | None -> fun db -> Hint_db.map_none ~secvars db | Some hdc -> fun db -> if Hint_db.use_dn db then (* Using dnet *) - Hint_db.map_eauto sigma secvars hdc concl db - else Hint_db.map_existential sigma secvars hdc concl db + Hint_db.map_eauto sigma ~secvars hdc concl db + else Hint_db.map_existential sigma ~secvars hdc concl db (** Hack to properly solve dependent evars that are typeclasses *) let rec e_trivial_fail_db only_classes db_list local_db secvars = diff --git a/tactics/hints.ml b/tactics/hints.ml index bcc068d3d..240178c9a 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1435,7 +1435,7 @@ let pr_hints_db (name,db,hintlist) = let pr_hint_list_for_head c = let dbs = current_db () in let validate (name, db) = - let hints = List.map (fun v -> 0, v) (Hint_db.map_all Id.Pred.full c db) in + let hints = List.map (fun v -> 0, v) (Hint_db.map_all ~secvars:Id.Pred.full c db) in (name, db, hints) in let valid_dbs = List.map validate dbs in diff --git a/tools/coqc.ml b/tools/coqc.ml index b12d48710..552a943c8 100644 --- a/tools/coqc.ml +++ b/tools/coqc.ml @@ -77,12 +77,12 @@ let parse_args () = | ("-v"|"--version") :: _ -> Usage.version 0 | ("-where") :: _ -> - Envars.set_coqlib (fun x -> x); + Envars.set_coqlib ~fail:(fun x -> x); print_endline (Envars.coqlib ()); exit 0 | ("-config" | "--config") :: _ -> - Envars.set_coqlib (fun x -> x); + Envars.set_coqlib ~fail:(fun x -> x); Usage.print_config (); exit 0 diff --git a/tools/coqdep.ml b/tools/coqdep.ml index a9f1b7376..1c1c1be6a 100644 --- a/tools/coqdep.ml +++ b/tools/coqdep.ml @@ -502,7 +502,7 @@ let coqdep () = let user = coqlib//"user-contrib" in if Sys.file_exists user then add_rec_dir_no_import add_coqlib_known user []; List.iter (fun s -> add_rec_dir_no_import add_coqlib_known s []) - (Envars.xdg_dirs (fun x -> Feedback.msg_warning (Pp.str x))); + (Envars.xdg_dirs ~warn:(fun x -> Feedback.msg_warning (Pp.str x))); List.iter (fun s -> add_rec_dir_no_import add_coqlib_known s []) Envars.coqpath; end; List.iter (fun (f,d) -> add_mli_known f d ".mli") !mliAccu; diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index 9cf075065..c3a755860 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -292,7 +292,7 @@ let init_gc () = between coqtop and coqc. *) let usage () = - Envars.set_coqlib CErrors.error; + Envars.set_coqlib ~fail:CErrors.error; init_load_path (); if !batch_mode then Usage.print_usage_coqc () else begin @@ -609,7 +609,7 @@ let init_toplevel arglist = (* If we have been spawned by the Spawn module, this has to be done * early since the master waits us to connect back *) Spawned.init_channels (); - Envars.set_coqlib CErrors.error; + Envars.set_coqlib ~fail:CErrors.error; if !print_where then (print_endline(Envars.coqlib ()); exit(exitcode ())); if !print_config then (Usage.print_config (); exit (exitcode ())); if !print_tags then (print_style_tags (); exit (exitcode ())); diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 9ff320ee8..48cd70f69 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -125,7 +125,7 @@ let rec interp_vernac sid (loc,com) = let f = Loadpath.locate_file fname in load_vernac verbosely sid f | v -> - let nsid, ntip = Stm.add sid (not !Flags.quiet) (loc,v) in + let nsid, ntip = Stm.add ~ontop:sid (not !Flags.quiet) (loc,v) in (* Main STM interaction *) if ntip <> `NewTip then diff --git a/vernac/class.ml b/vernac/class.ml index 95114ec39..104f3c1db 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -311,7 +311,7 @@ let add_coercion_hook poly local ref = | Global -> false | Discharge -> assert false in - let () = try_add_new_coercion ref stre poly in + let () = try_add_new_coercion ref ~local:stre poly in let msg = pr_global_env Id.Set.empty ref ++ str " is now a coercion" in Flags.if_verbose Feedback.msg_info msg @@ -324,6 +324,6 @@ let add_subclass_hook poly local ref = | Discharge -> assert false in let cl = class_of_global ref in - try_add_new_coercion_subclass cl stre poly + try_add_new_coercion_subclass cl ~local:stre poly let add_subclass_hook poly = Lemmas.mk_hook (add_subclass_hook poly) diff --git a/vernac/command.ml b/vernac/command.ml index b27d8a0a3..b2f5755ce 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -258,7 +258,7 @@ match local with let () = Universes.register_universe_binders gr pl in let () = assumption_message ident in let () = Typeclasses.declare_instance None false gr in - let () = if is_coe then Class.try_add_new_coercion gr local p in + let () = if is_coe then Class.try_add_new_coercion gr ~local p in let inst = if p (* polymorphic *) then Univ.UContext.instance ctx else Univ.Instance.empty @@ -752,7 +752,7 @@ let do_mutual_inductive indl poly prv finite = (* Declare the possible notations of inductive types *) List.iter Metasyntax.add_notation_interpretation ntns; (* Declare the coercions *) - List.iter (fun qid -> Class.try_add_new_coercion (locate qid) false poly) coes; + List.iter (fun qid -> Class.try_add_new_coercion (locate qid) ~local:false poly) coes; (* If positivity is assumed declares itself as unsafe. *) if Environ.deactivated_guard (Global.env ()) then Feedback.feedback Feedback.AddedAxiom else () diff --git a/vernac/ind_tables.ml b/vernac/ind_tables.ml index 85d0b6194..c6588684a 100644 --- a/vernac/ind_tables.ml +++ b/vernac/ind_tables.ml @@ -151,7 +151,7 @@ let define_individual_scheme_base kind suff f mode idopt (mind,i as ind) = let const = define mode id c mib.mind_polymorphic ctx in declare_scheme kind [|ind,const|]; const, Safe_typing.add_private - (Safe_typing.private_con_of_scheme kind (Global.safe_env()) [ind,const]) eff + (Safe_typing.private_con_of_scheme ~kind (Global.safe_env()) [ind,const]) eff let define_individual_scheme kind mode names (mind,i as ind) = match Hashtbl.find scheme_object_table kind with @@ -172,7 +172,7 @@ let define_mutual_scheme_base kind suff f mode names mind = consts, Safe_typing.add_private (Safe_typing.private_con_of_scheme - kind (Global.safe_env()) (Array.to_list schemes)) + ~kind (Global.safe_env()) (Array.to_list schemes)) eff let define_mutual_scheme kind mode names mind = @@ -185,7 +185,7 @@ let find_scheme_on_env_too kind ind = let s = String.Map.find kind (Indmap.find ind !scheme_map) in s, Safe_typing.add_private (Safe_typing.private_con_of_scheme - kind (Global.safe_env()) [ind, s]) + ~kind (Global.safe_env()) [ind, s]) Safe_typing.empty_private_constants let find_scheme ?(mode=InternalTacticRequest) kind (mind,i as ind) = diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index f805eeaa9..f86a0ef92 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -1196,7 +1196,7 @@ let inNotation : notation_obj -> obj = (**********************************************************************) let with_lib_stk_protection f x = - let fs = Lib.freeze `No in + let fs = Lib.freeze ~marshallable:`No in try let a = f x in Lib.unfreeze fs; a with reraise -> let reraise = CErrors.push reraise in -- cgit v1.2.3 From f67f84fc4831923aa9a2323b3a71c3a69fe61480 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:50:27 +0200 Subject: Use [method!] to override methods (warning 7) --- ide/preferences.ml | 2 +- ide/wg_Detachable.ml | 4 ++-- ide/wg_Notebook.ml | 2 +- ide/wg_ScriptView.ml | 8 ++++---- 4 files changed, 8 insertions(+), 8 deletions(-) diff --git a/ide/preferences.ml b/ide/preferences.ml index e1fc4005b..9fe9c6337 100644 --- a/ide/preferences.ml +++ b/ide/preferences.ml @@ -360,7 +360,7 @@ object ~name:["doc_url"] ~init:Coq_config.wwwrefman ~repr:Repr.(string) as super - method set v = + method! set v = if not (Flags.is_standard_doc_url v) && v <> use_default_doc_url && (* Extra hack to support links to last released doc version *) diff --git a/ide/wg_Detachable.ml b/ide/wg_Detachable.ml index 3d1b63dfa..cbc34462e 100644 --- a/ide/wg_Detachable.ml +++ b/ide/wg_Detachable.ml @@ -26,8 +26,8 @@ class detachable (obj : ([> Gtk.box] as 'a) Gobject.obj) = val mutable attached_cb = (fun _ -> ()) method child = frame#child - method add = frame#add - method pack ?from ?expand ?fill ?padding w = + method! add = frame#add + method! pack ?from ?expand ?fill ?padding w = if frame#all_children = [] then self#add w else raise (Invalid_argument "detachable#pack") diff --git a/ide/wg_Notebook.ml b/ide/wg_Notebook.ml index 08d7d1983..0e5284c2f 100644 --- a/ide/wg_Notebook.ml +++ b/ide/wg_Notebook.ml @@ -50,7 +50,7 @@ object(self) method pages = term_list - method remove_page index = + method! remove_page index = term_list <- Util.List.filteri (fun i x -> if i = index then kill_page x; i <> index) term_list; super#remove_page index diff --git a/ide/wg_ScriptView.ml b/ide/wg_ScriptView.ml index 7c058febd..7430f07d4 100644 --- a/ide/wg_ScriptView.ml +++ b/ide/wg_ScriptView.ml @@ -301,28 +301,28 @@ object (self) ~use_align:false ~yalign:0.75 ~within_margin:0.25 `INSERT (* HACK: missing gtksourceview features *) - method right_margin_position = + method! right_margin_position = let prop = { Gobject.name = "right-margin-position"; conv = Gobject.Data.int; } in Gobject.get prop obj - method set_right_margin_position pos = + method! set_right_margin_position pos = let prop = { Gobject.name = "right-margin-position"; conv = Gobject.Data.int; } in Gobject.set prop obj pos - method show_right_margin = + method! show_right_margin = let prop = { Gobject.name = "show-right-margin"; conv = Gobject.Data.boolean; } in Gobject.get prop obj - method set_show_right_margin show = + method! set_show_right_margin show = let prop = { Gobject.name = "show-right-margin"; conv = Gobject.Data.boolean; -- cgit v1.2.3 From a5f33b5fda592c2dfaed0c16d3773b35e7acab28 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:52:33 +0200 Subject: Disambiguate Polynomial.Hyp and Mfourier.Hyp -> Assum --- plugins/micromega/mfourier.ml | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml index f4f9b3c2f..377994415 100644 --- a/plugins/micromega/mfourier.ml +++ b/plugins/micromega/mfourier.ml @@ -99,7 +99,7 @@ module PSet = ISet module System = Hashtbl.Make(Vect) type proof = -| Hyp of int +| Assum of int | Elim of var * proof * proof | And of proof * proof @@ -134,7 +134,7 @@ exception SystemContradiction of proof let hyps prf = let rec hyps prf acc = match prf with - | Hyp i -> ISet.add i acc + | Assum i -> ISet.add i acc | Elim(_,prf1,prf2) | And(prf1,prf2) -> hyps prf1 (hyps prf2 acc) in hyps prf ISet.empty @@ -143,7 +143,7 @@ let hyps prf = (** Pretty printing *) let rec pp_proof o prf = match prf with - | Hyp i -> Printf.fprintf o "H%i" i + | Assum i -> Printf.fprintf o "H%i" i | Elim(v, prf1,prf2) -> Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2 | And(prf1,prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2 @@ -270,7 +270,7 @@ let norm_cstr {coeffs = v ; op = o ; cst = c} idx = (match o with | Eq -> Some c , Some c | Ge -> Some c , None) ; - prf = Hyp idx } + prf = Assum idx } (** [load_system l] takes a list of constraints of type [cstr_compat] @@ -285,7 +285,7 @@ let load_system l = let vars = List.fold_left (fun vrs (cstr,i) -> match norm_cstr cstr i with - | Contradiction -> raise (SystemContradiction (Hyp i)) + | Contradiction -> raise (SystemContradiction (Assum i)) | Redundant -> vrs | Cstr(vect,info) -> xadd_cstr vect info sys ; @@ -867,7 +867,7 @@ let mk_proof hyps prf = let rec mk_proof prf = match prf with - | Hyp i -> [ ([i, Int 1] , List.nth hyps i) ] + | Assum i -> [ ([i, Int 1] , List.nth hyps i) ] | Elim(v,prf1,prf2) -> let prfsl = mk_proof prf1 -- cgit v1.2.3 From 95239fa33c5da60080833dabc3ced246680dfdf9 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:56:05 +0200 Subject: Rename Sos_lib.(||) -> parser_or to avoid (deprecated) Pervasives.or --- plugins/micromega/sos.ml | 19 +++++++++++-------- plugins/micromega/sos_lib.ml | 2 +- 2 files changed, 12 insertions(+), 9 deletions(-) diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml index cc89e2b9d..7e716258c 100644 --- a/plugins/micromega/sos.ml +++ b/plugins/micromega/sos.ml @@ -282,14 +282,14 @@ let poly_variables (p:poly) = (* Order monomials for human presentation. *) (* ------------------------------------------------------------------------- *) -let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 or x1 = x2 && k1 > k2;; +let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 || x1 = x2 && k1 > k2;; let humanorder_monomial = let rec ord l1 l2 = match (l1,l2) with _,[] -> true | [],_ -> false - | h1::t1,h2::t2 -> humanorder_varpow h1 h2 or h1 = h2 && ord t1 t2 in - fun m1 m2 -> m1 = m2 or + | h1::t1,h2::t2 -> humanorder_varpow h1 h2 || h1 = h2 && ord t1 t2 in + fun m1 m2 -> m1 = m2 || ord (sort humanorder_varpow (graph m1)) (sort humanorder_varpow (graph m2));; @@ -466,6 +466,7 @@ let token s = >> (fun ((_,t),_) -> t);; let decimal = + let (||) = parser_or in let numeral = some isnum in let decimalint = atleast 1 numeral >> ((o) Num.num_of_string implode) in let decimalfrac = atleast 1 numeral @@ -492,6 +493,7 @@ let parse_decimal = mkparser decimal;; (* ------------------------------------------------------------------------- *) let parse_sdpaoutput,parse_csdpoutput = + let (||) = parser_or in let vector = token "{" ++ listof decimal (token ",") "decimal" ++ token "}" >> (fun ((_,v),_) -> vector_of_list v) in @@ -512,6 +514,7 @@ let parse_sdpaoutput,parse_csdpoutput = (* ------------------------------------------------------------------------- *) let sdpa_run_succeeded = + let (||) = parser_or in let rec skipupto dscr prs inp = (dscr ++ prs >> snd || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in @@ -602,7 +605,7 @@ let run_csdp dbg obj mats = let csdp obj mats = let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (* Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline() *) @@ -653,7 +656,7 @@ let linear_program_basic a = let mats = List.map (fun j -> diagonal (column j a)) (1--n) and obj = vector_const (Int 1) m in let rv,res = run_csdp false obj mats in - if rv = 1 or rv = 2 then false + if rv = 1 || rv = 2 then false else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; @@ -667,7 +670,7 @@ let linear_program a b = let mats = diagonal b :: List.map (fun j -> diagonal (column j a)) (1--n) and obj = vector_const (Int 1) m in let rv,res = run_csdp false obj mats in - if rv = 1 or rv = 2 then false + if rv = 1 || rv = 2 then false else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; @@ -968,7 +971,7 @@ let run_csdp dbg nblocks blocksizes obj mats = let csdp nblocks blocksizes obj mats = let rv,res = run_csdp (!debugging) nblocks blocksizes obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (*Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline() *) @@ -1439,7 +1442,7 @@ let run_csdp dbg obj mats = let csdp obj mats = let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (* (Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline()) *) diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml index f54914f25..9ee3db6e0 100644 --- a/plugins/micromega/sos_lib.ml +++ b/plugins/micromega/sos_lib.ml @@ -525,7 +525,7 @@ let isspace,issep,isbra,issymb,isalpha,isnum,isalnum = and isalnum c = Array.get ctable (charcode c) >= 16 in isspace,issep,isbra,issymb,isalpha,isnum,isalnum;; -let (||) parser1 parser2 input = +let parser_or parser1 parser2 input = try parser1 input with Noparse -> parser2 input;; -- cgit v1.2.3 From 02d2f34e5c84f0169e884c07054a6fbfef9f365c Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:04:58 +0200 Subject: Remove some unused values and types --- checker/closure.ml | 16 -------- checker/declarations.ml | 11 ----- checker/inductive.ml | 7 ---- checker/typeops.ml | 3 -- checker/univ.ml | 7 ---- engine/termops.ml | 24 ----------- ide/coqOps.ml | 3 -- ide/ideutils.ml | 11 ----- interp/notation_ops.ml | 4 -- kernel/names.ml | 2 - kernel/uGraph.mli | 3 ++ kernel/univ.ml | 4 -- lib/feedback.ml | 2 - lib/stateid.ml | 1 - parsing/cLexer.ml4 | 8 ---- parsing/egramcoq.ml | 4 -- parsing/g_vernac.ml4 | 5 --- parsing/pcoq.ml | 7 ---- plugins/funind/functional_principles_proofs.ml | 6 --- plugins/ltac/extratactics.ml4 | 2 - plugins/ltac/profile_ltac.ml | 1 - plugins/ltac/tacentries.ml | 3 -- plugins/ltac/tacintern.ml | 6 --- plugins/ltac/tacinterp.ml | 6 --- plugins/nsatz/ideal.ml | 57 -------------------------- plugins/nsatz/nsatz.ml | 28 ------------- plugins/setoid_ring/newring.ml | 2 - plugins/ssrmatching/ssrmatching.ml4 | 21 ---------- plugins/ssrmatching/ssrmatching.mli | 1 - pretyping/unification.ml | 3 -- printing/ppconstr.ml | 4 -- printing/ppvernac.ml | 4 -- printing/printer.ml | 5 --- proofs/clenv.ml | 5 --- proofs/logic.ml | 5 --- proofs/tacmach.ml | 3 -- stm/stm.ml | 7 ---- tactics/hipattern.ml | 1 - tactics/tactics.ml | 2 - tools/coqdoc/output.mli | 1 - toplevel/vernac.ml | 1 - vernac/auto_ind_decl.ml | 2 - vernac/metasyntax.ml | 2 +- vernac/record.ml | 2 +- vernac/topfmt.ml | 12 ------ 45 files changed, 5 insertions(+), 309 deletions(-) diff --git a/checker/closure.ml b/checker/closure.ml index cef1d31a6..b8294e795 100644 --- a/checker/closure.ml +++ b/checker/closure.ml @@ -651,22 +651,6 @@ let drop_parameters depth n argstk = (** Projections and eta expansion *) -let rec get_parameters depth n argstk = - match argstk with - Zapp args::s -> - let q = Array.length args in - if n > q then Array.append args (get_parameters depth (n-q) s) - else if Int.equal n q then [||] - else Array.sub args 0 n - | Zshift(k)::s -> - get_parameters (depth-k) n s - | [] -> (* we know that n < stack_args_size(argstk) (if well-typed term) *) - if Int.equal n 0 then [||] - else raise Not_found (* Trying to eta-expand a partial application..., should do - eta expansion first? *) - | _ -> assert false - (* strip_update_shift_app only produces Zapp and Zshift items *) - let eta_expand_ind_stack env ind m s (f, s') = let mib = lookup_mind (fst ind) env in match mib.mind_record with diff --git a/checker/declarations.ml b/checker/declarations.ml index 56d437c16..ad93146d5 100644 --- a/checker/declarations.ml +++ b/checker/declarations.ml @@ -463,13 +463,6 @@ let is_opaque cb = match cb.const_body with let opaque_univ_context cb = force_lazy_constr_univs cb.const_body -let subst_rel_declaration sub (id,copt,t as x) = - let copt' = Option.smartmap (subst_mps sub) copt in - let t' = subst_mps sub t in - if copt == copt' && t == t' then x else (id,copt',t') - -let subst_rel_context sub = List.smartmap (subst_rel_declaration sub) - let subst_recarg sub r = match r with | Norec -> r | (Mrec(kn,i)|Imbr (kn,i)) -> let kn' = subst_ind sub kn in @@ -517,10 +510,6 @@ let subst_decl_arity f g sub ar = if x' == x then ar else TemplateArity x' -let map_decl_arity f g = function - | RegularArity a -> RegularArity (f a) - | TemplateArity a -> TemplateArity (g a) - let subst_rel_declaration sub = Term.map_rel_decl (subst_mps sub) diff --git a/checker/inductive.ml b/checker/inductive.ml index ae2f7de8a..8f23a38af 100644 --- a/checker/inductive.ml +++ b/checker/inductive.ml @@ -436,13 +436,6 @@ let eq_recarg r1 r2 = match r1, r2 with let eq_wf_paths = Rtree.equal eq_recarg -let pp_recarg = function - | Norec -> Pp.str "Norec" - | Mrec i -> Pp.str ("Mrec "^MutInd.to_string (fst i)) - | Imbr i -> Pp.str ("Imbr "^MutInd.to_string (fst i)) - -let pp_wf_paths = Rtree.pp_tree pp_recarg - let inter_recarg r1 r2 = match r1, r2 with | Norec, Norec -> Some r1 | Mrec i1, Mrec i2 diff --git a/checker/typeops.ml b/checker/typeops.ml index 173e19ce1..02d801741 100644 --- a/checker/typeops.ml +++ b/checker/typeops.ml @@ -85,9 +85,6 @@ let type_of_constant_knowing_parameters env cst paramtyps = let type_of_constant_type env t = type_of_constant_type_knowing_parameters env t [||] -let type_of_constant env cst = - type_of_constant_knowing_parameters env cst [||] - let judge_of_constant_knowing_parameters env (kn,u as cst) paramstyp = let _cb = try lookup_constant kn env diff --git a/checker/univ.ml b/checker/univ.ml index 668f3a058..fb1a0faa7 100644 --- a/checker/univ.ml +++ b/checker/univ.ml @@ -87,7 +87,6 @@ module HList = struct val exists : (elt -> bool) -> t -> bool val for_all : (elt -> bool) -> t -> bool val for_all2 : (elt -> elt -> bool) -> t -> t -> bool - val remove : elt -> t -> t val to_list : t -> elt list end @@ -128,12 +127,6 @@ module HList = struct | Nil -> [] | Cons (x, _, l) -> x :: to_list l - let rec remove x = function - | Nil -> nil - | Cons (y, _, l) -> - if H.eq x y then l - else cons y (remove x l) - end end diff --git a/engine/termops.ml b/engine/termops.ml index 64f4c6dc5..29dcddbb0 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -612,30 +612,6 @@ let adjust_app_array_size f1 l1 f2 l2 = let extras,restl1 = Array.chop (len1-len2) l1 in (mkApp (f1,extras), restl1, f2, l2) -(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate - subterms of [c]; it carries an extra data [l] (typically a name - list) which is processed by [g na] (which typically cons [na] to - [l]) at each binder traversal (with name [na]); it is not recursive - and the order with which subterms are processed is not specified *) - -let map_constr_with_named_binders g f l c = match kind_of_term c with - | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _) -> c - | Cast (c,k,t) -> mkCast (f l c, k, f l t) - | Prod (na,t,c) -> mkProd (na, f l t, f (g na l) c) - | Lambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c) - | LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c) - | App (c,al) -> mkApp (f l c, Array.map (f l) al) - | Proj (p,c) -> mkProj (p, f l c) - | Evar (e,al) -> mkEvar (e, Array.map (f l) al) - | Case (ci,p,c,bl) -> mkCase (ci, f l p, f l c, Array.map (f l) bl) - | Fix (ln,(lna,tl,bl)) -> - let l' = Array.fold_left (fun l na -> g na l) l lna in - mkFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl)) - | CoFix(ln,(lna,tl,bl)) -> - let l' = Array.fold_left (fun l na -> g na l) l lna in - mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl)) - (* [map_constr_with_binders_left_to_right g f n c] maps [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 259557f40..b180aa556 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -128,9 +128,6 @@ end = struct end open SentenceId -let log_pp msg : unit task = - Coq.lift (fun () -> Minilib.log_pp msg) - let log msg : unit task = Coq.lift (fun () -> Minilib.log msg) diff --git a/ide/ideutils.ml b/ide/ideutils.ml index d7d38a5ec..a08ab07b5 100644 --- a/ide/ideutils.ml +++ b/ide/ideutils.ml @@ -35,17 +35,6 @@ let flash_info = let flash_context = status#new_context ~name:"Flash" in (fun ?(delay=5000) s -> flash_context#flash ~delay s) -let xml_to_string xml = - let open Xml_datatype in - let buf = Buffer.create 1024 in - let rec iter = function - | PCData s -> Buffer.add_string buf s - | Element (_, _, children) -> - List.iter iter children - in - let () = iter xml in - Buffer.contents buf - let insert_with_tags (buf : #GText.buffer_skel) mark rmark tags text = (** FIXME: LablGTK2 does not export the C insert_with_tags function, so that it has to reimplement its own helper function. Unluckily, it relies on diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 8b4fadb5a..d08fb107b 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -1141,10 +1141,6 @@ let term_of_binder = function | Name id -> GVar (Loc.ghost,id) | Anonymous -> GHole (Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) -type glob_decl2 = - (name, cases_pattern) Util.union * Decl_kinds.binding_kind * - glob_constr option * glob_constr - let match_notation_constr u c (metas,pat) = let terms,binders,termlists,binderlists = match_ false u ([],[]) metas ([],[],[],[]) c pat in diff --git a/kernel/names.ml b/kernel/names.ml index 5c10badbe..811b4a62a 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -542,7 +542,6 @@ module KerPair = struct end module SyntacticOrd = struct - type t = kernel_pair let compare x y = match x, y with | Same knx, Same kny -> KerName.compare knx kny | Dual (knux,kncx), Dual (knuy,kncy) -> @@ -865,7 +864,6 @@ struct let hash (c, b) = (if b then 0 else 1) + Constant.hash c module SyntacticOrd = struct - type t = constant * bool let compare (c, b) (c', b') = if b = b' then Constant.SyntacticOrd.compare c c' else -1 let equal (c, b as x) (c', b' as x') = diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli index e95cf4d1c..c8ac7df5c 100644 --- a/kernel/uGraph.mli +++ b/kernel/uGraph.mli @@ -61,3 +61,6 @@ val pr_universes : (Level.t -> Pp.std_ppcmds) -> universes -> Pp.std_ppcmds val dump_universes : (constraint_type -> string -> string -> unit) -> universes -> unit + +(** {6 Debugging} *) +val check_universes_invariants : universes -> unit diff --git a/kernel/univ.ml b/kernel/univ.ml index 09f884ecd..afe9cbe8d 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -440,10 +440,6 @@ struct let set = make Level.set let type1 = hcons (Level.set, 1) - let is_prop = function - | (l,0) -> Level.is_prop l - | _ -> false - let is_small = function | (l,0) -> Level.is_small l | _ -> false diff --git a/lib/feedback.ml b/lib/feedback.ml index df6fe3a62..0846e419b 100644 --- a/lib/feedback.ml +++ b/lib/feedback.ml @@ -40,8 +40,6 @@ type feedback = { contents : feedback_content; } -let default_route = 0 - (** Feeders *) let feeders : (int, feedback -> unit) Hashtbl.t = Hashtbl.create 7 diff --git a/lib/stateid.ml b/lib/stateid.ml index ae25735c5..c153f0e80 100644 --- a/lib/stateid.ml +++ b/lib/stateid.ml @@ -32,7 +32,6 @@ let compare = Int.compare module Self = struct type t = int let compare = compare - let equal = equal end module Set = Set.Make(Self) diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4 index f25b394f0..18942ac29 100644 --- a/parsing/cLexer.ml4 +++ b/parsing/cLexer.ml4 @@ -114,9 +114,6 @@ module Error = struct | UnsupportedUnicode x -> Printf.sprintf "Unsupported Unicode character (0x%x)" x) - (* Require to fix the Camlp4 signature *) - let print ppf x = Pp.pp_with ppf (Pp.str (to_string x)) - end open Error @@ -153,10 +150,6 @@ let bump_loc_line_last loc bol_pos = in Ploc.encl loc loc' -let set_loc_file loc fname = - Ploc.make_loc fname (Ploc.line_nb loc) (Ploc.bol_pos loc) - (Ploc.first_pos loc, Ploc.last_pos loc) (Ploc.comment loc) - (* For some reason, the [Ploc.after] function of Camlp5 does not update line numbers, so we define our own function that does it. *) let after loc = @@ -434,7 +427,6 @@ let push_char c = real_push_char c let push_string s = Buffer.add_string current_comment s -let push_bytes s = Buffer.add_bytes current_comment s let null_comment s = let rec null i = diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 496b20002..65e99d1b9 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -80,10 +80,6 @@ let create_pos = function | None -> Extend.First | Some lev -> Extend.After (constr_level lev) -type gram_level = - gram_position option * gram_assoc option * string option * - (** for reinitialization: *) gram_reinit option - let find_position_gen current ensure assoc lev = match lev with | None -> diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 011565d86..3438635cf 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -756,11 +756,6 @@ GEXTEND Gram implicit_status = MaximallyImplicit}) items ] ]; - name_or_bang: [ - [ b = OPT "!"; id = name -> - not (Option.is_empty b), id - ] - ]; (* Same as [argument_spec_block], but with only implicit status and names *) more_implicits_block: [ [ name = name -> [(snd name, Vernacexpr.NotImplicit)] diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml index dad98e2e9..3f014c4c8 100644 --- a/parsing/pcoq.ml +++ b/parsing/pcoq.ml @@ -171,7 +171,6 @@ module Symbols : sig val sopt : G.symbol -> G.symbol val snterml : G.internal_entry * string -> G.symbol val snterm : G.internal_entry -> G.symbol - val snterml_level : G.symbol -> string end = struct let stoken tok = @@ -199,10 +198,6 @@ end = struct let slist1 x = Gramext.Slist1 x let sopt x = Gramext.Sopt x - let snterml_level = function - | Gramext.Snterml (_, l) -> l - | _ -> failwith "snterml_level" - end let camlp4_verbosity silent f x = @@ -211,8 +206,6 @@ let camlp4_verbosity silent f x = f x; warning_verbose := a -let camlp4_verbose f x = camlp4_verbosity (not !Flags.quiet) f x - (** Grammar extensions *) (** NB: [extend_statment = diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 8dae17d69..df81bc78c 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -19,12 +19,6 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration -let local_assum (na, t) = - RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t) - -let local_def (na, b, t) = - RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) - (* let msgnl = Pp.msgnl *) (* diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 38fdfb759..35cfe8b54 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -52,8 +52,6 @@ let replace_in_clause_maybe_by ist c1 c2 cl tac = let replace_term ist dir_opt c cl = with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) -let clause = Pltac.clause_dft_concl - TACTIC EXTEND replace ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] -> [ replace_in_clause_maybe_by ist c1 c2 cl tac ] diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml index bcb28f77c..a853576f2 100644 --- a/plugins/ltac/profile_ltac.ml +++ b/plugins/ltac/profile_ltac.ml @@ -136,7 +136,6 @@ let feedback_results results = let format_sec x = (Printf.sprintf "%.3fs" x) let format_ratio x = (Printf.sprintf "%.1f%%" (100. *. x)) let padl n s = ws (max 0 (n - utf8_length s)) ++ str s -let padr n s = str s ++ ws (max 0 (n - utf8_length s)) let padr_with c n s = let ulength = utf8_length s in str (utf8_sub s 0 n) ++ str (String.make (max 0 (n - ulength)) c) diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index 8cda73b4b..ef1d69d35 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -88,9 +88,6 @@ let rec parse_user_entry s sep = else Uentry s -let arg_list = function Rawwit t -> Rawwit (ListArg t) -let arg_opt = function Rawwit t -> Rawwit (OptArg t) - let interp_entry_name interp symb = let rec eval = function | Ulist1 e -> Ulist1 (eval e) diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 3f83f104e..75227def0 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -118,12 +118,6 @@ let intern_constr_reference strict ist = function GRef (loc,locate_global_with_alias lqid,None), if strict then None else Some (CRef (r,None)) -let intern_move_location ist = function - | MoveAfter id -> MoveAfter (intern_hyp ist id) - | MoveBefore id -> MoveBefore (intern_hyp ist id) - | MoveFirst -> MoveFirst - | MoveLast -> MoveLast - (* Internalize an isolated reference in position of tactic *) let intern_isolated_global_tactic_reference r = diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 50f43931e..fcdf7bb2c 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -436,12 +436,6 @@ let interp_hyp_list_as_list ist env sigma (loc,id as x) = let interp_hyp_list ist env sigma l = List.flatten (List.map (interp_hyp_list_as_list ist env sigma) l) -let interp_move_location ist env sigma = function - | MoveAfter id -> MoveAfter (interp_hyp ist env sigma id) - | MoveBefore id -> MoveBefore (interp_hyp ist env sigma id) - | MoveFirst -> MoveFirst - | MoveLast -> MoveLast - let interp_reference ist env sigma = function | ArgArg (_,r) -> r | ArgVar (loc, id) -> diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml index b1ff59e78..41f2edfcf 100644 --- a/plugins/nsatz/ideal.ml +++ b/plugins/nsatz/ideal.ml @@ -153,7 +153,6 @@ module Make (P:Polynom.S) = struct type coef = P.t let coef0 = P.of_num (Num.Int 0) let coef1 = P.of_num (Num.Int 1) - let coefm1 = P.of_num (Num.Int (-1)) let string_of_coef c = "["^(P.to_string c)^"]" (*********************************************************************** @@ -252,59 +251,6 @@ let string_of_pol zeroP hdP tlP coefterm monterm string_of_coef in (stringP p true) - - -let print_pol zeroP hdP tlP coefterm monterm string_of_coef - dimmon string_of_exp lvar p = - - let rec print_mon m coefone = - let s=ref [] in - for i=1 to (dimmon m) do - (match (string_of_exp m i) with - "0" -> () - | "1" -> s:= (!s) @ [(getvar lvar (i-1))] - | e -> s:= (!s) @ [((getvar lvar (i-1)) ^ "^" ^ e)]); - done; - (match !s with - [] -> if coefone - then print_string "1" - else () - | l -> if coefone - then print_string (String.concat "*" l) - else (print_string "*"; - print_string (String.concat "*" l))) - and print_term t start = let a = coefterm t and m = monterm t in - match (string_of_coef a) with - "0" -> () - | "1" ->(match start with - true -> print_mon m true - |false -> (print_string "+ "; - print_mon m true)) - | "-1" ->(print_string "-";print_space();print_mon m true) - | c -> if (String.get c 0)='-' - then (print_string "- "; - print_string (String.sub c 1 - ((String.length c)-1)); - print_mon m false) - else (match start with - true -> (print_string c;print_mon m false) - |false -> (print_string "+ "; - print_string c;print_mon m false)) - and printP p start = - if (zeroP p) - then (if start - then print_string("0") - else ()) - else (print_term (hdP p) start; - if start then open_hovbox 0; - print_space(); - print_cut(); - printP (tlP p) false) - in open_hovbox 3; - printP p true; - print_flush() - - let stringP metadata (p : poly) = string_of_pol (fun p -> match p with [] -> true | _ -> false) @@ -595,9 +541,6 @@ let addS x l = l @ [x] (* oblige de mettre en queue sinon le certificat decon critical pairs/s-polynomials *) -let ordcpair ((i1,j1),m1) ((i2,j2),m2) = - compare_mon m1 m2 - module CPair = struct type t = (int * int) * Monomial.t diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index a5b016611..632b9dac1 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -22,7 +22,6 @@ open Utile let num_0 = Int 0 and num_1 = Int 1 and num_2 = Int 2 -and num_10 = Int 10 let numdom r = let r' = Ratio.normalize_ratio (ratio_of_num r) in @@ -35,7 +34,6 @@ module BigInt = struct type t = big_int let of_int = big_int_of_int let coef0 = of_int 0 - let coef1 = of_int 1 let of_num = Num.big_int_of_num let to_num = Num.num_of_big_int let equal = eq_big_int @@ -49,7 +47,6 @@ module BigInt = struct let div = div_big_int let modulo = mod_big_int let to_string = string_of_big_int - let to_int x = int_of_big_int x let hash x = try (int_of_big_int x) with Failure _ -> 1 @@ -61,15 +58,6 @@ module BigInt = struct then a else if lt a b then pgcd b a else pgcd b (modulo a b) - - (* signe du pgcd = signe(a)*signe(b) si non nuls. *) - let pgcd2 a b = - if equal a coef0 then b - else if equal b coef0 then a - else let c = pgcd (abs a) (abs b) in - if ((lt coef0 a)&&(lt b coef0)) - ||((lt coef0 b)&&(lt a coef0)) - then opp c else c end (* @@ -146,8 +134,6 @@ let mul = function | (Const n,q) when eq_num n num_1 -> q | (p,q) -> Mul(p,q) -let unconstr = mkRel 1 - let tpexpr = lazy (gen_constant "CC" ["setoid_ring";"Ring_polynom"] "PExpr") let ttconst = lazy (gen_constant "CC" ["setoid_ring";"Ring_polynom"] "PEc") @@ -271,20 +257,6 @@ let set_nvars_term nvars t = | Pow (t1,n) -> aux t1 nvars in aux t nvars -let string_of_term p = - let rec aux p = - match p with - | Zero -> "0" - | Const r -> string_of_num r - | Var v -> "x"^v - | Opp t1 -> "(-"^(aux t1)^")" - | Add (t1,t2) -> "("^(aux t1)^"+"^(aux t2)^")" - | Sub (t1,t2) -> "("^(aux t1)^"-"^(aux t2)^")" - | Mul (t1,t2) -> "("^(aux t1)^"*"^(aux t2)^")" - | Pow (t1,n) -> (aux t1)^"^"^(string_of_int n) - in aux p - - (*********************************************************************** Coefficients: recursive polynomials *) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index dd68eac24..6e072e831 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -279,8 +279,6 @@ let my_constant c = let my_reference c = lazy (Coqlib.gen_reference_in_modules "Ring" plugin_modules c) -let new_ring_path = - DirPath.make (List.map Id.of_string ["Ring_tac";plugin_dir;"Coq"]) let znew_ring_path = DirPath.make (List.map Id.of_string ["InitialRing";plugin_dir;"Coq"]) let zltac s = diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index fc7963b2d..60c199bf5 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -90,8 +90,6 @@ let pp s = !pp_ref s let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] -let get_index = function ArgArg i -> i | _ -> - CErrors.anomaly (str"Uninterpreted index") (* Toplevel constr must be globalized twice ! *) let glob_constr ist genv = function | _, Some ce -> @@ -304,8 +302,6 @@ let unif_EQ_args env sigma pa a = let unif_HO env ise p c = Evarconv.the_conv_x env p c ise -let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise - let unif_HO_args env ise0 pa i ca = let n = Array.length pa in let rec loop ise j = @@ -371,11 +367,6 @@ let unif_end env sigma0 ise0 pt ok = let s, uc', t = nf_open_term sigma0 ise2 t in s, Evd.union_evar_universe_context uc uc', t -let pf_unif_HO gl sigma pt p c = - let env = pf_env gl in - let ise = unif_HO env (create_evar_defs sigma) p c in - unif_end env (project gl) ise pt (fun _ -> true) - let unify_HO env sigma0 t1 t2 = let sigma = unif_HO env sigma0 t1 t2 in let sigma, uc, _ = unif_end env sigma0 sigma t2 (fun _ -> true) in @@ -440,10 +431,6 @@ let all_ok _ _ = true let proj_nparams c = try 1 + Recordops.find_projection_nparams (ConstRef c) with _ -> 0 -let isFixed c = match kind_of_term c with - | Var _ | Ind _ | Construct _ | Const _ | Proj _ -> true - | _ -> false - let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false @@ -917,13 +904,6 @@ let pp_pattern (sigma, p) = let pr_cpattern = pr_term let pr_rpattern _ _ _ = pr_pattern -let pr_option f = function None -> mt() | Some x -> f x -let pr_ssrpattern _ _ _ = pr_option pr_pattern -let pr_pattern_squarep = pr_option (fun r -> str "[" ++ pr_pattern r ++ str "]") -let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep -let pr_pattern_roundp = pr_option (fun r -> str "(" ++ pr_pattern r ++ str ")") -let pr_ssrpattern_roundp _ _ _ = pr_pattern_roundp - let wit_rpatternty = add_genarg "rpatternty" pr_pattern let glob_ssrterm gs = function @@ -1045,7 +1025,6 @@ let interp_wit wit ist gl x = let arg = interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in sigma, Value.cast (topwit wit) arg -let interp_constr = interp_wit wit_constr let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 894cdb943..1f984b160 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -226,7 +226,6 @@ val loc_of_cpattern : cpattern -> Loc.t val id_of_pattern : pattern -> Names.variable option val is_wildcard : cpattern -> bool val cpattern_of_id : Names.variable -> cpattern -val cpattern_of_id : Names.variable -> cpattern val pr_constr_pat : constr -> Pp.std_ppcmds val pf_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma val pf_unsafe_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 532cc8baa..4a3836d86 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1535,9 +1535,6 @@ let indirectly_dependent sigma c d decls = way to see that the second hypothesis depends indirectly over 2 *) List.exists (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) decls -let indirect_dependency sigma d decls = - decls |> List.filter (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id - let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) = let current_sigma = Sigma.to_evar_map current_sigma in let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 38eeda9b9..3041ac259 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -212,10 +212,6 @@ let tag_var = tag Tag.variable | Some (_,ExplByName id) -> str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")" - let pr_opt_type pr = function - | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt () - | t -> cut () ++ str ":" ++ pr t - let pr_opt_type_spc pr = function | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt () | t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index e4a87739b..58475630e 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -395,10 +395,6 @@ open Decl_kinds pr_opt (pr_guard_annot pr_lconstr_expr bl) guard ++ str":" ++ pr_spc_lconstr c) - let pr_priority = function - | None -> mt () - | Some i -> spc () ++ str "|" ++ spc () ++ int i - (**************************************) (* Pretty printer for vernac commands *) (**************************************) diff --git a/printing/printer.ml b/printing/printer.ml index 35ddf2e8c..93d221f03 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -370,11 +370,6 @@ let pr_context_limit_compact ?n env sigma = env ~init:(mt ()) in (sign_env ++ db_env) -(* compact printing an env (variables and de Bruijn). Separator: three - spaces between simple hyps, and newline otherwise *) -let pr_context_unlimited_compact env sigma = - pr_context_limit_compact env sigma - (* The number of printed hypothesis in a goal *) (* If [None], no limit *) let print_hyps_limit = ref (None : int option) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index f9ebc4233..605914a01 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -27,11 +27,6 @@ open Unification open Misctypes open Sigma.Notations -(* Abbreviations *) - -let pf_env = Refiner.pf_env -let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma c - (******************************************************************) (* Clausal environments *) diff --git a/proofs/logic.ml b/proofs/logic.ml index e6024785d..9ab8c34d8 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -97,11 +97,6 @@ let check_typability env sigma c = (instead of iterating on the list of identifier to be removed, which forces the user to give them in order). *) -let clear_hyps env sigma ids sign cl = - let evdref = ref (Evd.clear_metas sigma) in - let (hyps,cl) = Evarutil.clear_hyps_in_evi env evdref sign cl ids in - (hyps, cl, !evdref) - let clear_hyps2 env sigma ids sign t cl = let evdref = ref (Evd.clear_metas sigma) in let (hyps,t,cl) = Evarutil.clear_hyps2_in_evi env evdref sign t cl ids in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index b55f8ef11..97c5cda77 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -179,9 +179,6 @@ module New = struct let pf_unsafe_type_of gl t = pf_apply unsafe_type_of gl t - let pf_get_type_of gl t = - pf_apply (Retyping.get_type_of ~lax:true) gl t - let pf_type_of gl t = pf_apply type_of gl t diff --git a/stm/stm.ml b/stm/stm.ml index 3738aa94a..f773d60c5 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1041,13 +1041,6 @@ end = struct (* {{{ *) | `Stop x -> x | `Cont acc -> next acc - let back_safe () = - let id = - fold_until (fun n (id,_,_,_,_) -> - if n >= 0 && State.is_cached_and_valid id then `Stop id else `Cont (succ n)) - 0 (VCS.get_branch_pos (VCS.current_branch ())) in - backto id - let undo_vernac_classifier v = try match v with diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 851554b83..15b40b42d 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -40,7 +40,6 @@ let mkmeta n = Nameops.make_ident "X" (Some n) let meta1 = mkmeta 1 let meta2 = mkmeta 2 let meta3 = mkmeta 3 -let meta4 = mkmeta 4 let op2bool = function Some _ -> true | None -> false diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 9c2a1d850..53f8e4d5f 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5039,8 +5039,6 @@ end module New = struct open Proofview.Notations - let exact_proof c = exact_proof c - open Genredexpr open Locus diff --git a/tools/coqdoc/output.mli b/tools/coqdoc/output.mli index 853bc29aa..235f2588c 100644 --- a/tools/coqdoc/output.mli +++ b/tools/coqdoc/output.mli @@ -64,7 +64,6 @@ val keyword : string -> loc -> unit val ident : string -> loc option -> unit val sublexer : char -> loc -> unit val sublexer_in_doc : char -> unit -val initialize : unit -> unit val proofbox : unit -> unit diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 48cd70f69..9f6f77ef1 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -41,7 +41,6 @@ let vernac_echo loc in_chan = let open Loc in (* vernac parses the given stream, executes interpfun on the syntax tree it * parses, and is verbose on "primitives" commands if verbosely is true *) -let chan_beautify = ref stdout let beautify_suffix = ".beautified" let set_formatter_translator ch = diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index b9c4c6cc5..015552d68 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -28,8 +28,6 @@ open Proofview.Notations module RelDecl = Context.Rel.Declaration -let out_punivs = Univ.out_punivs - (**********************************************************************) (* Generic synthesis of boolean equality *) diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index f86a0ef92..bb5be4cb0 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -527,7 +527,7 @@ let warn_skip_spaces_curly = (fun () ->strbrk "Skipping spaces inside curly brackets") let rec drop_spacing = function - | UnpCut _ as u :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt + | UnpCut _ :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt | UnpTerminal s' :: fmt when String.equal s' (String.make (String.length s') ' ') -> warn_skip_spaces_curly (); drop_spacing fmt | fmt -> fmt diff --git a/vernac/record.ml b/vernac/record.ml index 8b4b7606f..53722b8f6 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -216,7 +216,7 @@ let warning_or_error coe indsp err = (pr_id fi ++ strbrk " cannot be defined because it is not typable.") in if coe then user_err ~hdr:"structure" st; - Flags.if_verbose Feedback.msg_info (hov 0 st) + warn_cannot_define_projection (hov 0 st) type field_status = | NoProjection of Name.t diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml index a1835959c..e44b585d7 100644 --- a/vernac/topfmt.ml +++ b/vernac/topfmt.ml @@ -106,8 +106,6 @@ module Tag = struct end -type logger = ?loc:Loc.t -> level -> std_ppcmds -> unit - let msgnl_with ?pre_hdr fmt strm = pp_with fmt (strm ++ fnl ()); Format.pp_print_flush fmt () @@ -283,16 +281,6 @@ let print_err_exn ?extra any = let msg = CErrors.iprint (e, info) ++ fnl () in std_logger ~pre_hdr Feedback.Error msg -(* Output to file, used only in extraction so a candidate for removal *) -let ft_logger old_logger ft ?loc level mesg = - let id x = x in - match level with - | Debug -> msgnl_with ft (make_body id dbg_hdr mesg) - | Info -> msgnl_with ft (make_body id info_hdr mesg) - | Notice -> msgnl_with ft mesg - | Warning -> old_logger ?loc level mesg - | Error -> old_logger ?loc level mesg - let with_output_to_file fname func input = (* XXX FIXME: redirect std_ft *) (* let old_logger = !logger in *) -- cgit v1.2.3 From 528c237b658dbba896a1fe0041990cc7fec9c4c8 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:07:32 +0200 Subject: Add [_] prefix to unused values which maybe should be kept --- engine/universes.ml | 2 +- plugins/ltac/tauto.ml | 2 +- tactics/term_dnet.ml | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/engine/universes.ml b/engine/universes.ml index ad5ff827b..ab561784c 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -732,7 +732,7 @@ let instantiate_with_lbound u lbound lower alg enforce (ctx, us, algs, insts, cs type constraints_map = (Univ.constraint_type * Univ.LMap.key) list Univ.LMap.t -let pr_constraints_map cmap = +let _pr_constraints_map (cmap:constraints_map) = LMap.fold (fun l cstrs acc -> Level.pr l ++ str " => " ++ prlist_with_sep spc (fun (d,r) -> pr_constraint_type d ++ Level.pr r) cstrs ++ diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index dc7ee6a23..e86d1c728 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -242,7 +242,7 @@ let tauto_uniform_unit_flags = { } (* This is the compatibility mode (not used) *) -let tauto_legacy_flags = { +let _tauto_legacy_flags = { binary_mode = true; binary_mode_bugged_detection = true; strict_in_contravariant_hyp = true; diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index 2c863c42a..726fd23b6 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -49,7 +49,7 @@ struct | DNil (* debug *) - let pr_dconstr f : 'a t -> std_ppcmds = function + let _pr_dconstr f : 'a t -> std_ppcmds = function | DRel -> str "*" | DSort -> str "Sort" | DRef _ -> str "Ref" -- cgit v1.2.3 From 9be835c4f16b3bc08ff9540a6854ced2d8185cb2 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:09:05 +0200 Subject: Remove unused constructors --- parsing/cLexer.ml4 | 5 +---- plugins/ssrmatching/ssrmatching.ml4 | 2 -- stm/stm.ml | 1 - 3 files changed, 1 insertion(+), 7 deletions(-) diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4 index 18942ac29..462905808 100644 --- a/parsing/cLexer.ml4 +++ b/parsing/cLexer.ml4 @@ -99,7 +99,6 @@ module Error = struct | Unterminated_string | Undefined_token | Bad_token of string - | UnsupportedUnicode of int exception E of t @@ -110,9 +109,7 @@ module Error = struct | Unterminated_comment -> "Unterminated comment" | Unterminated_string -> "Unterminated string" | Undefined_token -> "Undefined token" - | Bad_token tok -> Format.sprintf "Bad token %S" tok - | UnsupportedUnicode x -> - Printf.sprintf "Unsupported Unicode character (0x%x)" x) + | Bad_token tok -> Format.sprintf "Bad token %S" tok) end open Error diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 60c199bf5..f146fbb11 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -435,8 +435,6 @@ let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false -exception UndefPat - let hole_var = mkVar (id_of_string "_") let pr_constr_pat c0 = let rec wipe_evar c = diff --git a/stm/stm.ml b/stm/stm.ml index f773d60c5..8c1185b6d 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1274,7 +1274,6 @@ end = struct (* {{{ *) | RespBuiltProof of Proof_global.closed_proof_output * float | RespError of error | RespStates of (Stateid.t * State.partial_state) list - | RespDone let name = ref "proofworker" let extra_env () = !async_proofs_workers_extra_env -- cgit v1.2.3 From e1d2a898feacbe4bd363818f259bce5fd74c2ee7 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:09:23 +0200 Subject: Micromega: do not use Filename.temp_dir_path, remove unused values --- plugins/micromega/sos.ml | 257 ++----------------------------------------- plugins/micromega/sos_lib.ml | 2 +- 2 files changed, 11 insertions(+), 248 deletions(-) diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml index 7e716258c..e1ceabe9e 100644 --- a/plugins/micromega/sos.ml +++ b/plugins/micromega/sos.ml @@ -21,8 +21,6 @@ let debugging = ref false;; exception Sanity;; -exception Unsolvable;; - (* ------------------------------------------------------------------------- *) (* Turn a rational into a decimal string with d sig digits. *) (* ------------------------------------------------------------------------- *) @@ -99,28 +97,11 @@ let vector_const c n = if c =/ Int 0 then vector_0 n else (n,itlist (fun k -> k |-> c) (1--n) undefined :vector);; -let vector_1 = vector_const (Int 1);; - let vector_cmul c (v:vector) = let n = dim v in if c =/ Int 0 then vector_0 n else n,mapf (fun x -> c */ x) (snd v) -let vector_neg (v:vector) = (fst v,mapf minus_num (snd v) :vector);; - -let vector_add (v1:vector) (v2:vector) = - let m = dim v1 and n = dim v2 in - if m <> n then failwith "vector_add: incompatible dimensions" else - (n,combine (+/) (fun x -> x =/ Int 0) (snd v1) (snd v2) :vector);; - -let vector_sub v1 v2 = vector_add v1 (vector_neg v2);; - -let vector_dot (v1:vector) (v2:vector) = - let m = dim v1 and n = dim v2 in - if m <> n then failwith "vector_add: incompatible dimensions" else - foldl (fun a i x -> x +/ a) (Int 0) - (combine ( */ ) (fun x -> x =/ Int 0) (snd v1) (snd v2));; - let vector_of_list l = let n = List.length l in (n,itlist2 (|->) (1--n) l undefined :vector);; @@ -133,13 +114,6 @@ let matrix_0 (m,n) = ((m,n),undefined:matrix);; let dimensions (m:matrix) = fst m;; -let matrix_const c (m,n as mn) = - if m <> n then failwith "matrix_const: needs to be square" - else if c =/ Int 0 then matrix_0 mn - else (mn,itlist (fun k -> (k,k) |-> c) (1--n) undefined :matrix);; - -let matrix_1 = matrix_const (Int 1);; - let matrix_cmul c (m:matrix) = let (i,j) = dimensions m in if c =/ Int 0 then matrix_0 (i,j) @@ -152,8 +126,6 @@ let matrix_add (m1:matrix) (m2:matrix) = if d1 <> d2 then failwith "matrix_add: incompatible dimensions" else (d1,combine (+/) (fun x -> x =/ Int 0) (snd m1) (snd m2) :matrix);; -let matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);; - let row k (m:matrix) = let i,j = dimensions m in (j, @@ -166,20 +138,10 @@ let column k (m:matrix) = foldl (fun a (i,j) c -> if j = k then (i |-> c) a else a) undefined (snd m) : vector);; -let transp (m:matrix) = - let i,j = dimensions m in - ((j,i),foldl (fun a (i,j) c -> ((j,i) |-> c) a) undefined (snd m) :matrix);; - let diagonal (v:vector) = let n = dim v in ((n,n),foldl (fun a i c -> ((i,i) |-> c) a) undefined (snd v) : matrix);; -let matrix_of_list l = - let m = List.length l in - if m = 0 then matrix_0 (0,0) else - let n = List.length (List.hd l) in - (m,n),itern 1 l (fun v i -> itern 1 v (fun c j -> (i,j) |-> c)) undefined;; - (* ------------------------------------------------------------------------- *) (* Monomials. *) (* ------------------------------------------------------------------------- *) @@ -195,24 +157,8 @@ let monomial_var x = (x |=> 1 :monomial);; let (monomial_mul:monomial->monomial->monomial) = combine (+) (fun x -> false);; -let monomial_pow (m:monomial) k = - if k = 0 then monomial_1 - else mapf (fun x -> k * x) m;; - -let monomial_divides (m1:monomial) (m2:monomial) = - foldl (fun a x k -> tryapplyd m2 x 0 >= k && a) true m1;; - -let monomial_div (m1:monomial) (m2:monomial) = - let m = combine (+) (fun x -> x = 0) m1 (mapf (fun x -> -x) m2) in - if foldl (fun a x k -> k >= 0 && a) true m then m - else failwith "monomial_div: non-divisible";; - let monomial_degree x (m:monomial) = tryapplyd m x 0;; -let monomial_lcm (m1:monomial) (m2:monomial) = - (itlist (fun x -> x |-> max (monomial_degree x m1) (monomial_degree x m2)) - (union (dom m1) (dom m2)) undefined :monomial);; - let monomial_multidegree (m:monomial) = foldl (fun a x k -> k + a) 0 m;; let monomial_variables m = dom m;; @@ -252,12 +198,6 @@ let poly_cmmul (c,m) (p:poly) = let poly_mul (p1:poly) (p2:poly) = foldl (fun a m c -> poly_add (poly_cmmul (c,m) p2) a) poly_0 p1;; -let poly_div (p1:poly) (p2:poly) = - if not(poly_isconst p2) then failwith "poly_div: non-constant" else - let c = eval undefined p2 in - if c =/ Int 0 then failwith "poly_div: division by zero" - else poly_cmul (Int 1 // c) p1;; - let poly_square p = poly_mul p p;; let rec poly_pow p k = @@ -266,10 +206,6 @@ let rec poly_pow p k = else let q = poly_square(poly_pow p (k / 2)) in if k mod 2 = 1 then poly_mul p q else q;; -let poly_exp p1 p2 = - if not(poly_isconst p2) then failwith "poly_exp: not a constant" else - poly_pow p1 (Num.int_of_num (eval undefined p2));; - let degree x (p:poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p;; let multidegree (p:poly) = @@ -297,42 +233,8 @@ let humanorder_monomial = (* Conversions to strings. *) (* ------------------------------------------------------------------------- *) -let string_of_vector min_size max_size (v:vector) = - let n_raw = dim v in - if n_raw = 0 then "[]" else - let n = max min_size (min n_raw max_size) in - let xs = List.map ((o) string_of_num (element v)) (1--n) in - "[" ^ end_itlist (fun s t -> s ^ ", " ^ t) xs ^ - (if n_raw > max_size then ", ...]" else "]");; - -let string_of_matrix max_size (m:matrix) = - let i_raw,j_raw = dimensions m in - let i = min max_size i_raw and j = min max_size j_raw in - let rstr = List.map (fun k -> string_of_vector j j (row k m)) (1--i) in - "["^end_itlist(fun s t -> s^";\n "^t) rstr ^ - (if j > max_size then "\n ...]" else "]");; - let string_of_vname (v:vname): string = (v: string);; -let rec string_of_term t = - match t with - Opp t1 -> "(- " ^ string_of_term t1 ^ ")" -| Add (t1, t2) -> - "(" ^ (string_of_term t1) ^ " + " ^ (string_of_term t2) ^ ")" -| Sub (t1, t2) -> - "(" ^ (string_of_term t1) ^ " - " ^ (string_of_term t2) ^ ")" -| Mul (t1, t2) -> - "(" ^ (string_of_term t1) ^ " * " ^ (string_of_term t2) ^ ")" -| Inv t1 -> "(/ " ^ string_of_term t1 ^ ")" -| Div (t1, t2) -> - "(" ^ (string_of_term t1) ^ " / " ^ (string_of_term t2) ^ ")" -| Pow (t1, n1) -> - "(" ^ (string_of_term t1) ^ " ^ " ^ (string_of_int n1) ^ ")" -| Zero -> "0" -| Var v -> "x" ^ (string_of_vname v) -| Const x -> string_of_num x;; - - let string_of_varpow x k = if k = 1 then string_of_vname x else string_of_vname x^"^"^string_of_int k;; @@ -363,6 +265,7 @@ let string_of_poly (p:poly) = (* Printers. *) (* ------------------------------------------------------------------------- *) +(* let print_vector v = Format.print_string(string_of_vector 0 20 v);; let print_matrix m = Format.print_string(string_of_matrix 20 m);; @@ -371,7 +274,6 @@ let print_monomial m = Format.print_string(string_of_monomial m);; let print_poly m = Format.print_string(string_of_poly m);; -(* #install_printer print_vector;; #install_printer print_matrix;; #install_printer print_monomial;; @@ -410,19 +312,6 @@ let sdpa_of_vector (v:vector) = let strs = List.map (o (decimalize 20) (element v)) (1--n) in end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";; -(* ------------------------------------------------------------------------- *) -(* String for block diagonal matrix numbered k. *) -(* ------------------------------------------------------------------------- *) - -let sdpa_of_blockdiagonal k m = - let pfx = string_of_int k ^" " in - let ents = - foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in - let entss = sort (increasing fst) ents in - itlist (fun ((b,i,j),c) a -> - pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ - " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; - (* ------------------------------------------------------------------------- *) (* String for a matrix numbered k, in SDPA sparse format. *) (* ------------------------------------------------------------------------- *) @@ -486,13 +375,11 @@ let mkparser p s = let x,rst = p(explode s) in if rst = [] then x else failwith "mkparser: unparsed input";; -let parse_decimal = mkparser decimal;; - (* ------------------------------------------------------------------------- *) (* Parse back a vector. *) (* ------------------------------------------------------------------------- *) -let parse_sdpaoutput,parse_csdpoutput = +let _parse_sdpaoutput, parse_csdpoutput = let (||) = parser_or in let vector = token "{" ++ listof decimal (token ",") "decimal" ++ token "}" @@ -509,25 +396,11 @@ let parse_sdpaoutput,parse_csdpoutput = (a " " ++ a "\n" ++ ignore) >> ((o) vector_of_list fst) in mkparser sdpaoutput,mkparser csdpoutput;; -(* ------------------------------------------------------------------------- *) -(* Also parse the SDPA output to test success (CSDP yields a return code). *) -(* ------------------------------------------------------------------------- *) - -let sdpa_run_succeeded = - let (||) = parser_or in - let rec skipupto dscr prs inp = - (dscr ++ prs >> snd - || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in - let prs = skipupto (word "phase.value" ++ token "=") - (possibly (a "p") ++ possibly (a "d") ++ - (word "OPT" || word "FEAS")) in - fun s -> try ignore (prs (explode s)); true with Noparse -> false;; - (* ------------------------------------------------------------------------- *) (* The default parameters. Unfortunately this goes to a fixed file. *) (* ------------------------------------------------------------------------- *) -let sdpa_default_parameters = +let _sdpa_default_parameters = "100 unsigned int maxIteration;\ \n1.0E-7 double 0.0 < epsilonStar;\ \n1.0E2 double 0.0 < lambdaStar;\ @@ -558,7 +431,7 @@ let sdpa_alt_parameters = \n1.0E-7 double 0.0 < epsilonDash;\ \n";; -let sdpa_params = sdpa_alt_parameters;; +let _sdpa_params = sdpa_alt_parameters;; (* ------------------------------------------------------------------------- *) (* CSDP parameters; so far I'm sticking with the defaults. *) @@ -591,10 +464,10 @@ let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in @@ -603,16 +476,6 @@ let run_csdp dbg obj mats = else (Sys.remove input_file; Sys.remove output_file)); rv,res);; -let csdp obj mats = - let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" - else if rv = 3 then () - (* Format.print_string "csdp warning: Reduced accuracy"; - Format.print_newline() *) - else if rv <> 0 then failwith("csdp: error "^string_of_int rv) - else ()); - res;; - (* ------------------------------------------------------------------------- *) (* Try some apparently sensible scaling first. Note that this is purely to *) (* get a cleaner translation to floating-point, and doesn't affect any of *) @@ -660,20 +523,6 @@ let linear_program_basic a = else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; -(* ------------------------------------------------------------------------- *) -(* Alternative interface testing A x >= b for matrix A, vector b. *) -(* ------------------------------------------------------------------------- *) - -let linear_program a b = - let m,n = dimensions a in - if dim b <> m then failwith "linear_program: incompatible dimensions" else - let mats = diagonal b :: List.map (fun j -> diagonal (column j a)) (1--n) - and obj = vector_const (Int 1) m in - let rv,res = run_csdp false obj mats in - if rv = 1 || rv = 2 then false - else if rv = 0 then true - else failwith "linear_program: An error occurred in the SDP solver";; - (* ------------------------------------------------------------------------- *) (* Test whether a point is in the convex hull of others. Rather than use *) (* computational geometry, express as linear inequalities and call CSDP. *) @@ -718,40 +567,6 @@ let equation_eval assig eq = let value v = apply assig v in foldl (fun a v c -> a +/ value(v) */ c) (Int 0) eq;; -(* ------------------------------------------------------------------------- *) -(* Eliminate among linear equations: return unconstrained variables and *) -(* assignments for the others in terms of them. We give one pseudo-variable *) -(* "one" that's used for a constant term. *) -(* ------------------------------------------------------------------------- *) - -let failstore = ref [];; - -let eliminate_equations = - let rec extract_first p l = - match l with - [] -> failwith "extract_first" - | h::t -> if p(h) then h,t else - let k,s = extract_first p t in - k,h::s in - let rec eliminate vars dun eqs = - match vars with - [] -> if forall is_undefined eqs then dun - else (failstore := [vars,dun,eqs]; raise Unsolvable) - | v::vs -> - try let eq,oeqs = extract_first (fun e -> defined e v) eqs in - let a = apply eq v in - let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in - let elim e = - let b = tryapplyd e v (Int 0) in - if b =/ Int 0 then e else - equation_add e (equation_cmul (minus_num b // a) eq) in - eliminate vs ((v |-> eq') (mapf elim dun)) (List.map elim oeqs) - with Failure _ -> eliminate vs dun eqs in - fun one vars eqs -> - let assig = eliminate vars undefined eqs in - let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in - setify vs,assig;; - (* ------------------------------------------------------------------------- *) (* Eliminate all variables, in an essentially arbitrary order. *) (* ------------------------------------------------------------------------- *) @@ -782,18 +597,6 @@ let eliminate_all_equations one = let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in setify vs,assig;; -(* ------------------------------------------------------------------------- *) -(* Solve equations by assigning arbitrary numbers. *) -(* ------------------------------------------------------------------------- *) - -let solve_equations one eqs = - let vars,assigs = eliminate_all_equations one eqs in - let vfn = itlist (fun v -> (v |-> Int 0)) vars (one |=> Int(-1)) in - let ass = - combine (+/) (fun c -> false) (mapf (equation_eval vfn) assigs) vfn in - if forall (fun e -> equation_eval ass e =/ Int 0) eqs - then undefine one ass else raise Sanity;; - (* ------------------------------------------------------------------------- *) (* Hence produce the "relevant" monomials: those whose squares lie in the *) (* Newton polytope of the monomials in the input. (This is enough according *) @@ -900,19 +703,6 @@ let epoly_pmul p q acc = (m |-> equation_add (equation_cmul c e) es) b) a q) acc p;; -(* ------------------------------------------------------------------------- *) -(* Usual operations on equation-parametrized poly. *) -(* ------------------------------------------------------------------------- *) - -let epoly_cmul c l = - if c =/ Int 0 then undefined else mapf (equation_cmul c) l;; - -let epoly_neg = epoly_cmul (Int(-1));; - -let epoly_add = combine equation_add is_undefined;; - -let epoly_sub p q = epoly_add p (epoly_neg q);; - (* ------------------------------------------------------------------------- *) (* Convert regular polynomial. Note that we treat (0,0,0) as -1. *) (* ------------------------------------------------------------------------- *) @@ -956,11 +746,11 @@ let run_csdp dbg nblocks blocksizes obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_blockproblem "" nblocks blocksizes obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in @@ -991,8 +781,6 @@ let bmatrix_cmul c bm = let bmatrix_neg = bmatrix_cmul (Int(-1));; -let bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);; - (* ------------------------------------------------------------------------- *) (* Smash a block matrix into components. *) (* ------------------------------------------------------------------------- *) @@ -1104,15 +892,6 @@ let real_positivnullstellensatz_general linf d eqs leqs pol = if not(is_undefined sanity) then raise Sanity else cfs,List.map (fun (a,b) -> snd a,b) msq;; -(* ------------------------------------------------------------------------- *) -(* Iterative deepening. *) -(* ------------------------------------------------------------------------- *) - -let rec deepen f n = - try print_string "Searching with depth limit "; - print_int n; print_newline(); f n - with Failure _ -> deepen f (n + 1);; - (* ------------------------------------------------------------------------- *) (* The ordering so we can create canonical HOL polynomials. *) (* ------------------------------------------------------------------------- *) @@ -1139,10 +918,6 @@ let monomial_order = if deg1 < deg2 then false else if deg1 > deg2 then true else lexorder mon1 mon2;; -let dest_poly p = - List.map (fun (m,c) -> c,dest_monomial m) - (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p));; - (* ------------------------------------------------------------------------- *) (* Map back polynomials and their composites to HOL. *) (* ------------------------------------------------------------------------- *) @@ -1376,9 +1151,6 @@ let rec allpermutations l = itlist (fun h acc -> List.map (fun t -> h::t) (allpermutations (subtract l [h])) @ acc) l [];; -let allvarorders l = - List.map (fun vlis x -> index x vlis) (allpermutations l);; - let changevariables_monomial zoln (m:monomial) = foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m;; @@ -1395,15 +1167,6 @@ let sdpa_of_vector (v:vector) = let strs = List.map (o (decimalize 20) (element v)) (1--n) in end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";; -let sdpa_of_blockdiagonal k m = - let pfx = string_of_int k ^" " in - let ents = - foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in - let entss = sort (increasing fst) ents in - itlist (fun ((b,i,j),c) a -> - pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ - " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; - let sdpa_of_matrix k (m:matrix) = let pfx = string_of_int k ^ " 1 " in let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a) @@ -1428,10 +1191,10 @@ let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml index 9ee3db6e0..6b8b820ac 100644 --- a/plugins/micromega/sos_lib.ml +++ b/plugins/micromega/sos_lib.ml @@ -571,7 +571,7 @@ let finished input = (* ------------------------------------------------------------------------- *) -let temp_path = ref Filename.temp_dir_name;; +let temp_path = Filename.get_temp_dir_name ();; (* ------------------------------------------------------------------------- *) (* Convenient conversion between files and (lists of) strings. *) -- cgit v1.2.3 From 4e84e83911c1cf7613a35b921b1e68e097f84b5a Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:11:47 +0200 Subject: Remove unused [open] statements --- dev/top_printers.ml | 1 - engine/eConstr.ml | 1 - engine/eConstr.mli | 1 - engine/evarutil.ml | 1 - engine/evd.ml | 4 ---- engine/proofview.mli | 1 - engine/termops.ml | 1 - engine/universes.ml | 1 - interp/constrintern.ml | 2 -- interp/constrintern.mli | 1 - interp/stdarg.ml | 2 -- interp/stdarg.mli | 1 - interp/syntax_def.ml | 2 -- intf/tactypes.mli | 1 - library/goptions.ml | 1 - library/libobject.ml | 1 - parsing/egramcoq.ml | 1 - parsing/egramcoq.mli | 8 -------- parsing/g_prim.ml4 | 1 - parsing/g_proofs.ml4 | 1 - plugins/firstorder/instances.ml | 1 - plugins/funind/functional_principles_proofs.mli | 1 - plugins/funind/g_indfun.ml4 | 1 - plugins/funind/indfun_common.ml | 1 - plugins/funind/invfun.ml | 1 - plugins/funind/merge.ml | 1 - plugins/ltac/evar_tactics.ml | 1 - plugins/ltac/extratactics.ml4 | 1 - plugins/ltac/g_auto.ml4 | 1 - plugins/ltac/g_class.ml4 | 3 --- plugins/ltac/pltac.ml | 1 - plugins/ltac/rewrite.mli | 1 - plugins/ltac/taccoerce.mli | 1 - plugins/ltac/tacentries.ml | 1 - plugins/ltac/tacenv.mli | 1 - plugins/ltac/tacinterp.ml | 1 - plugins/ltac/tacinterp.mli | 1 - plugins/ltac/tactic_debug.mli | 1 - plugins/micromega/coq_micromega.ml | 1 - plugins/nsatz/ideal.ml | 2 -- plugins/setoid_ring/newring.ml | 1 - plugins/setoid_ring/newring.mli | 3 --- plugins/ssrmatching/ssrmatching.ml4 | 14 -------------- plugins/ssrmatching/ssrmatching.mli | 1 - pretyping/cbv.mli | 1 - pretyping/classops.ml | 1 - pretyping/classops.mli | 1 - pretyping/coercion.ml | 1 - pretyping/coercion.mli | 1 - pretyping/detyping.ml | 1 - pretyping/evarconv.ml | 1 - pretyping/evarconv.mli | 1 - pretyping/evardefine.ml | 1 - pretyping/find_subterm.mli | 1 - pretyping/inductiveops.ml | 1 - pretyping/patternops.ml | 1 - pretyping/patternops.mli | 1 - pretyping/pretype_errors.ml | 1 - pretyping/pretyping.ml | 1 - pretyping/program.ml | 1 - pretyping/tacred.mli | 1 - pretyping/typeclasses_errors.ml | 1 - pretyping/typeclasses_errors.mli | 1 - printing/prettyp.mli | 1 - proofs/clenvtac.mli | 1 - proofs/goal.ml | 1 - proofs/proof_type.mli | 3 --- proofs/refiner.ml | 1 - proofs/tacmach.mli | 1 - stm/stm.mli | 3 --- tactics/auto.ml | 1 - tactics/auto.mli | 1 - tactics/btermdn.mli | 1 - tactics/class_tactics.ml | 3 --- tactics/class_tactics.mli | 1 - tactics/contradiction.mli | 1 - tactics/eauto.mli | 2 -- tactics/elim.ml | 1 - tactics/elim.mli | 1 - tactics/eqdecide.ml | 1 - tactics/equality.ml | 1 - tactics/equality.mli | 1 - tactics/hints.mli | 1 - tactics/hipattern.mli | 1 - tactics/inv.ml | 1 - tactics/inv.mli | 1 - tactics/leminv.mli | 1 - tactics/tactics.ml | 2 -- toplevel/coqloop.mli | 2 -- vernac/auto_ind_decl.ml | 2 -- vernac/classes.ml | 2 -- vernac/obligations.mli | 1 - vernac/search.ml | 2 -- vernac/search.mli | 1 - 94 files changed, 137 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 918e98f77..7caaf2d9d 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -503,7 +503,6 @@ VERNAC COMMAND EXTEND PrintConstr END *) -open Pcoq open Genarg open Stdarg open Egramml diff --git a/engine/eConstr.ml b/engine/eConstr.ml index bb9075e74..54d3ce6cf 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open CSig open CErrors open Util open Names diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 3a9469e55..693b592fd 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -9,7 +9,6 @@ open CSig open Names open Constr -open Context open Environ type t diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 1624dc93e..fba466107 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -10,7 +10,6 @@ open CErrors open Util open Names open Term -open Vars open Termops open Namegen open Pre_env diff --git a/engine/evd.ml b/engine/evd.ml index 5419a10a8..6b1e1a855 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -14,8 +14,6 @@ open Nameops open Term open Vars open Environ -open Globnames -open Context.Named.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration @@ -360,8 +358,6 @@ module EvMap = Evar.Map module EvNames : sig -open Misctypes - type t val empty : t diff --git a/engine/proofview.mli b/engine/proofview.mli index a3b0304b1..da8a8fecd 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -13,7 +13,6 @@ state and returning a value of type ['a]. *) open Util -open Term open EConstr (** Main state of tactics *) diff --git a/engine/termops.ml b/engine/termops.ml index 29dcddbb0..19e62f8e6 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1427,7 +1427,6 @@ let dependency_closure env sigma sign hyps = List.rev lh let global_app_of_constr sigma c = - let open Univ in let open Globnames in match EConstr.kind sigma c with | Const (c, u) -> (ConstRef c, u), None diff --git a/engine/universes.ml b/engine/universes.ml index ab561784c..1900112dd 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -13,7 +13,6 @@ open Term open Environ open Univ open Globnames -open Decl_kinds let pr_with_global_universes l = try Nameops.pr_id (LMap.find l (snd (Global.global_universe_names ()))) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d75487ecf..389880c5c 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -2045,8 +2045,6 @@ let interp_binder_evars env evdref na t = let t' = locate_if_hole (loc_of_glob_constr t) na t in understand_tcc_evars env evdref ~expected_type:IsType t' -open Environ - let my_intern_constr env lvar acc c = internalize env acc false lvar c diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 758d4e650..fdd50c6a1 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -18,7 +18,6 @@ open Constrexpr open Notation_term open Pretyping open Misctypes -open Decl_kinds (** Translation from front abstract syntax of term to untyped terms (glob_constr) *) diff --git a/interp/stdarg.ml b/interp/stdarg.ml index 341ff5662..5920b0d50 100644 --- a/interp/stdarg.ml +++ b/interp/stdarg.ml @@ -6,9 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Misctypes -open Tactypes open Genarg open Geninterp diff --git a/interp/stdarg.mli b/interp/stdarg.mli index 113fe40ba..ac40a2328 100644 --- a/interp/stdarg.mli +++ b/interp/stdarg.mli @@ -10,7 +10,6 @@ open Loc open Names -open Term open EConstr open Libnames open Globnames diff --git a/interp/syntax_def.ml b/interp/syntax_def.ml index c3f4c4f30..ed7b0b70d 100644 --- a/interp/syntax_def.ml +++ b/interp/syntax_def.ml @@ -106,5 +106,3 @@ let search_syntactic_definition kn = let def = out_pat pat in verbose_compat kn def v; def - -open Goptions diff --git a/intf/tactypes.mli b/intf/tactypes.mli index 02cfc44e2..ef90b911c 100644 --- a/intf/tactypes.mli +++ b/intf/tactypes.mli @@ -13,7 +13,6 @@ open Loc open Names open Constrexpr -open Glob_term open Pattern open Misctypes diff --git a/library/goptions.ml b/library/goptions.ml index 1c08b9539..c111113ca 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -235,7 +235,6 @@ with Not_found -> then error "Sorry, this option name is already used." open Libobject -open Lib let warn_deprecated_option = CWarnings.create ~name:"deprecated-option" ~category:"deprecated" diff --git a/library/libobject.ml b/library/libobject.ml index 8757ca08c..897591762 100644 --- a/library/libobject.ml +++ b/library/libobject.ml @@ -8,7 +8,6 @@ open Libnames open Pp -open Util module Dyn = Dyn.Make(struct end) diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 65e99d1b9..86c66ec5f 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -10,7 +10,6 @@ open CErrors open Util open Pcoq open Constrexpr -open Notation open Notation_term open Extend open Libnames diff --git a/parsing/egramcoq.mli b/parsing/egramcoq.mli index 6dda3817a..0a0430ba6 100644 --- a/parsing/egramcoq.mli +++ b/parsing/egramcoq.mli @@ -6,14 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Names -open Constrexpr -open Notation_term -open Pcoq -open Extend -open Genarg -open Egramml - (** Mapping of grammar productions to camlp4 actions *) (** This is the part specific to Coq-level Notation and Tactic Notation. diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index 2af4ed533..abb463f82 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -8,7 +8,6 @@ open Names open Libnames -open Tok (* necessary for camlp4 *) open Pcoq open Pcoq.Prim diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 7ca2e4a4f..68b8be6b8 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -9,7 +9,6 @@ open Constrexpr open Vernacexpr open Misctypes -open Tok open Pcoq open Pcoq.Prim diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 62f811546..2b624b983 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -17,7 +17,6 @@ open Tacmach.New open Tactics open Tacticals.New open Proofview.Notations -open Termops open Reductionops open Formula open Sequent diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli index 7ddc84d01..61752aa33 100644 --- a/plugins/funind/functional_principles_proofs.mli +++ b/plugins/funind/functional_principles_proofs.mli @@ -1,5 +1,4 @@ open Names -open Term val prove_princ_for_struct : Evd.evar_map ref -> diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 0dccd25d7..b5eacee81 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -8,7 +8,6 @@ (*i camlp4deps: "grammar/grammar.cma" i*) open Ltac_plugin open Util -open Term open Pp open Constrexpr open Indfun_common diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 7b0d7d27d..c6f5c2745 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -508,7 +508,6 @@ let list_rewrite (rev:bool) (eqs: (EConstr.constr*bool) list) = (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));; let decompose_lam_n sigma n = - let open EConstr in if n < 0 then CErrors.error "decompose_lam_n: integer parameter must be positive"; let rec lamdec_rec l n c = if Int.equal n 0 then l,c diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 94ec0a898..29472cdef 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -7,7 +7,6 @@ (************************************************************************) open Ltac_plugin -open Tacexpr open Declarations open CErrors open Util diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index f1ca57585..0af0898a0 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -19,7 +19,6 @@ open Pp open Names open Term open Vars -open Termops open Declarations open Glob_term open Glob_termops diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5d3f6df03..bc9c300e2 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -9,7 +9,6 @@ open Util open Names open Term -open EConstr open CErrors open Evar_refiner open Tacmach diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 35cfe8b54..21419d1f9 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -21,7 +21,6 @@ open Tacexpr open Glob_ops open CErrors open Util -open Evd open Termops open Equality open Misctypes diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4 index dfa8331ff..50e8255a6 100644 --- a/plugins/ltac/g_auto.ml4 +++ b/plugins/ltac/g_auto.ml4 @@ -16,7 +16,6 @@ open Pcoq.Constr open Pltac open Hints open Tacexpr -open Proofview.Notations open Names DECLARE PLUGIN "g_auto" diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4 index ff5e7d5ff..23ce368ee 100644 --- a/plugins/ltac/g_class.ml4 +++ b/plugins/ltac/g_class.ml4 @@ -8,9 +8,7 @@ (*i camlp4deps: "grammar/grammar.cma" i*) -open Misctypes open Class_tactics -open Pltac open Stdarg open Tacarg open Names @@ -95,7 +93,6 @@ END (** TODO: DEPRECATE *) (* A progress test that allows to see if the evars have changed *) open Term -open Proofview.Goal open Proofview.Notations let rec eq_constr_mod_evars sigma x y = diff --git a/plugins/ltac/pltac.ml b/plugins/ltac/pltac.ml index 1d21118ae..7e979d269 100644 --- a/plugins/ltac/pltac.ml +++ b/plugins/ltac/pltac.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Names open Pcoq (* Main entry for extensions *) diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli index 7a20838a2..6683d753b 100644 --- a/plugins/ltac/rewrite.mli +++ b/plugins/ltac/rewrite.mli @@ -14,7 +14,6 @@ open Constrexpr open Tacexpr open Misctypes open Evd -open Proof_type open Tacinterp (** TODO: document and clean me! *) diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli index 9c4ac5265..4a44f86d9 100644 --- a/plugins/ltac/taccoerce.mli +++ b/plugins/ltac/taccoerce.mli @@ -8,7 +8,6 @@ open Util open Names -open Term open EConstr open Misctypes open Pattern diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index ef1d69d35..32750383b 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -15,7 +15,6 @@ open Genarg open Extend open Pcoq open Egramml -open Egramcoq open Vernacexpr open Libnames open Nameops diff --git a/plugins/ltac/tacenv.mli b/plugins/ltac/tacenv.mli index 94e14223a..d1e2a7bbe 100644 --- a/plugins/ltac/tacenv.mli +++ b/plugins/ltac/tacenv.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Genarg open Names open Tacexpr open Geninterp diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index fcdf7bb2c..b8c021f18 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -25,7 +25,6 @@ open Refiner open Tacmach.New open Tactic_debug open Constrexpr -open Term open Termops open Tacexpr open Genarg diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli index 1e5f6bd42..494f36a95 100644 --- a/plugins/ltac/tacinterp.mli +++ b/plugins/ltac/tacinterp.mli @@ -8,7 +8,6 @@ open Names open Tactic_debug -open Term open EConstr open Tacexpr open Genarg diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli index 7745d9b7b..0b4d35a22 100644 --- a/plugins/ltac/tactic_debug.mli +++ b/plugins/ltac/tactic_debug.mli @@ -10,7 +10,6 @@ open Environ open Pattern open Names open Tacexpr -open Term open EConstr open Evd diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index eb26c5431..a36607ec3 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -331,7 +331,6 @@ module M = struct open Coqlib - open Term open Constr open EConstr diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml index 41f2edfcf..a120d4efb 100644 --- a/plugins/nsatz/ideal.ml +++ b/plugins/nsatz/ideal.ml @@ -196,8 +196,6 @@ module Hashpol = Hashtbl.Make( (* A pretty printer for polynomials, with Maple-like syntax. *) -open Format - let getvar lv i = try (List.nth lv i) with Failure _ -> (List.fold_left (fun r x -> r^" "^x) "lv= " lv) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 6e072e831..d59ffee54 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -19,7 +19,6 @@ open Environ open Libnames open Globnames open Glob_term -open Tacticals open Tacexpr open Coqlib open Mod_subst diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli index 4367d021c..d9d32c681 100644 --- a/plugins/setoid_ring/newring.mli +++ b/plugins/setoid_ring/newring.mli @@ -7,13 +7,10 @@ (************************************************************************) open Names -open Constr open EConstr open Libnames open Globnames open Constrexpr -open Tacexpr -open Proof_type open Newring_ast val protect_tac_in : string -> Id.t -> unit Proofview.tactic diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f146fbb11..72c70750c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -21,30 +21,21 @@ open Pp open Pcoq open Genarg open Stdarg -open Tacarg open Term open Vars -open Topconstr open Libnames open Tactics open Tacticals open Termops -open Namegen open Recordops open Tacmach -open Coqlib open Glob_term open Util open Evd -open Extend -open Goptions open Tacexpr -open Proofview.Notations open Tacinterp open Pretyping open Constr -open Pltac -open Extraargs open Ppconstr open Printer @@ -54,14 +45,9 @@ open Decl_kinds open Evar_kinds open Constrexpr open Constrexpr_ops -open Notation_term -open Notation_ops -open Locus -open Locusops DECLARE PLUGIN "ssrmatching_plugin" -type loc = Loc.t let dummy_loc = Loc.ghost let errorstrm = CErrors.user_err ~hdr:"ssrmatching" let loc_error loc msg = CErrors.user_err ~loc ~hdr:msg (str msg) diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 1f984b160..638b4e254 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -4,7 +4,6 @@ open Genarg open Tacexpr open Environ -open Tacmach open Evd open Proof_type open Term diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index b014af2c7..eb25994be 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Environ open CClosure diff --git a/pretyping/classops.ml b/pretyping/classops.ml index e9b3d197b..32da81f96 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -17,7 +17,6 @@ open Nametab open Environ open Libobject open Term -open Termops open Mod_subst (* usage qque peu general: utilise aussi dans record *) diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 0d741a5a5..c4238e8b0 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Environ open EConstr open Evd diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 542db7fdf..c26e7458e 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -22,7 +22,6 @@ open Environ open EConstr open Vars open Reductionops -open Typeops open Pretype_errors open Classops open Evarutil diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index bc63d092d..ea3d3f0fa 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -8,7 +8,6 @@ open Evd open Names -open Term open Environ open EConstr open Glob_term diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 8ba408679..483e2b432 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -13,7 +13,6 @@ open CErrors open Util open Names open Term -open Environ open EConstr open Vars open Inductiveops diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4bb66b8e9..305eae15a 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -21,7 +21,6 @@ open Recordops open Evarutil open Evardefine open Evarsolve -open Globnames open Evd open Pretype_errors open Sigma.Notations diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli index fc07f0fbe..7cee1e8a7 100644 --- a/pretyping/evarconv.mli +++ b/pretyping/evarconv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Environ open Reductionops diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index c5ae684e3..5fd104c78 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -11,7 +11,6 @@ open Pp open Names open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index e3d3b74f1..d22f94e4e 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -7,7 +7,6 @@ (************************************************************************) open Locus -open Term open Evd open Pretype_errors open Environ diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 5b42add28..429e5005e 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -459,7 +459,6 @@ let extract_mrectype sigma t = | _ -> raise Not_found let find_mrectype_vect env sigma c = - let open EConstr in let (t, l) = Termops.decompose_app_vect sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind ind -> (ind, l) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index b16d04495..33a68589c 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -20,7 +20,6 @@ open Mod_subst open Misctypes open Decl_kinds open Pattern -open Evd open Environ let case_info_pattern_eq i1 i2 = diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli index 5694d345c..791fd74ed 100644 --- a/pretyping/patternops.mli +++ b/pretyping/patternops.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr open Globnames open Glob_term diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 24f6d1689..f9cf6b83b 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Util open Names open Term open Environ diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 497bde340..68ef97659 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -33,7 +33,6 @@ open EConstr open Vars open Reductionops open Type_errors -open Typeops open Typing open Globnames open Nameops diff --git a/pretyping/program.ml b/pretyping/program.ml index caa5a5c8a..42acc5705 100644 --- a/pretyping/program.ml +++ b/pretyping/program.ml @@ -10,7 +10,6 @@ open Pp open CErrors open Util open Names -open Term let make_dir l = DirPath.make (List.rev_map Id.of_string l) diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 76d0bc241..c31212e26 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Environ open Evd open EConstr diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml index 2db0e9e88..754dacd19 100644 --- a/pretyping/typeclasses_errors.ml +++ b/pretyping/typeclasses_errors.ml @@ -8,7 +8,6 @@ (*i*) open Names -open Term open EConstr open Environ open Constrexpr diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli index 9bd430e4d..558575ccc 100644 --- a/pretyping/typeclasses_errors.mli +++ b/pretyping/typeclasses_errors.mli @@ -8,7 +8,6 @@ open Loc open Names -open Term open EConstr open Environ open Constrexpr diff --git a/printing/prettyp.mli b/printing/prettyp.mli index 38e111034..6841781cc 100644 --- a/printing/prettyp.mli +++ b/printing/prettyp.mli @@ -8,7 +8,6 @@ open Pp open Names -open Term open Environ open Reductionops open Libnames diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli index 5b7164705..26069207e 100644 --- a/proofs/clenvtac.mli +++ b/proofs/clenvtac.mli @@ -8,7 +8,6 @@ (** Legacy components of the previous proof engine. *) -open Term open Clenv open EConstr open Unification diff --git a/proofs/goal.ml b/proofs/goal.ml index 9046f4534..fc8e635a0 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -8,7 +8,6 @@ open Util open Pp -open Term open Sigma.Notations module NamedDecl = Context.Named.Declaration diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli index 03bc5e471..e59db9e42 100644 --- a/proofs/proof_type.mli +++ b/proofs/proof_type.mli @@ -11,9 +11,6 @@ open Evd open Names open Term -open Glob_term -open Nametab -open Misctypes (** This module defines the structure of proof tree and the tactic type. So, it is used by [Proof_tree] and [Refiner] *) diff --git a/proofs/refiner.ml b/proofs/refiner.ml index bc12b3ba7..259e96a27 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -10,7 +10,6 @@ open Pp open CErrors open Util open Evd -open Environ open Proof_type open Logic diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 627a8e0e7..e6e60e27f 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -15,7 +15,6 @@ open Proof_type open Redexpr open Pattern open Locus -open Misctypes (** Operations for handling terms under a local typing context. *) diff --git a/stm/stm.mli b/stm/stm.mli index 30e9629c6..d2bee4496 100644 --- a/stm/stm.mli +++ b/stm/stm.mli @@ -6,10 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Vernacexpr open Names -open Feedback -open Loc (** state-transaction-machine interface *) diff --git a/tactics/auto.ml b/tactics/auto.ml index 42230dff1..324c551d0 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -12,7 +12,6 @@ open Pp open Util open CErrors open Names -open Vars open Termops open EConstr open Environ diff --git a/tactics/auto.mli b/tactics/auto.mli index 32710e347..9ed9f0ae2 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -9,7 +9,6 @@ (** This files implements auto and related automation tactics *) open Names -open Term open EConstr open Clenv open Pattern diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli index 2a5e7c345..27f624f71 100644 --- a/tactics/btermdn.mli +++ b/tactics/btermdn.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open Pattern open Names diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index c430e776a..6c724a1eb 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -18,7 +18,6 @@ open Names open Term open Termops open EConstr -open Reduction open Proof_type open Tacticals open Tacmach @@ -1219,7 +1218,6 @@ module Search = struct let intro_tac info kont gl = let open Proofview in - let open Proofview.Notations in let env = Goal.env gl in let sigma = Goal.sigma gl in let s = Sigma.to_evar_map sigma in @@ -1257,7 +1255,6 @@ module Search = struct let search_tac_gl ?st only_classes dep hints depth i sigma gls gl : unit Proofview.tactic = let open Proofview in - let open Proofview.Notations in if false (* In 8.6, still allow non-class goals only_classes && not (is_class_type sigma (Goal.concl gl)) *) then Tacticals.New.tclZEROMSG (str"Not a subgoal for a class") else diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 738cc0feb..4ab29f260 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -9,7 +9,6 @@ (** This files implements typeclasses eauto *) open Names -open Constr open EConstr open Tacmach diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli index 510b135b0..2cf5a6829 100644 --- a/tactics/contradiction.mli +++ b/tactics/contradiction.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr open Misctypes diff --git a/tactics/eauto.mli b/tactics/eauto.mli index e2006ac1e..c952f4e72 100644 --- a/tactics/eauto.mli +++ b/tactics/eauto.mli @@ -6,9 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr -open Proof_type open Hints open Tactypes diff --git a/tactics/elim.ml b/tactics/elim.ml index e37ec6bce..855cb206f 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -8,7 +8,6 @@ open Util open Names -open Term open Termops open EConstr open Inductiveops diff --git a/tactics/elim.mli b/tactics/elim.mli index dc1af79ba..fb7cc7b83 100644 --- a/tactics/elim.mli +++ b/tactics/elim.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Tacticals open Misctypes diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index bac3980d2..5012b0ef7 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -25,7 +25,6 @@ open Constr_matching open Misctypes open Tactypes open Hipattern -open Pretyping open Proofview.Notations open Tacmach.New open Coqlib diff --git a/tactics/equality.ml b/tactics/equality.ml index 25c28cf4a..cc7701ad5 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -14,7 +14,6 @@ open Names open Nameops open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/tactics/equality.mli b/tactics/equality.mli index 5467b4af2..d979c580a 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -8,7 +8,6 @@ (*i*) open Names -open Term open Evd open EConstr open Environ diff --git a/tactics/hints.mli b/tactics/hints.mli index 467fd46d5..3a0339ff5 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -9,7 +9,6 @@ open Pp open Util open Names -open Term open EConstr open Environ open Globnames diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index dd09c3a4d..82a3d47b5 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Evd open EConstr open Coqlib diff --git a/tactics/inv.ml b/tactics/inv.ml index 904a17417..266cac5c7 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -13,7 +13,6 @@ open Names open Nameops open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/tactics/inv.mli b/tactics/inv.mli index 446a62f6d..5835e763d 100644 --- a/tactics/inv.mli +++ b/tactics/inv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Misctypes open Tactypes diff --git a/tactics/leminv.mli b/tactics/leminv.mli index 26d4ac994..a343fc81a 100644 --- a/tactics/leminv.mli +++ b/tactics/leminv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Constrexpr open Misctypes diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 53f8e4d5f..f3f7d16b7 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5037,8 +5037,6 @@ end (** Tacticals defined directly in term of Proofview *) module New = struct - open Proofview.Notations - open Genredexpr open Locus diff --git a/toplevel/coqloop.mli b/toplevel/coqloop.mli index 66bbf43f6..13e860a88 100644 --- a/toplevel/coqloop.mli +++ b/toplevel/coqloop.mli @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp - (** The Coq toplevel loop. *) (** A buffer for the character read from a channel. We store the command diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 015552d68..25091f783 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -9,7 +9,6 @@ (* This file is about the automatic generation of schemes about decidable equality, created by Vincent Siles, Oct 2007 *) -open Tacmach open CErrors open Util open Pp @@ -716,7 +715,6 @@ let compute_lb_goal ind lnamesparrec nparrec = ))), eff let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = - let open EConstr in let list_id = list_id lnamesparrec in let avoid = ref [] in let first_intros = diff --git a/vernac/classes.ml b/vernac/classes.ml index 833719965..d515b0c9b 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -66,8 +66,6 @@ let _ = Hook.set Typeclasses.classes_transparent_state_hook (fun () -> Hints.Hint_db.transparent_state (Hints.searchtable_map typeclasses_db)) -open Vernacexpr - (** TODO: add subinstances *) let existing_instance glob g info = let c = global g in diff --git a/vernac/obligations.mli b/vernac/obligations.mli index 11366fe91..a276f9f9a 100644 --- a/vernac/obligations.mli +++ b/vernac/obligations.mli @@ -12,7 +12,6 @@ open Evd open Names open Pp open Globnames -open Vernacexpr open Decl_kinds (** Forward declaration. *) diff --git a/vernac/search.ml b/vernac/search.ml index 6279b17ae..5b6e9a9c3 100644 --- a/vernac/search.ml +++ b/vernac/search.ml @@ -14,11 +14,9 @@ open Declarations open Libobject open Environ open Pattern -open Printer open Libnames open Globnames open Nametab -open Goptions module NamedDecl = Context.Named.Declaration diff --git a/vernac/search.mli b/vernac/search.mli index 82b79f75d..e34522d8a 100644 --- a/vernac/search.mli +++ b/vernac/search.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp open Names open Term open Environ -- cgit v1.2.3 From b20d52da0d040fe37bb75b0b739ad7686f9af127 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Sat, 22 Apr 2017 12:55:46 +0200 Subject: Warning 29: non escaped end of line may be non portable --- ide/coqide_ui.ml | 284 ++++++++++++++++++++++++------------------------- ide/ide_slave.ml | 4 +- interp/constrintern.ml | 6 +- tools/coq_makefile.ml | 116 ++++++++++---------- toplevel/usage.ml | 2 +- 5 files changed, 206 insertions(+), 206 deletions(-) diff --git a/ide/coqide_ui.ml b/ide/coqide_ui.ml index 2ae18593a..91c281c8d 100644 --- a/ide/coqide_ui.ml +++ b/ide/coqide_ui.ml @@ -28,148 +28,148 @@ let list_queries menu li = res_buf let init () = - let theui = Printf.sprintf " - - - - - - - - - - - - - - - - - - - %s - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - %s - - - - - - - - - - - %s - - - - - - - - - - %s - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - -" + let theui = Printf.sprintf "\ +\n\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n %s\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n %s\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n %s\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n %s\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n\ +\n\ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n \ +\n\ +\n" (if Coq_config.gtk_platform <> `QUARTZ then "" else "") (Buffer.contents (list_items "Tactic" Coq_commands.tactics)) (Buffer.contents (list_items "Template" Coq_commands.commands)) diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index ca0ef38d3..dbfe4256b 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -512,5 +512,5 @@ let () = Coqtop.toploop_init := (fun args -> let () = Coqtop.toploop_run := loop let () = Usage.add_to_usage "coqidetop" -" --xml_format=Ppcmds serialize pretty printing messages using the std_ppcmds format - --help-XML-protocol print the documentation of the XML protocol used by CoqIDE\n" +" --xml_format=Ppcmds serialize pretty printing messages using the std_ppcmds format\ +\n --help-XML-protocol print the documentation of the XML protocol used by CoqIDE\n" diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 389880c5c..3f99a3c7c 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -855,9 +855,9 @@ let intern_qualid loc qid intern env lvar us args = | Some _, GApp (loc, GRef (loc', ref, None), arg) -> GApp (loc, GRef (loc', ref, us), arg) | Some _, _ -> - user_err ~loc (str "Notation " ++ pr_qualid qid ++ - str " cannot have a universe instance, its expanded head - does not start with a reference") + user_err ~loc (str "Notation " ++ pr_qualid qid + ++ str " cannot have a universe instance," + ++ str " its expanded head does not start with a reference") in c, projapp, args2 diff --git a/tools/coq_makefile.ml b/tools/coq_makefile.ml index ed89bda2c..4875cb62b 100644 --- a/tools/coq_makefile.ml +++ b/tools/coq_makefile.ml @@ -56,48 +56,48 @@ let lib_dirs = let usage () = - output_string stderr "Usage summary: - -coq_makefile .... [file.v] ... [file.ml[i4]?] ... [file.ml{lib,pack}] - ... [any] ... [-extra[-phony] result dependencies command] - ... [-I dir] ... [-R physicalpath logicalpath] - ... [-Q physicalpath logicalpath] ... [VARIABLE = value] - ... [-arg opt] ... [-opt|-byte] [-no-install] [-f file] [-o file] - [-h] [--help] - -[file.v]: Coq file to be compiled -[file.ml[i4]?]: Objective Caml file to be compiled -[file.ml{lib,pack}]: ocamlbuild file that describes a Objective Caml - library/module -[any] : subdirectory that should be \"made\" and has a Makefile itself - to do so. Very fragile and discouraged. -[-extra result dependencies command]: add target \"result\" with command - \"command\" and dependencies \"dependencies\". If \"result\" is not - generic (do not contains a %), \"result\" is built by _make all_ and - deleted by _make clean_. -[-extra-phony result dependencies command]: add a PHONY target \"result\" - with command \"command\" and dependencies \"dependencies\". Note that - _-extra-phony foo bar \"\"_ is a regular way to add the target \"bar\" as - as a dependencies of an already defined target \"foo\". -[-I dir]: look for Objective Caml dependencies in \"dir\" -[-R physicalpath logicalpath]: look for Coq dependencies resursively - starting from \"physicalpath\". The logical path associated to the - physical path is \"logicalpath\". -[-Q physicalpath logicalpath]: look for Coq dependencies starting from - \"physicalpath\". The logical path associated to the physical path - is \"logicalpath\". -[VARIABLE = value]: Add the variable definition \"VARIABLE=value\" -[-byte]: compile with byte-code version of coq -[-opt]: compile with native-code version of coq -[-arg opt]: send option \"opt\" to coqc -[-install opt]: where opt is \"user\" to force install into user directory, - \"none\" to build a makefile with no install target or - \"global\" to force install in $COQLIB directory -[-f file]: take the contents of file as arguments -[-o file]: output should go in file file - Output file outside the current directory is forbidden. -[-h]: print this usage summary -[--help]: equivalent to [-h]\n"; + output_string stderr "Usage summary:\ +\n\ +\ncoq_makefile .... [file.v] ... [file.ml[i4]?] ... [file.ml{lib,pack}]\ +\n ... [any] ... [-extra[-phony] result dependencies command]\ +\n ... [-I dir] ... [-R physicalpath logicalpath]\ +\n ... [-Q physicalpath logicalpath] ... [VARIABLE = value]\ +\n ... [-arg opt] ... [-opt|-byte] [-no-install] [-f file] [-o file]\ +\n [-h] [--help]\ +\n\ +\n[file.v]: Coq file to be compiled\ +\n[file.ml[i4]?]: Objective Caml file to be compiled\ +\n[file.ml{lib,pack}]: ocamlbuild file that describes a Objective Caml\ +\n library/module\ +\n[any] : subdirectory that should be \"made\" and has a Makefile itself\ +\n to do so. Very fragile and discouraged.\ +\n[-extra result dependencies command]: add target \"result\" with command\ +\n \"command\" and dependencies \"dependencies\". If \"result\" is not\ +\n generic (do not contains a %), \"result\" is built by _make all_ and\ +\n deleted by _make clean_.\ +\n[-extra-phony result dependencies command]: add a PHONY target \"result\"\ +\n with command \"command\" and dependencies \"dependencies\". Note that\ +\n _-extra-phony foo bar \"\"_ is a regular way to add the target \"bar\" as\ +\n as a dependencies of an already defined target \"foo\".\ +\n[-I dir]: look for Objective Caml dependencies in \"dir\"\ +\n[-R physicalpath logicalpath]: look for Coq dependencies resursively\ +\n starting from \"physicalpath\". The logical path associated to the\ +\n physical path is \"logicalpath\".\ +\n[-Q physicalpath logicalpath]: look for Coq dependencies starting from\ +\n \"physicalpath\". The logical path associated to the physical path\ +\n is \"logicalpath\".\ +\n[VARIABLE = value]: Add the variable definition \"VARIABLE=value\"\ +\n[-byte]: compile with byte-code version of coq\ +\n[-opt]: compile with native-code version of coq\ +\n[-arg opt]: send option \"opt\" to coqc\ +\n[-install opt]: where opt is \"user\" to force install into user directory,\ +\n \"none\" to build a makefile with no install target or\ +\n \"global\" to force install in $COQLIB directory\ +\n[-f file]: take the contents of file as arguments\ +\n[-o file]: output should go in file file\ +\n Output file outside the current directory is forbidden.\ +\n[-h]: print this usage summary\ +\n[--help]: equivalent to [-h]\n"; exit 1 let is_genrule r = (* generic rule (like bar%foo: ...) *) @@ -264,8 +264,8 @@ let where_put_doc = function then physical_dir_of_logical_dir pr else - let () = prerr_string "Warning: -Q options don't have a correct common prefix, - install-doc will put anything in $INSTALLDEFAULTROOT\n" in + let () = prerr_string ("Warning: -Q options don't have a correct common prefix," + ^ " install-doc will put anything in $INSTALLDEFAULTROOT\n") in "$(INSTALLDEFAULTROOT)" |_,inc_i,((_,lp,_)::q as inc_r) -> let pr = List.fold_left (fun a (_,b,_) -> string_prefix a b) lp q in @@ -277,8 +277,8 @@ let where_put_doc = function then physical_dir_of_logical_dir pr else - let () = prerr_string "Warning: -R/-Q options don't have a correct common prefix, - install-doc will put anything in $INSTALLDEFAULTROOT\n" in + let () = prerr_string ("Warning: -R/-Q options don't have a correct common prefix," + ^ " install-doc will put anything in $INSTALLDEFAULTROOT\n") in "$(INSTALLDEFAULTROOT)" let install (vfiles,(mlis,ml4s,mls,mllibs,mlpacks),_,sds) inc = function @@ -518,8 +518,8 @@ let variables is_install opt (args,defs) = if !some_ml4file || !some_mlfile || !some_mlifile then begin print "COQSRCLIBS?=" ; List.iter (fun c -> print "-I \"$(COQLIB)"; print c ; print "\" \\\n") lib_dirs ; - List.iter (fun c -> print " \\ - -I \"$(COQLIB)/"; print c; print "\"") Coq_config.plugins_dirs; print "\n"; + List.iter (fun c -> print " \\\ +\n -I \"$(COQLIB)/"; print c; print "\"") Coq_config.plugins_dirs; print "\n"; print "ZFLAGS=$(OCAMLLIBS) $(COQSRCLIBS) -I $(CAMLP4LIB)\n\n"; print "CAMLC?=$(OCAMLFIND) ocamlc -c -rectypes -thread -safe-string\n"; print "CAMLOPTC?=$(OCAMLFIND) opt -c -rectypes -thread -safe-string\n"; @@ -529,8 +529,8 @@ let variables is_install opt (args,defs) = print "CAMLLIB?=$(shell $(OCAMLFIND) printconf stdlib)\n"; print "GRAMMARS?=grammar.cma\n"; print "CAMLP4EXTEND=pa_extend.cmo q_MLast.cmo pa_macro.cmo\n"; - print "PP?=-pp '$(CAMLP4O) -I $(CAMLLIB) -I $(COQLIB)/grammar compat5.cmo \\ - $(CAMLP4EXTEND) $(GRAMMARS) $(CAMLP4OPTIONS) -impl'\n\n"; + print "PP?=-pp '$(CAMLP4O) -I $(CAMLLIB) -I $(COQLIB)/grammar compat5.cmo \\\ +\n $(CAMLP4EXTEND) $(GRAMMARS) $(CAMLP4OPTIONS) -impl'\n\n"; end; match is_install with | Project_file.NoInstall -> () @@ -816,14 +816,14 @@ let all_target (vfiles, (_,_,_,_,mlpackfiles as mlfiles), sps, sds) inc = let banner () = print (Printf.sprintf -"############################################################################# -## v # The Coq Proof Assistant ## -## Date: Tue, 25 Apr 2017 14:03:54 +0200 Subject: Remove uses of [Flags.make_silent] --- checker/checker.ml | 2 +- checker/include | 2 +- ide/ide_slave.ml | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/checker/checker.ml b/checker/checker.ml index 57e1806cf..5cadfe7b9 100644 --- a/checker/checker.ml +++ b/checker/checker.ml @@ -354,7 +354,7 @@ let parse_args argv = | "-norec" :: [] -> usage () | "-silent" :: rem -> - Flags.make_silent true; parse rem + Flags.quiet := true; parse rem | s :: _ when s<>"" && s.[0]='-' -> fatal_error (str "Unknown option " ++ str s) false diff --git a/checker/include b/checker/include index 6bea3c91a..09bf2826c 100644 --- a/checker/include +++ b/checker/include @@ -116,7 +116,7 @@ let prsub s = #install_printer prsub;;*) Checker.init_with_argv [|"";"-coqlib";"."|];; -Flags.make_silent false;; +Flags.quiet := false;; Flags.debug := true;; Sys.catch_break true;; diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index dbfe4256b..966a94da9 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -505,7 +505,7 @@ let rec parse = function let () = Coqtop.toploop_init := (fun args -> let args = parse args in - Flags.make_silent true; + Flags.quiet := true; CoqworkmgrApi.(init Flags.High); args) -- cgit v1.2.3 From 87910d7be9bd50de4db80f70c6e287c7c7994460 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Tue, 25 Apr 2017 14:31:15 +0200 Subject: Fix 4.04 warnings --- engine/evd.ml | 2 +- interp/notation.ml | 1 - kernel/constr.ml | 22 ---------------------- lib/cErrors.ml | 2 ++ parsing/pcoq.ml | 1 - plugins/firstorder/g_ground.ml4 | 1 - plugins/funind/functional_principles_proofs.ml | 5 +++-- plugins/funind/indfun_common.ml | 9 +++------ plugins/funind/invfun.ml | 10 +++++----- plugins/funind/recdef.ml | 5 +++-- plugins/ltac/rewrite.ml | 2 +- plugins/ltac/tauto.ml | 1 - plugins/omega/omega.ml | 10 ++++++---- pretyping/evarsolve.ml | 1 - pretyping/reductionops.ml | 6 +++++- pretyping/reductionops.mli | 5 ++++- pretyping/unification.ml | 2 +- printing/prettyp.ml | 2 +- stm/stm.ml | 3 +++ stm/vcs.ml | 2 -- tools/coqwc.mll | 2 ++ 21 files changed, 40 insertions(+), 54 deletions(-) diff --git a/engine/evd.ml b/engine/evd.ml index 6b1e1a855..db048bbd6 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -15,7 +15,7 @@ open Term open Vars open Environ -module RelDecl = Context.Rel.Declaration +(* module RelDecl = Context.Rel.Declaration *) module NamedDecl = Context.Named.Declaration (** Generic filters *) diff --git a/interp/notation.ml b/interp/notation.ml index 90ac7f729..6aa6f54c0 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -22,7 +22,6 @@ open Glob_ops open Ppextend open Context.Named.Declaration -module NamedDecl = Context.Named.Declaration (*i*) (*s A scope is a set of notations; it includes diff --git a/kernel/constr.ml b/kernel/constr.ml index 5a7561bf5..71efe9032 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -978,28 +978,6 @@ module Hcaseinfo = Hashcons.Make(CaseinfoHash) let case_info_hash = CaseinfoHash.hash -module Hsorts = - Hashcons.Make( - struct - open Sorts - - type t = Sorts.t - type u = universe -> universe - let hashcons huniv = function - Prop c -> Prop c - | Type u -> Type (huniv u) - let eq s1 s2 = - s1 == s2 || - match (s1,s2) with - (Prop c1, Prop c2) -> c1 == c2 - | (Type u1, Type u2) -> u1 == u2 - |_ -> false - let hash = function - | Prop Null -> 0 | Prop Pos -> 1 - | Type u -> 2 + Universe.hash u - end) - -(* let hcons_sorts = Hashcons.simple_hcons Hsorts.generate hcons_univ *) let hcons_caseinfo = Hashcons.simple_hcons Hcaseinfo.generate Hcaseinfo.hcons hcons_ind let hcons = diff --git a/lib/cErrors.ml b/lib/cErrors.ml index 1c1ff7e2f..b55fd80c6 100644 --- a/lib/cErrors.ml +++ b/lib/cErrors.ml @@ -121,12 +121,14 @@ end by inner functions during a [vernacinterp]. They should be handled only at the very end of interp, to be displayed to the user. *) +[@@@ocaml.warning "-52"] let noncritical = function | Sys.Break | Out_of_memory | Stack_overflow | Assert_failure _ | Match_failure _ | Anomaly _ | Timeout | Drop | Quit -> false | Invalid_argument "equal: functional value" -> false | _ -> true +[@@@ocaml.warning "+52"] (** Check whether an exception is handled *) diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml index 3f014c4c8..9a4766c0b 100644 --- a/parsing/pcoq.ml +++ b/parsing/pcoq.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp open CErrors open Util open Extend diff --git a/plugins/firstorder/g_ground.ml4 b/plugins/firstorder/g_ground.ml4 index 8ef6a09d0..b25017535 100644 --- a/plugins/firstorder/g_ground.ml4 +++ b/plugins/firstorder/g_ground.ml4 @@ -123,7 +123,6 @@ let normalize_evaluables= unfold_in_hyp (Lazy.force defined_connectives) (Tacexpr.InHypType id)) *) -open Pp open Genarg open Ppconstr open Printer diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index df81bc78c..55d361e3d 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -229,12 +229,13 @@ let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta Reductionops.clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty +exception NoChange let change_eq env sigma hyp_id (context:rel_context) x t end_of_type = let nochange ?t' msg = begin observe (str ("Not treating ( "^msg^" )") ++ pr_leconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr t ); - failwith "NoChange"; + raise NoChange; end in let eq_constr c1 c2 = Evarconv.e_conv env (ref sigma) c1 c2 in @@ -536,7 +537,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = tclTHEN tac (scan_type new_context new_t') - with Failure "NoChange" -> + with NoChange -> (* Last thing todo : push the rel in the context and continue *) scan_type (LocalAssum (x,t_x) :: context) t' end diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index c6f5c2745..848b44a60 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -21,12 +21,9 @@ let get_name avoid ?(default="H") = function | Name n -> Name n let array_get_start a = - try - Array.init - (Array.length a - 1) - (fun i -> a.(i)) - with Invalid_argument "index out of bounds" -> - invalid_arg "array_get_start" + Array.init + (Array.length a - 1) + (fun i -> a.(i)) let id_of_name = function Name id -> id diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 29472cdef..6c0c28905 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -1025,7 +1025,7 @@ let invfun qhyp f = | Not_found -> error "No graph found" | Option.IsNone -> error "Cannot use equivalence with graph!" - +exception NoFunction let invfun qhyp f g = match f with | Some f -> invfun qhyp f g @@ -1040,23 +1040,23 @@ let invfun qhyp f g = begin let f1,_ = decompose_app sigma args.(1) in try - if not (isConst sigma f1) then failwith ""; + if not (isConst sigma f1) then raise NoFunction; let finfos = find_Function_infos (fst (destConst sigma f1)) in let f_correct = mkConst(Option.get finfos.correctness_lemma) and kn = fst finfos.graph_ind in functional_inversion kn hid f1 f_correct g - with | Failure "" | Option.IsNone | Not_found -> + with | NoFunction | Option.IsNone | Not_found -> try let f2,_ = decompose_app sigma args.(2) in - if not (isConst sigma f2) then failwith ""; + if not (isConst sigma f2) then raise NoFunction; let finfos = find_Function_infos (fst (destConst sigma f2)) in let f_correct = mkConst(Option.get finfos.correctness_lemma) and kn = fst finfos.graph_ind in functional_inversion kn hid f2 f_correct g with - | Failure "" -> + | NoFunction -> user_err (str "Hypothesis " ++ Ppconstr.pr_id hid ++ str " must contain at least one Function") | Option.IsNone -> if do_observe () diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 1e405d2c9..bd30f1159 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1225,6 +1225,7 @@ let get_current_subgoals_types () = let { Evd.it=sgs ; sigma=sigma } = Proof.V82.subgoals p in sigma, List.map (Goal.V82.abstract_type sigma) sgs +exception EmptySubgoals let build_and_l sigma l = let and_constr = Coqlib.build_coq_and () in let conj_constr = coq_conj () in @@ -1246,7 +1247,7 @@ let build_and_l sigma l = in let l = List.sort compare l in let rec f = function - | [] -> failwith "empty list of subgoals!" + | [] -> raise EmptySubgoals | [p] -> p,tclIDTAC,1 | p1::pl -> let c,tac,nb = f pl in @@ -1432,7 +1433,7 @@ let com_terminate using_lemmas tcc_lemma_ref (Some tcc_lemma_name) (new_goal_type); - with Failure "empty list of subgoals!" -> + with EmptySubgoals -> (* a non recursive function declared with measure ! *) tcc_lemma_ref := Not_needed; defined () diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 12a1566e2..9a1615d3f 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -39,7 +39,7 @@ open Proofview.Notations open Context.Named.Declaration module NamedDecl = Context.Named.Declaration -module RelDecl = Context.Rel.Declaration +(* module RelDecl = Context.Rel.Declaration *) (** Typeclass-based generalized rewriting. *) diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml index e86d1c728..4de2081cf 100644 --- a/plugins/ltac/tauto.ml +++ b/plugins/ltac/tauto.ml @@ -10,7 +10,6 @@ open Term open EConstr open Hipattern open Names -open Pp open Geninterp open Misctypes open Tacexpr diff --git a/plugins/omega/omega.ml b/plugins/omega/omega.ml index bd991a955..334b03de1 100644 --- a/plugins/omega/omega.ml +++ b/plugins/omega/omega.ml @@ -330,11 +330,13 @@ let omega_mod a b = a - b * floor_div (two * a + b) (two * b) let banerjee_step (new_eq_id,new_var_id,print_var) original l1 l2 = let e = original.body in let sigma = new_var_id () in + if e == [] then begin + display_system print_var [original] ; failwith "TL" + end; let smallest,var = - try - List.fold_left (fun (v,p) c -> if v >? (abs c.c) then abs c.c,c.v else (v,p)) - (abs (List.hd e).c, (List.hd e).v) (List.tl e) - with Failure "tl" -> display_system print_var [original] ; failwith "TL" in + List.fold_left (fun (v,p) c -> if v >? (abs c.c) then abs c.c,c.v else (v,p)) + (abs (List.hd e).c, (List.hd e).v) (List.tl e) + in let m = smallest + one in let new_eq = { constant = omega_mod original.constant m; diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 77086d046..f0d011477 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -module CVars = Vars open Util open CErrors open Names diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index da3add2e5..52f424f75 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -239,6 +239,9 @@ sig | Shift of int | Update of 'a and 'a t = 'a member list + + exception IncompatibleFold2 + val pr : ('a -> Pp.std_ppcmds) -> 'a t -> Pp.std_ppcmds val empty : 'a t val is_empty : 'a t -> bool @@ -413,6 +416,7 @@ struct | (_,_) -> false in compare_rec 0 stk1 stk2 + exception IncompatibleFold2 let fold2 f o sk1 sk2 = let rec aux o lft1 sk1 lft2 sk2 = let fold_array = @@ -442,7 +446,7 @@ struct aux o lft1 (List.rev params1) lft2 (List.rev params2) in aux o' lft1' q1 lft2' q2 | (((Update _|App _|Case _|Proj _|Fix _| Cst _) :: _|[]), _) -> - raise (Invalid_argument "Reductionops.Stack.fold2") + raise IncompatibleFold2 in aux o 0 (List.rev sk1) 0 (List.rev sk2) let rec map f x = List.map (function diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 752c30a8a..af8048156 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -81,8 +81,11 @@ module Stack : sig val decomp_node_last : 'a app_node -> 'a t -> ('a * 'a t) val compare_shape : 'a t -> 'a t -> bool + + exception IncompatibleFold2 (** [fold2 f x sk1 sk2] folds [f] on any pair of term in [(sk1,sk2)]. - @return the result and the lifts to apply on the terms *) + @return the result and the lifts to apply on the terms + @raise IncompatibleFold2 when [sk1] and [sk2] have incompatible shapes *) val fold2 : ('a -> constr -> constr -> 'a) -> 'a -> constr t -> constr t -> 'a * int * int val map : ('a -> 'a) -> 'a t -> 'a t diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 4a3836d86..661c1d865 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1095,7 +1095,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e let app = mkApp (c, Array.rev_of_list ks) in (* let substn = unirec_rec curenvnb pb b false substn t cN in *) unirec_rec curenvnb pb opt' substn c1 app - with Invalid_argument "Reductionops.Stack.fold2" -> + with Reductionops.Stack.IncompatibleFold2 -> error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) in diff --git a/printing/prettyp.ml b/printing/prettyp.ml index aa422e36a..381af83c7 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -29,7 +29,7 @@ open Printer open Printmod open Context.Rel.Declaration -module RelDecl = Context.Rel.Declaration +(* module RelDecl = Context.Rel.Declaration *) module NamedDecl = Context.Named.Declaration type object_pr = { diff --git a/stm/stm.ml b/stm/stm.ml index 8c1185b6d..84c8aa9a9 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1205,6 +1205,8 @@ let detect_proof_block id name = (****************************** THE SCHEDULER *********************************) (******************************************************************************) +(* Unused module warning doesn't understand [module rec] *) +[@@@ocaml.warning "-60"] module rec ProofTask : sig type competence = Stateid.t list @@ -2318,6 +2320,7 @@ let known_state ?(redefine_qed=false) ~cache id = reach ~redefine_qed id end (* }}} *) +[@@@ocaml.warning "+60"] (********************************* STM API ************************************) (******************************************************************************) diff --git a/stm/vcs.ml b/stm/vcs.ml index d3886464d..88f860eb6 100644 --- a/stm/vcs.ml +++ b/stm/vcs.ml @@ -74,8 +74,6 @@ module Dag = Dag.Make(OT) type id = OT.t -module NodeSet = Dag.NodeSet - module Branch = struct type t = string diff --git a/tools/coqwc.mll b/tools/coqwc.mll index b4fc738d0..cd07d4216 100644 --- a/tools/coqwc.mll +++ b/tools/coqwc.mll @@ -239,6 +239,7 @@ let process_channel ch = if !skip_header then read_header lb; spec lb +[@@@ocaml.warning "-52"] let process_file f = try let ch = open_in f in @@ -251,6 +252,7 @@ let process_file f = flush stdout; eprintf "coqwc: %s: Is a directory\n" f; flush stderr | Sys_error s -> flush stdout; eprintf "coqwc: %s\n" s; flush stderr +[@@@ocaml.warning "+52"] (*s Parsing of the command line. *) -- cgit v1.2.3 From a77734ad64fff2396b161308099923037fb5d8a1 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Mon, 24 Apr 2017 13:43:23 +0200 Subject: Enable more warnings, and add -warn-error configure flag --- .merlin | 2 +- configure.ml | 31 ++++++++++++++++++++++++++++++- 2 files changed, 31 insertions(+), 2 deletions(-) diff --git a/.merlin b/.merlin index 5cae15f5f..f91e1b8fd 100644 --- a/.merlin +++ b/.merlin @@ -1,4 +1,4 @@ -FLG -rectypes -thread -safe-string +FLG -rectypes -thread -safe-string -w +a-4-9-27-41-42-44-45-48-50 S ltac B ltac diff --git a/configure.ml b/configure.ml index befd67262..5330da7d3 100644 --- a/configure.ml +++ b/configure.ml @@ -270,6 +270,7 @@ module Prefs = struct let nativecompiler = ref (not (os_type_win32 || os_type_cygwin)) let coqwebsite = ref "http://coq.inria.fr/" let force_caml_version = ref false + let warn_error = ref false end (* TODO : earlier any option -foo was also available as --foo *) @@ -352,6 +353,8 @@ let args_options = Arg.align [ " URL of the coq website"; "-force-caml-version", Arg.Set Prefs.force_caml_version, " Force OCaml version"; + "-warn-error", Arg.Set Prefs.warn_error, + " Make OCaml warnings into errors"; ] let parse_args () = @@ -511,6 +514,32 @@ let camltag = match caml_version_list with | x::y::_ -> "OCAML"^x^y | _ -> assert false +(** Explanation of disabled warnings: + 3: deprecated warning (not error for non minimum supported ocaml) + 4: fragile pattern matching: too common in the code and too annoying to avoid in general + 9: missing fields in a record pattern: too common in the code and not worth the bother + 27: innocuous unused variable: innocuous + 41: ambiguous constructor or label: too common + 42: disambiguated counstructor or label: too common + 44: "open" shadowing already defined identifier: too common, especially when some are aliases + 45: "open" shadowing a label or constructor: see 44 + 48: implicit elimination of optional arguments: too common + 50: unexpected documentation comment: too common and annoying to avoid + 56: unreachable match case: the [_ -> .] syntax doesn't exist in 4.02.3 +*) +let coq_warn_flags = + let warnings = "-w +a-4-9-27-41-42-44-45-48-50" in + let errors = + if !Prefs.warn_error + then "-warn-error +a" + ^ (if caml_version_nums > [4;2;3] + then "-3-56" + else "") + else "" + in + warnings ^ " " ^ errors + + (** * CamlpX configuration *) @@ -1103,7 +1132,7 @@ let write_makefile f = pr "CAMLHLIB=%S\n\n" camllib; pr "# Caml link command and Caml make top command\n"; pr "# Caml flags\n"; - pr "CAMLFLAGS=-rectypes %s %s\n" coq_annotate_flag coq_safe_string; + pr "CAMLFLAGS=-rectypes %s %s %s\n" coq_warn_flags coq_annotate_flag coq_safe_string; pr "# User compilation flag\n"; pr "USERFLAGS=\n\n"; pr "# Flags for GCC\n"; -- cgit v1.2.3 From 9a48211ea8439a8502145e508b70ede9b5929b2f Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Thu, 27 Apr 2017 21:58:52 +0200 Subject: Post-rebase warnings (unused opens and 2 unused values) --- plugins/cc/cctac.ml | 2 -- plugins/cc/cctac.mli | 1 - plugins/firstorder/formula.ml | 1 - plugins/firstorder/instances.ml | 1 - plugins/firstorder/rules.ml | 1 - plugins/firstorder/rules.mli | 1 - plugins/firstorder/sequent.ml | 1 - plugins/firstorder/sequent.mli | 2 -- plugins/ltac/g_rewrite.ml4 | 1 - plugins/ltac/rewrite.ml | 1 - pretyping/constr_matching.ml | 1 - tactics/auto.ml | 3 --- tactics/class_tactics.ml | 1 - tactics/class_tactics.mli | 1 - vernac/topfmt.ml | 1 - 15 files changed, 19 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 00b31ccce..7c5efaea3 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -15,13 +15,11 @@ open Declarations open Term open EConstr open Vars -open Tacmach open Tactics open Typing open Ccalgo open Ccproof open Pp -open CErrors open Util open Proofview.Notations diff --git a/plugins/cc/cctac.mli b/plugins/cc/cctac.mli index 5099d847b..b4bb62be8 100644 --- a/plugins/cc/cctac.mli +++ b/plugins/cc/cctac.mli @@ -8,7 +8,6 @@ (************************************************************************) open EConstr -open Proof_type val proof_tac: Ccproof.proof -> unit Proofview.tactic diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index ade94e98e..9900792ca 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -12,7 +12,6 @@ open Term open EConstr open Vars open Termops -open Tacmach open Util open Declarations open Globnames diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 2b624b983..5a1e7c26a 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -10,7 +10,6 @@ open Unify open Rules open CErrors open Util -open Term open EConstr open Vars open Tacmach.New diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index e0d2c38a7..86a677007 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -9,7 +9,6 @@ open CErrors open Util open Names -open Term open EConstr open Vars open Tacmach.New diff --git a/plugins/firstorder/rules.mli b/plugins/firstorder/rules.mli index 80a7ae2c2..fb2173083 100644 --- a/plugins/firstorder/rules.mli +++ b/plugins/firstorder/rules.mli @@ -8,7 +8,6 @@ open Term open EConstr -open Tacmach open Names open Globnames diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 59b842c82..2d18b6605 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -12,7 +12,6 @@ open CErrors open Util open Formula open Unify -open Tacmach open Globnames open Pp diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli index 18d68f54f..6ed251f34 100644 --- a/plugins/firstorder/sequent.mli +++ b/plugins/firstorder/sequent.mli @@ -6,10 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr open Formula -open Tacmach open Globnames module OrderedConstr: Set.OrderedType with type t=Constr.t diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4 index fdcaedab3..ac979bcf8 100644 --- a/plugins/ltac/g_rewrite.ml4 +++ b/plugins/ltac/g_rewrite.ml4 @@ -18,7 +18,6 @@ open Glob_term open Geninterp open Extraargs open Tacmach -open Tacticals open Proofview.Notations open Rewrite open Stdarg diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 9a1615d3f..5630a2d7b 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -17,7 +17,6 @@ open EConstr open Vars open Reduction open Tacticals.New -open Tacmach open Tactics open Pretype_errors open Typeclasses diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index d55350622..2334be966 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -87,7 +87,6 @@ let rec build_lambda sigma vars ctx m = match vars with | n :: vars -> (* change [ x1 ... xn y z1 ... zm |- t ] into [ x1 ... xn z1 ... zm |- lam y. t ] *) - let len = List.length ctx in let pre, suf = List.chop (pred n) ctx in let (na, t, suf) = match suf with | [] -> assert false diff --git a/tactics/auto.ml b/tactics/auto.ml index 324c551d0..c2d12ebd0 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -10,15 +10,12 @@ module CVars = Vars open Pp open Util -open CErrors open Names open Termops open EConstr open Environ -open Tacmach open Genredexpr open Tactics -open Tacticals open Clenv open Locus open Proofview.Notations diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 6c724a1eb..2d6dffdd2 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -19,7 +19,6 @@ open Term open Termops open EConstr open Proof_type -open Tacticals open Tacmach open Tactics open Clenv diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 4ab29f260..c5731e377 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -10,7 +10,6 @@ open Names open EConstr -open Tacmach val catchable : exn -> bool diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml index e44b585d7..6d9d71a62 100644 --- a/vernac/topfmt.ml +++ b/vernac/topfmt.ml @@ -131,7 +131,6 @@ let dbg_hdr = tag Tag.debug (str "Debug:") ++ spc () let info_hdr = mt () let warn_hdr = tag Tag.warning (str "Warning:") ++ spc () let err_hdr = tag Tag.error (str "Error:") ++ spc () -let ann_hdr = tag Tag.error (str "Anomaly:") ++ spc () let make_body quoter info ?pre_hdr s = pr_opt_no_spc (fun x -> x ++ fnl ()) pre_hdr ++ quoter (hov 0 (info ++ s)) -- cgit v1.2.3 From 68fb8e13c44c5ee95dbc9256b1d74c7c83303d2d Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 28 Apr 2017 16:30:45 +0200 Subject: Fixing #5487 (v8.5 regression on ltac-matching expressions with evars). The fix follows an invariant enforced in proofview.ml on the kind of evars that are goals or that occur in goals. One day, evar kinds will need a little cleaning... --- pretyping/patternops.ml | 4 +++- test-suite/bugs/closed/5487.v | 9 +++++++++ 2 files changed, 12 insertions(+), 1 deletion(-) create mode 100644 test-suite/bugs/closed/5487.v diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 2090aad8a..75d3ed30b 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -160,7 +160,9 @@ let pattern_of_constr env sigma t = let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in let () = ignore (pattern_of_constr env ty) in assert (not b); PMeta (Some id) - | Evar_kinds.GoalEvar -> + | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ -> + (* These are the two evar kinds used for existing goals *) + (* see Proofview.mark_in_evm *) PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v new file mode 100644 index 000000000..9b995f450 --- /dev/null +++ b/test-suite/bugs/closed/5487.v @@ -0,0 +1,9 @@ +(* Was a collision between an ltac pattern variable and an evar *) + +Goal forall n, exists m, n = m :> nat. +Proof. + eexists. + Fail match goal with + | [ |- ?x = ?y ] + => match x with y => idtac end + end. -- cgit v1.2.3 From 7bdfa1a4e46acf11d199a07bfca0bc59381874c3 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Wed, 26 Apr 2017 12:15:26 +0200 Subject: Renaming allow_patvar flag of intern_gen into pattern_mode. This highlights that this is a binary mode changing the interpretation of "?x" rather than additionally allowing patvar. --- interp/constrintern.ml | 12 ++++++------ interp/constrintern.mli | 2 +- plugins/funind/indfun.ml | 2 +- plugins/ltac/tacintern.ml | 4 ++-- plugins/ltac/tacinterp.ml | 4 ++-- 5 files changed, 12 insertions(+), 12 deletions(-) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d75487ecf..94924374a 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1516,7 +1516,7 @@ let extract_explicit_arg imps args = (**********************************************************************) (* Main loop *) -let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = +let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = let rec intern env = function | CRef (ref,us) as x -> let (c,imp,subscopes,l),_ = @@ -1747,9 +1747,9 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = in GHole (loc, k, naming, solve) (* Parsing pattern variables *) - | CPatVar (loc, n) when allow_patvar -> + | CPatVar (loc, n) when pattern_mode -> GPatVar (loc, (true,n)) - | CEvar (loc, n, []) when allow_patvar -> + | CEvar (loc, n, []) when pattern_mode -> GPatVar (loc, (false,n)) (* end *) (* Parsing existential variables *) @@ -1934,13 +1934,13 @@ let empty_ltac_sign = { } let intern_gen kind env - ?(impls=empty_internalization_env) ?(allow_patvar=false) ?(ltacvars=empty_ltac_sign) + ?(impls=empty_internalization_env) ?(pattern_mode=false) ?(ltacvars=empty_ltac_sign) c = let tmp_scope = scope_of_type_kind kind in internalize env {ids = extract_ids env; unb = false; tmp_scope = tmp_scope; scopes = []; impls = impls} - allow_patvar (ltacvars, Id.Map.empty) c + pattern_mode (ltacvars, Id.Map.empty) c let intern_constr env c = intern_gen WithoutTypeConstraint env c @@ -2013,7 +2013,7 @@ let interp_type_evars env evdref ?(impls=empty_internalization_env) c = let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c = let c = intern_gen (if as_type then IsType else WithoutTypeConstraint) - ~allow_patvar:true ~ltacvars env c in + ~pattern_mode:true ~ltacvars env c in pattern_of_glob_constr c let interp_notation_constr ?(impls=empty_internalization_env) nenv a = diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 758d4e650..2142d42bc 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -82,7 +82,7 @@ val intern_constr : env -> constr_expr -> glob_constr val intern_type : env -> constr_expr -> glob_constr val intern_gen : typing_constraint -> env -> - ?impls:internalization_env -> ?allow_patvar:bool -> ?ltacvars:ltac_sign -> + ?impls:internalization_env -> ?pattern_mode:bool -> ?ltacvars:ltac_sign -> constr_expr -> glob_constr val intern_pattern : env -> cases_pattern_expr -> diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index d335836df..5056d605d 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -142,7 +142,7 @@ let rec abstract_glob_constr c = function let interp_casted_constr_with_implicits env sigma impls c = Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls - ~allow_patvar:false c + c (* Construct a fixpoint as a Glob_term diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 3f83f104e..db47a9857 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -198,7 +198,7 @@ let intern_binding_name ist x = and if a term w/o ltac vars, check the name is indeed quantified *) x -let intern_constr_gen allow_patvar isarity {ltacvars=lfun; genv=env} c = +let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env} c = let warn = if !strict_check then fun x -> x else Constrintern.for_grammar in let scope = if isarity then Pretyping.IsType else Pretyping.WithoutTypeConstraint in let ltacvars = { @@ -206,7 +206,7 @@ let intern_constr_gen allow_patvar isarity {ltacvars=lfun; genv=env} c = ltac_bound = Id.Set.empty; } in let c' = - warn (Constrintern.intern_gen scope ~allow_patvar ~ltacvars env) c + warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env) c in (c',if !strict_check then None else Some c) diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 50f43931e..2d5dc5821 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -597,7 +597,7 @@ let interp_uconstr ist env sigma = function } in { closure ; term = intern_gen WithoutTypeConstraint ~ltacvars env ce } -let interp_gen kind ist allow_patvar flags env sigma (c,ce) = +let interp_gen kind ist pattern_mode flags env sigma (c,ce) = let constrvars = extract_ltac_constr_context ist env sigma in let vars = { Pretyping.ltac_constrs = constrvars.typed; @@ -624,7 +624,7 @@ let interp_gen kind ist allow_patvar flags env sigma (c,ce) = } in let kind_for_intern = match kind with OfType _ -> WithoutTypeConstraint | _ -> kind in - intern_gen kind_for_intern ~allow_patvar ~ltacvars env c + intern_gen kind_for_intern ~pattern_mode ~ltacvars env c in (* Jason Gross: To avoid unnecessary modifications to tacinterp, as suggested by Arnaud Spiwack, we run push_trace immediately. We do -- cgit v1.2.3 From 7943b1fade775af48917d54878e65b80217be038 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 28 Apr 2017 16:15:10 +0200 Subject: Using a more explicit algebraic type for evars of kind "MatchingVar". A priori considered to be a good programming style. --- interp/constrextern.ml | 12 +++++++----- interp/constrintern.ml | 4 ++-- intf/evar_kinds.mli | 5 ++++- intf/glob_term.mli | 2 +- pretyping/glob_ops.ml | 9 +++++++-- pretyping/patternops.ml | 11 ++++++----- pretyping/pretyping.ml | 4 ++-- tactics/hipattern.ml | 2 +- 8 files changed, 30 insertions(+), 19 deletions(-) diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 59b8b4e5b..bdbfd8c27 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -654,9 +654,11 @@ let rec extern inctx scopes vars r = | GEvar (loc,n,l) -> extern_evar loc n (List.map (on_snd (extern false scopes vars)) l) - | GPatVar (loc,(b,n)) -> - if !print_meta_as_hole then CHole (loc, None, Misctypes.IntroAnonymous, None) else - if b then CPatVar (loc,n) else CEvar (loc,n,[]) + | GPatVar (loc,kind) -> + if !print_meta_as_hole then CHole (loc, None, Misctypes.IntroAnonymous, None) else + (match kind with + | Evar_kinds.SecondOrderPatVar n -> CPatVar (loc,n) + | Evar_kinds.FirstOrderPatVar n -> CEvar (loc,n,[])) | GApp (loc,f,args) -> (match f with @@ -1029,13 +1031,13 @@ let rec glob_of_pat env sigma = function with Not_found -> Id.of_string ("_UNBOUND_REL_"^(string_of_int n)) in GVar (loc,id) | PMeta None -> GHole (loc,Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None) - | PMeta (Some n) -> GPatVar (loc,(false,n)) + | PMeta (Some n) -> GPatVar (loc,Evar_kinds.FirstOrderPatVar n) | PProj (p,c) -> GApp (loc,GRef (loc, ConstRef (Projection.constant p),None), [glob_of_pat env sigma c]) | PApp (f,args) -> GApp (loc,glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args) | PSoApp (n,args) -> - GApp (loc,GPatVar (loc,(true,n)), + GApp (loc,GPatVar (loc,Evar_kinds.SecondOrderPatVar n), List.map (glob_of_pat env sigma) args) | PProd (na,t,c) -> GProd (loc,na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 94924374a..ea1802ccf 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1748,9 +1748,9 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = GHole (loc, k, naming, solve) (* Parsing pattern variables *) | CPatVar (loc, n) when pattern_mode -> - GPatVar (loc, (true,n)) + GPatVar (loc, Evar_kinds.SecondOrderPatVar n) | CEvar (loc, n, []) when pattern_mode -> - GPatVar (loc, (false,n)) + GPatVar (loc, Evar_kinds.FirstOrderPatVar n) (* end *) (* Parsing existential variables *) | CEvar (loc, n, l) -> diff --git a/intf/evar_kinds.mli b/intf/evar_kinds.mli index 470ad2a23..75dad2e91 100644 --- a/intf/evar_kinds.mli +++ b/intf/evar_kinds.mli @@ -8,6 +8,7 @@ open Names open Globnames +open Misctypes (** The kinds of existential variable *) @@ -16,6 +17,8 @@ open Globnames type obligation_definition_status = Define of bool | Expand +type matching_var_kind = FirstOrderPatVar of patvar | SecondOrderPatVar of patvar + type t = | ImplicitArg of global_reference * (int * Id.t option) * bool (** Force inference *) @@ -27,6 +30,6 @@ type t = | TomatchTypeParameter of inductive * int | GoalEvar | ImpossibleCase - | MatchingVar of bool * Id.t + | MatchingVar of matching_var_kind | VarInstance of Id.t | SubEvar of Constr.existential_key diff --git a/intf/glob_term.mli b/intf/glob_term.mli index ced5a8b44..ba4a47a36 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -38,7 +38,7 @@ type glob_constr = (** An identifier that cannot be regarded as "GRef". Bound variables are typically represented this way. *) | GEvar of Loc.t * existential_name * (Id.t * glob_constr) list - | GPatVar of Loc.t * (bool * patvar) (** Used for patterns only *) + | GPatVar of Loc.t * Evar_kinds.matching_var_kind (** Used for patterns only *) | GApp of Loc.t * glob_constr * glob_constr list | GLambda of Loc.t * Name.t * binding_kind * glob_constr * glob_constr | GProd of Loc.t * Name.t * binding_kind * glob_constr * glob_constr diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index ebbfa195f..7fb539fb5 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -12,6 +12,7 @@ open Nameops open Globnames open Misctypes open Glob_term +open Evar_kinds (* Untyped intermediate terms, after ASTs and before constr. *) @@ -61,14 +62,18 @@ let cast_type_eq eq t1 t2 = match t1, t2 with | CastNative t1, CastNative t2 -> eq t1 t2 | _ -> false +let matching_var_kind_eq k1 k2 = match k1, k2 with +| FirstOrderPatVar ido1, FirstOrderPatVar ido2 -> Id.equal ido1 ido2 +| SecondOrderPatVar id1, SecondOrderPatVar id2 -> Id.equal id1 id2 +| _ -> false + let rec glob_constr_eq c1 c2 = match c1, c2 with | GRef (_, gr1, _), GRef (_, gr2, _) -> eq_gr gr1 gr2 | GVar (_, id1), GVar (_, id2) -> Id.equal id1 id2 | GEvar (_, id1, arg1), GEvar (_, id2, arg2) -> Id.equal id1 id2 && List.equal instance_eq arg1 arg2 -| GPatVar (_, (b1, pat1)), GPatVar (_, (b2, pat2)) -> - (b1 : bool) == b2 && Id.equal pat1 pat2 +| GPatVar (_, k1), GPatVar (_, k2) -> matching_var_kind_eq k1 k2 | GApp (_, f1, arg1), GApp (_, f2, arg2) -> glob_constr_eq f1 f2 && List.equal glob_constr_eq arg1 arg2 | GLambda (_, na1, bk1, t1, c1), GLambda (_, na2, bk2, t2, c2) -> diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index b16d04495..b310da9a3 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -145,7 +145,7 @@ let pattern_of_constr env sigma t = match kind_of_term f with | Evar (evk,args) -> (match snd (Evd.evar_source evk sigma) with - Evar_kinds.MatchingVar (true,id) -> Some id + Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar id) -> Some id | _ -> None) | _ -> None with @@ -158,10 +158,11 @@ let pattern_of_constr env sigma t = pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) | Evar (evk,ctxt) -> (match snd (Evd.evar_source evk sigma) with - | Evar_kinds.MatchingVar (b,id) -> - assert (not b); PMeta (Some id) + | Evar_kinds.MatchingVar (Evar_kinds.FirstOrderPatVar id) -> + PMeta (Some id) | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt) + | Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar ido) -> assert false | _ -> PMeta None) | Case (ci,p,a,br) -> @@ -328,12 +329,12 @@ let rec pat_of_raw metas vars = function | GVar (_,id) -> (try PRel (List.index Name.equal (Name id) vars) with Not_found -> PVar id) - | GPatVar (_,(false,n)) -> + | GPatVar (_,Evar_kinds.FirstOrderPatVar n) -> metas := n::!metas; PMeta (Some n) | GRef (_,gr,_) -> PRef (canonical_gr gr) (* Hack to avoid rewriting a complete interpretation of patterns *) - | GApp (_, GPatVar (_,(true,n)), cl) -> + | GApp (_, GPatVar (_,Evar_kinds.SecondOrderPatVar n), cl) -> metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl) | GApp (_,c,cl) -> PApp (pat_of_raw metas vars c, diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index ae87cd8c0..6c490045c 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -596,13 +596,13 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in inh_conv_coerce_to_tycon loc env evdref j tycon - | GPatVar (loc,(someta,n)) -> + | GPatVar (loc,kind) -> let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with | Some ty -> ty | None -> new_type_evar env evdref loc in - let k = Evar_kinds.MatchingVar (someta,n) in + let k = Evar_kinds.MatchingVar kind in { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } | GHole (loc, k, naming, None) -> diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 851554b83..257c092d7 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -261,7 +261,7 @@ let mkGProd id c1 c2 = let mkGArrow c1 c2 = GProd (Loc.ghost, Anonymous, Explicit, c1, c2) let mkGVar id = GVar (Loc.ghost, Id.of_string id) -let mkGPatVar id = GPatVar(Loc.ghost, (false, Id.of_string id)) +let mkGPatVar id = GPatVar(Loc.ghost, Evar_kinds.FirstOrderPatVar (Id.of_string id)) let mkGRef r = GRef (Loc.ghost, Lazy.force r, None) let mkGAppRef r args = mkGApp (mkGRef r) args -- cgit v1.2.3 From e1fc5c4dae82acd2eb6618724599aca368c200b7 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sun, 4 Dec 2016 14:09:24 -0500 Subject: Add .dir-locals.el and _CoqProject files for emacs stdlib editing These set up PG to use the local coqtop, and the local coqlib, when editing files in the stdlib. As per https://github.com/coq/coq/pull/386#issuecomment-279012238, we can use `_CoqProject` for `theories/Init`, and this allows CoqIDE to edit those files. However, we cannot use it for `theories/`, because a `_CoqProject` file will override a `.dir-locals.el` in the same directory, and there is no way to get PG to pick up a valid `-coqlib` from `_CoqProject` (because it'll take the path relative to the current directory, not relative to the directory of `_CoqProject`). --- theories/.dir-locals.el | 4 ++++ theories/Init/_CoqProject | 2 ++ 2 files changed, 6 insertions(+) create mode 100644 theories/.dir-locals.el create mode 100644 theories/Init/_CoqProject diff --git a/theories/.dir-locals.el b/theories/.dir-locals.el new file mode 100644 index 000000000..4e8830f6c --- /dev/null +++ b/theories/.dir-locals.el @@ -0,0 +1,4 @@ +((coq-mode . ((eval . (let ((default-directory (locate-dominating-file + buffer-file-name ".dir-locals.el"))) + (setq-local coq-prog-args `("-coqlib" ,(expand-file-name "..") "-R" ,(expand-file-name ".") "Coq")) + (setq-local coq-prog-name (expand-file-name "../bin/coqtop"))))))) diff --git a/theories/Init/_CoqProject b/theories/Init/_CoqProject new file mode 100644 index 000000000..bff79d34b --- /dev/null +++ b/theories/Init/_CoqProject @@ -0,0 +1,2 @@ +-R .. Coq +-arg -noinit -- cgit v1.2.3 From 6deddd3e0c6cc1d7cbf8414f7553c5d7752a0801 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 28 Apr 2017 14:02:00 -0400 Subject: Allow interactive editing of {C,}Morphisms in PG --- theories/Classes/CMorphisms.v | 2 +- theories/Classes/Morphisms.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Classes/CMorphisms.v b/theories/Classes/CMorphisms.v index 1cfca4169..74ca5d4c8 100644 --- a/theories/Classes/CMorphisms.v +++ b/theories/Classes/CMorphisms.v @@ -1,4 +1,4 @@ -(* -*- coding: utf-8 -*- *) +(* -*- coding: utf-8; coq-prog-args: ("-coqlib" "../.." "-R" ".." "Coq" "-top" "Coq.Classes.CMorphisms") -*- *) (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* Date: Fri, 28 Apr 2017 22:14:09 +0200 Subject: Revert "Using a more explicit algebraic type for evars of kind "MatchingVar"." I'm sure this was pushed by accident, since testing shows immediately that it breaks the compilation of the ssreflect plugin, hence all developments relying on it in Travis. --- interp/constrextern.ml | 12 +++++------- interp/constrintern.ml | 4 ++-- intf/evar_kinds.mli | 5 +---- intf/glob_term.mli | 2 +- pretyping/glob_ops.ml | 9 ++------- pretyping/patternops.ml | 11 +++++------ pretyping/pretyping.ml | 4 ++-- tactics/hipattern.ml | 2 +- 8 files changed, 19 insertions(+), 30 deletions(-) diff --git a/interp/constrextern.ml b/interp/constrextern.ml index bdbfd8c27..59b8b4e5b 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -654,11 +654,9 @@ let rec extern inctx scopes vars r = | GEvar (loc,n,l) -> extern_evar loc n (List.map (on_snd (extern false scopes vars)) l) - | GPatVar (loc,kind) -> - if !print_meta_as_hole then CHole (loc, None, Misctypes.IntroAnonymous, None) else - (match kind with - | Evar_kinds.SecondOrderPatVar n -> CPatVar (loc,n) - | Evar_kinds.FirstOrderPatVar n -> CEvar (loc,n,[])) + | GPatVar (loc,(b,n)) -> + if !print_meta_as_hole then CHole (loc, None, Misctypes.IntroAnonymous, None) else + if b then CPatVar (loc,n) else CEvar (loc,n,[]) | GApp (loc,f,args) -> (match f with @@ -1031,13 +1029,13 @@ let rec glob_of_pat env sigma = function with Not_found -> Id.of_string ("_UNBOUND_REL_"^(string_of_int n)) in GVar (loc,id) | PMeta None -> GHole (loc,Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None) - | PMeta (Some n) -> GPatVar (loc,Evar_kinds.FirstOrderPatVar n) + | PMeta (Some n) -> GPatVar (loc,(false,n)) | PProj (p,c) -> GApp (loc,GRef (loc, ConstRef (Projection.constant p),None), [glob_of_pat env sigma c]) | PApp (f,args) -> GApp (loc,glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args) | PSoApp (n,args) -> - GApp (loc,GPatVar (loc,Evar_kinds.SecondOrderPatVar n), + GApp (loc,GPatVar (loc,(true,n)), List.map (glob_of_pat env sigma) args) | PProd (na,t,c) -> GProd (loc,na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index ea1802ccf..94924374a 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1748,9 +1748,9 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = GHole (loc, k, naming, solve) (* Parsing pattern variables *) | CPatVar (loc, n) when pattern_mode -> - GPatVar (loc, Evar_kinds.SecondOrderPatVar n) + GPatVar (loc, (true,n)) | CEvar (loc, n, []) when pattern_mode -> - GPatVar (loc, Evar_kinds.FirstOrderPatVar n) + GPatVar (loc, (false,n)) (* end *) (* Parsing existential variables *) | CEvar (loc, n, l) -> diff --git a/intf/evar_kinds.mli b/intf/evar_kinds.mli index 75dad2e91..470ad2a23 100644 --- a/intf/evar_kinds.mli +++ b/intf/evar_kinds.mli @@ -8,7 +8,6 @@ open Names open Globnames -open Misctypes (** The kinds of existential variable *) @@ -17,8 +16,6 @@ open Misctypes type obligation_definition_status = Define of bool | Expand -type matching_var_kind = FirstOrderPatVar of patvar | SecondOrderPatVar of patvar - type t = | ImplicitArg of global_reference * (int * Id.t option) * bool (** Force inference *) @@ -30,6 +27,6 @@ type t = | TomatchTypeParameter of inductive * int | GoalEvar | ImpossibleCase - | MatchingVar of matching_var_kind + | MatchingVar of bool * Id.t | VarInstance of Id.t | SubEvar of Constr.existential_key diff --git a/intf/glob_term.mli b/intf/glob_term.mli index ba4a47a36..ced5a8b44 100644 --- a/intf/glob_term.mli +++ b/intf/glob_term.mli @@ -38,7 +38,7 @@ type glob_constr = (** An identifier that cannot be regarded as "GRef". Bound variables are typically represented this way. *) | GEvar of Loc.t * existential_name * (Id.t * glob_constr) list - | GPatVar of Loc.t * Evar_kinds.matching_var_kind (** Used for patterns only *) + | GPatVar of Loc.t * (bool * patvar) (** Used for patterns only *) | GApp of Loc.t * glob_constr * glob_constr list | GLambda of Loc.t * Name.t * binding_kind * glob_constr * glob_constr | GProd of Loc.t * Name.t * binding_kind * glob_constr * glob_constr diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 7fb539fb5..ebbfa195f 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -12,7 +12,6 @@ open Nameops open Globnames open Misctypes open Glob_term -open Evar_kinds (* Untyped intermediate terms, after ASTs and before constr. *) @@ -62,18 +61,14 @@ let cast_type_eq eq t1 t2 = match t1, t2 with | CastNative t1, CastNative t2 -> eq t1 t2 | _ -> false -let matching_var_kind_eq k1 k2 = match k1, k2 with -| FirstOrderPatVar ido1, FirstOrderPatVar ido2 -> Id.equal ido1 ido2 -| SecondOrderPatVar id1, SecondOrderPatVar id2 -> Id.equal id1 id2 -| _ -> false - let rec glob_constr_eq c1 c2 = match c1, c2 with | GRef (_, gr1, _), GRef (_, gr2, _) -> eq_gr gr1 gr2 | GVar (_, id1), GVar (_, id2) -> Id.equal id1 id2 | GEvar (_, id1, arg1), GEvar (_, id2, arg2) -> Id.equal id1 id2 && List.equal instance_eq arg1 arg2 -| GPatVar (_, k1), GPatVar (_, k2) -> matching_var_kind_eq k1 k2 +| GPatVar (_, (b1, pat1)), GPatVar (_, (b2, pat2)) -> + (b1 : bool) == b2 && Id.equal pat1 pat2 | GApp (_, f1, arg1), GApp (_, f2, arg2) -> glob_constr_eq f1 f2 && List.equal glob_constr_eq arg1 arg2 | GLambda (_, na1, bk1, t1, c1), GLambda (_, na2, bk2, t2, c2) -> diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index b310da9a3..b16d04495 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -145,7 +145,7 @@ let pattern_of_constr env sigma t = match kind_of_term f with | Evar (evk,args) -> (match snd (Evd.evar_source evk sigma) with - Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar id) -> Some id + Evar_kinds.MatchingVar (true,id) -> Some id | _ -> None) | _ -> None with @@ -158,11 +158,10 @@ let pattern_of_constr env sigma t = pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) | Evar (evk,ctxt) -> (match snd (Evd.evar_source evk sigma) with - | Evar_kinds.MatchingVar (Evar_kinds.FirstOrderPatVar id) -> - PMeta (Some id) + | Evar_kinds.MatchingVar (b,id) -> + assert (not b); PMeta (Some id) | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt) - | Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar ido) -> assert false | _ -> PMeta None) | Case (ci,p,a,br) -> @@ -329,12 +328,12 @@ let rec pat_of_raw metas vars = function | GVar (_,id) -> (try PRel (List.index Name.equal (Name id) vars) with Not_found -> PVar id) - | GPatVar (_,Evar_kinds.FirstOrderPatVar n) -> + | GPatVar (_,(false,n)) -> metas := n::!metas; PMeta (Some n) | GRef (_,gr,_) -> PRef (canonical_gr gr) (* Hack to avoid rewriting a complete interpretation of patterns *) - | GApp (_, GPatVar (_,Evar_kinds.SecondOrderPatVar n), cl) -> + | GApp (_, GPatVar (_,(true,n)), cl) -> metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl) | GApp (_,c,cl) -> PApp (pat_of_raw metas vars c, diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 6c490045c..ae87cd8c0 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -596,13 +596,13 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in inh_conv_coerce_to_tycon loc env evdref j tycon - | GPatVar (loc,kind) -> + | GPatVar (loc,(someta,n)) -> let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with | Some ty -> ty | None -> new_type_evar env evdref loc in - let k = Evar_kinds.MatchingVar kind in + let k = Evar_kinds.MatchingVar (someta,n) in { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } | GHole (loc, k, naming, None) -> diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 257c092d7..851554b83 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -261,7 +261,7 @@ let mkGProd id c1 c2 = let mkGArrow c1 c2 = GProd (Loc.ghost, Anonymous, Explicit, c1, c2) let mkGVar id = GVar (Loc.ghost, Id.of_string id) -let mkGPatVar id = GPatVar(Loc.ghost, Evar_kinds.FirstOrderPatVar (Id.of_string id)) +let mkGPatVar id = GPatVar(Loc.ghost, (false, Id.of_string id)) let mkGRef r = GRef (Loc.ghost, Lazy.force r, None) let mkGAppRef r args = mkGApp (mkGRef r) args -- cgit v1.2.3 From 1c8ed18d5f67d7d5656342584b8dcf8a441cb87f Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 28 Apr 2017 22:15:06 +0200 Subject: Revert "Renaming allow_patvar flag of intern_gen into pattern_mode." This reverts commit 7bdfa1a4e46acf11d199a07bfca0bc59381874c3. --- interp/constrintern.ml | 12 ++++++------ interp/constrintern.mli | 2 +- plugins/funind/indfun.ml | 2 +- plugins/ltac/tacintern.ml | 4 ++-- plugins/ltac/tacinterp.ml | 4 ++-- 5 files changed, 12 insertions(+), 12 deletions(-) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 94924374a..d75487ecf 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1516,7 +1516,7 @@ let extract_explicit_arg imps args = (**********************************************************************) (* Main loop *) -let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = +let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = let rec intern env = function | CRef (ref,us) as x -> let (c,imp,subscopes,l),_ = @@ -1747,9 +1747,9 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c = in GHole (loc, k, naming, solve) (* Parsing pattern variables *) - | CPatVar (loc, n) when pattern_mode -> + | CPatVar (loc, n) when allow_patvar -> GPatVar (loc, (true,n)) - | CEvar (loc, n, []) when pattern_mode -> + | CEvar (loc, n, []) when allow_patvar -> GPatVar (loc, (false,n)) (* end *) (* Parsing existential variables *) @@ -1934,13 +1934,13 @@ let empty_ltac_sign = { } let intern_gen kind env - ?(impls=empty_internalization_env) ?(pattern_mode=false) ?(ltacvars=empty_ltac_sign) + ?(impls=empty_internalization_env) ?(allow_patvar=false) ?(ltacvars=empty_ltac_sign) c = let tmp_scope = scope_of_type_kind kind in internalize env {ids = extract_ids env; unb = false; tmp_scope = tmp_scope; scopes = []; impls = impls} - pattern_mode (ltacvars, Id.Map.empty) c + allow_patvar (ltacvars, Id.Map.empty) c let intern_constr env c = intern_gen WithoutTypeConstraint env c @@ -2013,7 +2013,7 @@ let interp_type_evars env evdref ?(impls=empty_internalization_env) c = let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c = let c = intern_gen (if as_type then IsType else WithoutTypeConstraint) - ~pattern_mode:true ~ltacvars env c in + ~allow_patvar:true ~ltacvars env c in pattern_of_glob_constr c let interp_notation_constr ?(impls=empty_internalization_env) nenv a = diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 2142d42bc..758d4e650 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -82,7 +82,7 @@ val intern_constr : env -> constr_expr -> glob_constr val intern_type : env -> constr_expr -> glob_constr val intern_gen : typing_constraint -> env -> - ?impls:internalization_env -> ?pattern_mode:bool -> ?ltacvars:ltac_sign -> + ?impls:internalization_env -> ?allow_patvar:bool -> ?ltacvars:ltac_sign -> constr_expr -> glob_constr val intern_pattern : env -> cases_pattern_expr -> diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 5056d605d..d335836df 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -142,7 +142,7 @@ let rec abstract_glob_constr c = function let interp_casted_constr_with_implicits env sigma impls c = Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls - c + ~allow_patvar:false c (* Construct a fixpoint as a Glob_term diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index db47a9857..3f83f104e 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -198,7 +198,7 @@ let intern_binding_name ist x = and if a term w/o ltac vars, check the name is indeed quantified *) x -let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env} c = +let intern_constr_gen allow_patvar isarity {ltacvars=lfun; genv=env} c = let warn = if !strict_check then fun x -> x else Constrintern.for_grammar in let scope = if isarity then Pretyping.IsType else Pretyping.WithoutTypeConstraint in let ltacvars = { @@ -206,7 +206,7 @@ let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env} c = ltac_bound = Id.Set.empty; } in let c' = - warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env) c + warn (Constrintern.intern_gen scope ~allow_patvar ~ltacvars env) c in (c',if !strict_check then None else Some c) diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 2d5dc5821..50f43931e 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -597,7 +597,7 @@ let interp_uconstr ist env sigma = function } in { closure ; term = intern_gen WithoutTypeConstraint ~ltacvars env ce } -let interp_gen kind ist pattern_mode flags env sigma (c,ce) = +let interp_gen kind ist allow_patvar flags env sigma (c,ce) = let constrvars = extract_ltac_constr_context ist env sigma in let vars = { Pretyping.ltac_constrs = constrvars.typed; @@ -624,7 +624,7 @@ let interp_gen kind ist pattern_mode flags env sigma (c,ce) = } in let kind_for_intern = match kind with OfType _ -> WithoutTypeConstraint | _ -> kind in - intern_gen kind_for_intern ~pattern_mode ~ltacvars env c + intern_gen kind_for_intern ~allow_patvar ~ltacvars env c in (* Jason Gross: To avoid unnecessary modifications to tacinterp, as suggested by Arnaud Spiwack, we run push_trace immediately. We do -- cgit v1.2.3 From db28e827d21658797418c320d566fb99570b44b6 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 28 Apr 2017 22:20:35 +0200 Subject: Revert "Fixing #5487 (v8.5 regression on ltac-matching expressions with evars)." One day I'll get bored of spending my nights fixing commits that were pushed without being tested, and I'll ask for removal of push rights. But for now let's pretend I haven't insisted enough: ~~~~ PLEASE TEST YOUR COMMITS BEFORE PUSHING ~~~~ Thank you! --- pretyping/patternops.ml | 4 +--- test-suite/bugs/closed/5487.v | 9 --------- 2 files changed, 1 insertion(+), 12 deletions(-) delete mode 100644 test-suite/bugs/closed/5487.v diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 75d3ed30b..2090aad8a 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -160,9 +160,7 @@ let pattern_of_constr env sigma t = let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in let () = ignore (pattern_of_constr env ty) in assert (not b); PMeta (Some id) - | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ -> - (* These are the two evar kinds used for existing goals *) - (* see Proofview.mark_in_evm *) + | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v deleted file mode 100644 index 9b995f450..000000000 --- a/test-suite/bugs/closed/5487.v +++ /dev/null @@ -1,9 +0,0 @@ -(* Was a collision between an ltac pattern variable and an evar *) - -Goal forall n, exists m, n = m :> nat. -Proof. - eexists. - Fail match goal with - | [ |- ?x = ?y ] - => match x with y => idtac end - end. -- cgit v1.2.3 From 991b78fd9627ee76f1a1a39b8460bf361c6af53d Mon Sep 17 00:00:00 2001 From: Guillaume Melquiond Date: Sat, 29 Apr 2017 09:09:35 +0200 Subject: Suppress warning message in stdlib. --- theories/Reals/Rfunctions.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v index 653fac8d9..606136333 100644 --- a/theories/Reals/Rfunctions.v +++ b/theories/Reals/Rfunctions.v @@ -704,7 +704,7 @@ Lemma powerRZ_ind (P : Z -> R -> R -> Prop) : forall x (m : Z), P m x (powerRZ x m)%R. Proof. intros ? ? ? x m. - destruct (intP m) as [n Hm| n Hm|n Hm]. + destruct (intP m) as [Hm|n Hm|n Hm]. - easy. - now rewrite <- pow_powerRZ. - unfold powerRZ. -- cgit v1.2.3 From edb21eab674b170d125a5b6fbc8213066b356d17 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 6 Apr 2017 10:10:03 +0200 Subject: Fixing "decide equality"'s bug #5449. The code was assuming that the terms t and u for which {t=u}+{t<>u} is proved were distinct. We refine an internal "generalize" of "u" so that it works on the two precise occurrences to abstract, even if other occurrences of u occur as subterm of t too. We also reuse the global constants found in the statement rather than reconstructing them (this seems better in case the global constants eventually get polymorphic universes?). --- tactics/eqdecide.ml | 57 ++++++++++++++++++++++++++++++++++--------- test-suite/bugs/closed/5449.v | 6 +++++ 2 files changed, 52 insertions(+), 11 deletions(-) create mode 100644 test-suite/bugs/closed/5449.v diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index bac3980d2..7d8b84f9c 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -67,9 +67,27 @@ let choose_noteq eqonleft = else left_with_bindings false Misctypes.NoBindings -let mkBranches c1 c2 = +open Sigma.Notations +open Context.Rel.Declaration + +(* A surgical generalize which selects the right occurrences by hand *) +(* This prevents issues where c2 is also a subterm of c1 (see e.g. #5449) *) + +let generalize_right mk typ c1 c2 = + Proofview.Goal.enter { Proofview.Goal.enter = begin fun gl -> + let env = Proofview.Goal.env gl in + let store = Proofview.Goal.extra gl in + Refine.refine ~unsafe:true { Sigma.run = begin fun sigma -> + let na = Name (next_name_away_with_default "x" Anonymous (Termops.ids_of_context env)) in + let newconcl = mkProd (na, typ, mk typ c1 (mkRel 1)) in + let Sigma (x, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store newconcl in + Sigma (mkApp (x, [|c2|]), sigma, p) + end } + end } + +let mkBranches (eqonleft,mk,c1,c2,typ) = tclTHENLIST - [generalize [c2]; + [generalize_right mk typ c1 c2; Simple.elim c1; intros; onLastHyp Simple.case; @@ -145,15 +163,32 @@ let diseqCase hyps eqonleft = open Proofview.Notations -(* spiwack: a small wrapper around [Hipattern]. *) +(* spiwack: a PatternMatchingFailure wrapper around [Hipattern]. *) let match_eqdec sigma c = - try Proofview.tclUNIT (match_eqdec sigma c) + try + let (eqonleft,_,c1,c2,ty) = match_eqdec sigma c in + let (op,eq1,noteq,eq2) = + match EConstr.kind sigma c with + | App (op,[|ty1;ty2|]) -> + let ty1, ty2 = if eqonleft then ty1, ty2 else ty2, ty1 in + (match EConstr.kind sigma ty1, EConstr.kind sigma ty2 with + | App (eq1,_), App (noteq,[|neq|]) -> + (match EConstr.kind sigma neq with + | App (eq2,_) -> op,eq1,noteq,eq2 + | _ -> assert false) + | _ -> assert false) + | _ -> assert false in + let mk t x y = + let eq = mkApp (eq1,[|t;x;y|]) in + let neq = mkApp (noteq,[|mkApp (eq2,[|t;x;y|])|]) in + if eqonleft then mkApp (op,[|eq;neq|]) else mkApp (op,[|neq;eq|]) in + Proofview.tclUNIT (eqonleft,mk,c1,c2,ty) with PatternMatchingFailure -> Proofview.tclZERO PatternMatchingFailure (* /spiwack *) -let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with +let rec solveArg hyps eqonleft mk largs rargs = match largs, rargs with | [], [] -> tclTHENLIST [ choose_eq eqonleft; @@ -163,8 +198,8 @@ let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with | a1 :: largs, a2 :: rargs -> Proofview.Goal.enter { enter = begin fun gl -> let rectype = pf_unsafe_type_of gl a1 in - let decide = mkDecideEqGoal eqonleft op rectype a1 a2 in - let tac hyp = solveArg (hyp :: hyps) eqonleft op largs rargs in + let decide = mk rectype a1 a2 in + let tac hyp = solveArg (hyp :: hyps) eqonleft mk largs rargs in let subtacs = if eqonleft then [eqCase tac;diseqCase hyps eqonleft;default_auto] else [diseqCase hyps eqonleft;eqCase tac;default_auto] in @@ -178,13 +213,13 @@ let solveEqBranch rectype = Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_concl gl in let sigma = project gl in - match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) -> + match_eqdec sigma concl >>= fun (eqonleft,mk,lhs,rhs,_) -> let (mib,mip) = Global.lookup_inductive rectype in let nparams = mib.mind_nparams in let getargs l = List.skipn nparams (snd (decompose_app sigma l)) in let rargs = getargs rhs and largs = getargs lhs in - solveArg [] eqonleft op largs rargs + solveArg [] eqonleft mk largs rargs end } end begin function (e, info) -> match e with @@ -204,14 +239,14 @@ let decideGralEquality = Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_concl gl in let sigma = project gl in - match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) -> + match_eqdec sigma concl >>= fun (eqonleft,mk,c1,c2,typ as data) -> let headtyp = hd_app sigma (pf_compute gl typ) in begin match EConstr.kind sigma headtyp with | Ind (mi,_) -> Proofview.tclUNIT mi | _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.") end >>= fun rectype -> (tclTHEN - (mkBranches c1 c2) + (mkBranches data) (tclORELSE (solveNoteqBranch eqonleft) (solveEqBranch rectype))) end } end diff --git a/test-suite/bugs/closed/5449.v b/test-suite/bugs/closed/5449.v new file mode 100644 index 000000000..d7fc2aaa0 --- /dev/null +++ b/test-suite/bugs/closed/5449.v @@ -0,0 +1,6 @@ +(* An example of decide equality which was failing due to a lhs dep into the rhs *) + +Require Import Coq.PArith.BinPos. +Goal forall x y, {Pos.compare_cont Gt x y = Gt} + {Pos.compare_cont Gt x y <> Gt}. +intros. +decide equality. -- cgit v1.2.3 From 5c7163d2ee1412fa5af523fbcd275518fc61fbde Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 30 Apr 2017 12:14:38 +0200 Subject: Fix bug #5501: Universe polymorphism breaks proof involving auto. A universe substitution was lacking as the normalized evar map was dropped. --- tactics/hints.ml | 7 ++----- test-suite/bugs/closed/5501.v | 21 +++++++++++++++++++++ 2 files changed, 23 insertions(+), 5 deletions(-) create mode 100644 test-suite/bugs/closed/5501.v diff --git a/tactics/hints.ml b/tactics/hints.ml index bcc068d3d..50b811edb 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1238,18 +1238,15 @@ let prepare_hint check (poly,local) env init (sigma,c) = (* We re-abstract over uninstantiated evars and universes. It is actually a bit stupid to generalize over evars since the first thing make_resolves will do is to re-instantiate the products *) - let sigma, subst = Evd.nf_univ_variables sigma in + let sigma, _ = Evd.nf_univ_variables sigma in let c = Evarutil.nf_evar sigma c in - let c = EConstr.Unsafe.to_constr c in - let c = CVars.subst_univs_constr subst c in - let c = EConstr.of_constr c in let c = drop_extra_implicit_args sigma c in let vars = ref (collect_vars sigma c) in let subst = ref [] in let rec find_next_evar c = match EConstr.kind sigma c with | Evar (evk,args as ev) -> (* We skip the test whether args is the identity or not *) - let t = existential_type sigma ev in + let t = Evarutil.nf_evar sigma (existential_type sigma ev) in let t = List.fold_right (fun (e,id) c -> replace_term sigma e id c) !subst t in if not (closed0 sigma c) then error "Hints with holes dependent on a bound variable not supported."; diff --git a/test-suite/bugs/closed/5501.v b/test-suite/bugs/closed/5501.v new file mode 100644 index 000000000..24739a365 --- /dev/null +++ b/test-suite/bugs/closed/5501.v @@ -0,0 +1,21 @@ +Set Universe Polymorphism. + +Record Pred@{A} := + { car :> Type@{A} + ; P : car -> Prop + }. + +Class All@{A} (A : Pred@{A}) : Type := + { proof : forall (a : A), P A a + }. + +Record Pred_All@{A} : Type := + { P' :> Pred@{A} + ; P'_All : All P' + }. + +Global Instance Pred_All_instance (A : Pred_All) : All A := P'_All A. + +Definition Pred_All_proof {A : Pred_All} (a : A) : P A a. +Proof. +solve[auto using proof]. -- cgit v1.2.3 From fdd5a8452bd2da22ffd1cab3b1888f2261f193b9 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Sun, 30 Apr 2017 13:10:48 +0200 Subject: Functional choice <-> [inhabited] and [forall] commute --- theories/Init/Logic.v | 5 +++++ theories/Init/Specif.v | 11 +++++++++++ theories/Logic/ChoiceFacts.v | 33 +++++++++++++++++++++++++++++++++ 3 files changed, 49 insertions(+) diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v index 9ae9dde28..3eefe9a84 100644 --- a/theories/Init/Logic.v +++ b/theories/Init/Logic.v @@ -609,6 +609,11 @@ Proof. destruct 1; auto. Qed. +Lemma inhabited_covariant (A B : Type) : (A -> B) -> inhabited A -> inhabited B. +Proof. + intros f [x];exact (inhabits (f x)). +Qed. + (** Declaration of stepl and stepr for eq and iff *) Lemma eq_stepl : forall (A : Type) (x y z : A), x = y -> x = z -> z = y. diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v index 2cc2ecbc2..43a441fc5 100644 --- a/theories/Init/Specif.v +++ b/theories/Init/Specif.v @@ -207,6 +207,17 @@ Definition sig2_eta {A P Q} (p : { a : A | P a & Q a }) : p = exist2 _ _ (proj1_sig (sig_of_sig2 p)) (proj2_sig (sig_of_sig2 p)) (proj3_sig p). Proof. destruct p; reflexivity. Defined. +(** [exists x : A, B] is equivalent to [inhabited {x : A | B}] *) +Lemma exists_to_inhabited_sig {A P} : (exists x : A, P x) -> inhabited {x : A | P x}. +Proof. + intros [x y]. exact (inhabits (exist _ x y)). +Qed. + +Lemma inhabited_sig_to_exists {A P} : inhabited {x : A | P x} -> exists x : A, P x. +Proof. + intros [[x y]];exists x;exact y. +Qed. + (** [sumbool] is a boolean type equipped with the justification of their value *) diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v index 07e8b6ef8..f1f20606b 100644 --- a/theories/Logic/ChoiceFacts.v +++ b/theories/Logic/ChoiceFacts.v @@ -760,6 +760,39 @@ Proof. destruct Heq using eq_indd; trivial. Qed. +(** Functional choice can be reformulated as a property on [inhabited] *) + +Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) := + (forall x, inhabited (B x)) -> inhabited (forall x, B x). + +Notation InhabitedForallCommute := + (forall A (B : A -> Type), InhabitedForallCommute_on B). + +Theorem functional_choice_to_inhabited_forall_commute : + FunctionalChoice -> InhabitedForallCommute. +Proof. + intros choose0 A B Hinhab. + pose proof (non_dep_dep_functional_choice choose0) as choose;clear choose0. + assert (Hexists : forall x, exists _ : B x, True). + { intros x;apply inhabited_sig_to_exists. + refine (inhabited_covariant _ (Hinhab x)). + intros y;exists y;exact I. } + apply choose in Hexists. + destruct Hexists as [f _]. + exact (inhabits f). +Qed. + +Theorem inhabited_forall_commute_to_functional_choice : + InhabitedForallCommute -> FunctionalChoice. +Proof. + intros choose A B R Hexists. + assert (Hinhab : forall x, inhabited {y : B | R x y}). + { intros x;apply exists_to_inhabited_sig;trivial. } + apply choose in Hinhab. destruct Hinhab as [f]. + exists (fun x => proj1_sig (f x)). + exact (fun x => proj2_sig (f x)). +Qed. + (** ** Reification of dependent and non dependent functional relation are equivalent *) Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) := -- cgit v1.2.3 -- cgit v1.2.3 From 540588e7c1e7edc3ced91f2c5b544025da0681f5 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 14:15:55 +0200 Subject: Avoiding registering files from _build_ci when not calling Makefile.ci. --- Makefile | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Makefile b/Makefile index e1d6e8e1d..826ed17b0 100644 --- a/Makefile +++ b/Makefile @@ -52,7 +52,8 @@ FIND_VCS_CLAUSE:='(' \ -name '.bzr' -o \ -name 'debian' -o \ -name "$${GIT_DIR}" -o \ - -name '_build' \ + -name '_build' -o \ + -name '_build_ci' \ ')' -prune -o define find -- cgit v1.2.3 From 5365971dfdf4136586527aa4f4c85fbfebeee0bd Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Mon, 1 May 2017 15:00:56 +0200 Subject: Fix for bug 5507. Mispelt de Bruijn. --- dev/doc/cic.dtd | 2 +- doc/refman/RefMan-tus.tex | 12 ++++++------ kernel/constr.ml | 2 +- kernel/constr.mli | 2 +- kernel/term.ml | 2 +- kernel/term.mli | 2 +- 6 files changed, 11 insertions(+), 11 deletions(-) diff --git a/dev/doc/cic.dtd b/dev/doc/cic.dtd index f2314e224..7a9d7eb1e 100644 --- a/dev/doc/cic.dtd +++ b/dev/doc/cic.dtd @@ -125,7 +125,7 @@ id ID #REQUIRED sort %sort; #REQUIRED> - + diff --git a/doc/refman/RefMan-tus.tex b/doc/refman/RefMan-tus.tex index 797b0aded..c55b3775c 100644 --- a/doc/refman/RefMan-tus.tex +++ b/doc/refman/RefMan-tus.tex @@ -299,7 +299,7 @@ vector of terms $vt$. \end{itemize} In this meta-language, bound variables are represented using the -so-called deBrujin's indexes. In this representation, an occurrence of +so-called deBruijn's indexes. In this representation, an occurrence of a bound variable is denoted by an integer, meaning the number of binders that must be traversed to reach its own binder\footnote{Actually, $(\texttt{Rel}\;n)$ means that $(n-1)$ binders @@ -482,7 +482,7 @@ following constructor functions: \begin{description} \fun{val Term.mkRel : int -> constr} - {$(\texttt{mkRel}\;n)$ represents deBrujin's index $n$.} + {$(\texttt{mkRel}\;n)$ represents deBruijn's index $n$.} \fun{val Term.mkVar : identifier -> constr} {$(\texttt{mkVar}\;id)$ @@ -545,7 +545,7 @@ following constructor functions: \fun{val Term.mkProd : name ->constr ->constr -> constr} {$(\texttt{mkProd}\;x\;A\;B)$ represents the product $(x:A)B$. - The free ocurrences of $x$ in $B$ are represented by deBrujin's + The free ocurrences of $x$ in $B$ are represented by deBruijn's indexes.} \fun{val Term.mkNamedProd : identifier -> constr -> constr -> constr} @@ -553,14 +553,14 @@ following constructor functions: but the bound occurrences of $x$ in $B$ are denoted by the identifier $(\texttt{mkVar}\;x)$. The function automatically changes each occurrences of this identifier into the corresponding - deBrujin's index.} + deBruijn's index.} \fun{val Term.mkArrow : constr -> constr -> constr} {$(\texttt{arrow}\;A\;B)$ represents the type $(A\rightarrow B)$.} \fun{val Term.mkLambda : name -> constr -> constr -> constr} {$(\texttt{mkLambda}\;x\;A\;b)$ represents the lambda abstraction - $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by deBrujin's + $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by deBruijn's indexes.} \fun{val Term.mkNamedLambda : identifier -> constr -> constr -> constr} @@ -666,7 +666,7 @@ use the primitive \textsl{Case} described in Chapter \ref{Cic} \item Restoring type coercions and synthesizing the implicit arguments (the one denoted by question marks in {\Coq} syntax: see Section~\ref{Coercions}). -\item Transforming the named bound variables into deBrujin's indexes. +\item Transforming the named bound variables into deBruijn's indexes. \item Classifying the global names into the different classes of constants (defined constants, constructors, inductive types, etc). \end{enumerate} diff --git a/kernel/constr.ml b/kernel/constr.ml index e2630c3f0..49e17d74d 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -124,7 +124,7 @@ type types = constr (* Term constructors *) (*********************) -(* Constructs a DeBrujin index with number n *) +(* Constructs a DeBruijn index with number n *) let rels = [|Rel 1;Rel 2;Rel 3;Rel 4;Rel 5;Rel 6;Rel 7; Rel 8; Rel 9;Rel 10;Rel 11;Rel 12;Rel 13;Rel 14;Rel 15; Rel 16|] diff --git a/kernel/constr.mli b/kernel/constr.mli index 700c235e6..d90afea8f 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -70,7 +70,7 @@ type types = constr (** {6 Term constructors. } *) -(** Constructs a DeBrujin index (DB indices begin at 1) *) +(** Constructs a DeBruijn index (DB indices begin at 1) *) val mkRel : int -> constr (** Constructs a Variable *) diff --git a/kernel/term.ml b/kernel/term.ml index e5a681375..7c6136f08 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -169,7 +169,7 @@ let hcons_types = Constr.hcons exception DestKO -(* Destructs a DeBrujin index *) +(* Destructs a DeBruijn index *) let destRel c = match kind_of_term c with | Rel n -> n | _ -> raise DestKO diff --git a/kernel/term.mli b/kernel/term.mli index a9bb67705..623d2c8a6 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -127,7 +127,7 @@ val is_small : sorts -> bool exception DestKO -(** Destructs a DeBrujin index *) +(** Destructs a DeBruijn index *) val destRel : constr -> int (** Destructs an existential variable *) -- cgit v1.2.3 From 12f1c409daf2cdbd7d0323f0d61723819532b362 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 16:56:25 +0200 Subject: Really fixing #2602 which was wrongly working because of #5487 hiding the cause. The cause was a missing evar/evar clause in ltac pattern-matching function (constr_matching.ml). --- pretyping/constr_matching.ml | 3 +++ test-suite/success/ltac.v | 13 +++++++++++++ 2 files changed, 16 insertions(+) diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 3c47cfdc4..afdf601c2 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -347,6 +347,9 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst | PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst + | PEvar (c1,args1), Evar (c2,args2) when c1 = c2 -> + Array.fold_left2 (sorec ctx env) subst args1 args2 + | _ -> raise PatternMatchingFailure in diff --git a/test-suite/success/ltac.v b/test-suite/success/ltac.v index ce9099059..d7ec092d7 100644 --- a/test-suite/success/ltac.v +++ b/test-suite/success/ltac.v @@ -317,3 +317,16 @@ let T := constr:(fun a b : nat => a) in end. exact (eq_refl n). Qed. + +(* A variant of #2602 which was wrongly succeeding because "a", bound to + "?m", was then internally turned into a "_" in the second matching *) + +Goal exists m, S m > 0. +eexists. +Fail match goal with + | |- context [ S ?a ] => + match goal with + | |- S a > a => idtac + end +end. +Abort. -- cgit v1.2.3 From c3aec655a8a33fff676c79e12888d193cc2e237b Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 16:58:38 +0200 Subject: Fixing #5487 (v8.5 regression on ltac-matching expressions with evars). The fix follows an invariant enforced in proofview.ml on the kind of evars that are goals or that occur in goals. One day, evar kinds will need a little cleaning... PS: This is a second attempt, completing db28e82 which was missing the case PEvar in constr_matching.ml. Indeed the attached fix to #5487 alone made #2602 failing, revealing that the real cause for #2602 was actually not fixed and that if the test for #2602 was working it was because of #5487 hiding the real problem in #2602. --- pretyping/patternops.ml | 4 +++- test-suite/bugs/closed/5487.v | 9 +++++++++ 2 files changed, 12 insertions(+), 1 deletion(-) create mode 100644 test-suite/bugs/closed/5487.v diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 2090aad8a..75d3ed30b 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -160,7 +160,9 @@ let pattern_of_constr env sigma t = let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in let () = ignore (pattern_of_constr env ty) in assert (not b); PMeta (Some id) - | Evar_kinds.GoalEvar -> + | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ -> + (* These are the two evar kinds used for existing goals *) + (* see Proofview.mark_in_evm *) PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in diff --git a/test-suite/bugs/closed/5487.v b/test-suite/bugs/closed/5487.v new file mode 100644 index 000000000..9b995f450 --- /dev/null +++ b/test-suite/bugs/closed/5487.v @@ -0,0 +1,9 @@ +(* Was a collision between an ltac pattern variable and an evar *) + +Goal forall n, exists m, n = m :> nat. +Proof. + eexists. + Fail match goal with + | [ |- ?x = ?y ] + => match x with y => idtac end + end. -- cgit v1.2.3 From 2553e4bf5735a2bd127832e2d26609c6a8096fb7 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 1 May 2017 17:24:29 +0200 Subject: Removing dead code in Autorewrite. Since 260965d, an imperative code was semantically the identity because the closure allocation was not performed at the right moment. Because of it intricacy, I cannot really tell the original motivations of this piece of code, although it looks like it was for there for pretty-printing of errors. Anyway, both because the code was dubious and its effect not observed, it cannot hurt to remove it. --- tactics/autorewrite.ml | 44 ++++++-------------------------------------- 1 file changed, 6 insertions(+), 38 deletions(-) diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index e58ec5a31..3c430cb17 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -127,45 +127,13 @@ let autorewrite ?(conds=Naive) tac_main lbas = (Proofview.tclUNIT()) lbas)) let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> (* let's check at once if id exists (to raise the appropriate error) *) - let _ = List.map (fun id -> Tacmach.New.pf_get_hyp id gl) idl in - let general_rewrite_in id = - let id = ref id in - let to_be_cleared = ref false in - fun dir cstr tac gl -> - let last_hyp_id = - match Tacmach.pf_hyps gl with - d :: _ -> Context.Named.Declaration.get_id d - | _ -> (* even the hypothesis id is missing *) - raise (Logic.RefinerError (Logic.NoSuchHyp !id)) - in - let gl' = Proofview.V82.of_tactic (general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false !id cstr false) gl in - let gls = gl'.Evd.it in - match gls with - g::_ -> - (match Environ.named_context_of_val (Goal.V82.hyps gl'.Evd.sigma g) with - d ::_ -> - let lastid = Context.Named.Declaration.get_id d in - if not (Id.equal last_hyp_id lastid) then - begin - let gl'' = - if !to_be_cleared then - tclTHEN (fun _ -> gl') (tclTRY (Proofview.V82.of_tactic (clear [!id]))) gl - else gl' in - id := lastid ; - to_be_cleared := true ; - gl'' - end - else - begin - to_be_cleared := false ; - gl' - end - | _ -> assert false) (* there must be at least an hypothesis *) - | _ -> assert false (* rewriting cannot complete a proof *) - in - let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y (EConstr.of_constr z) w) in + let _ = List.map (fun id -> Tacmach.New.pf_get_hyp id gl) idl in + let general_rewrite_in id dir cstr tac = + let cstr = EConstr.of_constr cstr in + general_rewrite_in dir AllOccurrences true ~tac:(tac, conds) false id cstr false + in Tacticals.New.tclMAP (fun id -> Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS (List.fold_left (fun tac bas -> -- cgit v1.2.3 From f6856c5022ef27cdc492daadd1301cfcad025b01 Mon Sep 17 00:00:00 2001 From: Paul Steckler Date: Mon, 1 May 2017 11:34:00 -0400 Subject: remove unneeded -emacs flag to coq-prog-args --- test-suite/bugs/closed/3080.v | 2 +- test-suite/bugs/closed/3323.v | 2 +- test-suite/bugs/closed/3332.v | 2 +- test-suite/bugs/closed/3346.v | 2 +- test-suite/bugs/closed/3348.v | 2 +- test-suite/bugs/closed/3352.v | 2 +- test-suite/bugs/closed/3375.v | 2 +- test-suite/bugs/closed/3393.v | 2 +- test-suite/bugs/closed/3427.v | 2 +- test-suite/bugs/closed/3539.v | 2 +- test-suite/bugs/closed/3612.v | 2 +- test-suite/bugs/closed/3649.v | 2 +- test-suite/bugs/closed/3881.v | 2 +- test-suite/bugs/closed/3956.v | 2 +- test-suite/bugs/closed/4089.v | 2 +- test-suite/bugs/closed/4121.v | 2 +- test-suite/bugs/closed/4394.v | 2 +- test-suite/bugs/closed/4400.v | 2 +- test-suite/bugs/closed/4456.v | 2 +- test-suite/bugs/closed/4527.v | 2 +- test-suite/bugs/closed/4533.v | 2 +- test-suite/bugs/closed/4544.v | 2 +- test-suite/bugs/closed/4656.v | 2 +- test-suite/bugs/closed/4673.v | 2 +- test-suite/bugs/closed/4722.v | 2 +- test-suite/bugs/closed/4727.v | 2 +- test-suite/bugs/closed/4733.v | 2 +- test-suite/bugs/closed/4769.v | 2 +- test-suite/bugs/closed/4780.v | 2 +- test-suite/bugs/closed/4785_compat_85.v | 2 +- test-suite/bugs/closed/4811.v | 2 +- test-suite/bugs/closed/4818.v | 2 +- test-suite/bugs/closed/5198.v | 2 +- test-suite/bugs/closed/HoTT_coq_012.v | 2 +- test-suite/bugs/closed/HoTT_coq_032.v | 2 +- test-suite/bugs/closed/HoTT_coq_053.v | 2 +- test-suite/bugs/closed/HoTT_coq_054.v | 1 - test-suite/bugs/closed/HoTT_coq_055.v | 2 +- test-suite/bugs/closed/HoTT_coq_059.v | 2 +- test-suite/bugs/closed/HoTT_coq_062.v | 2 +- test-suite/bugs/closed/HoTT_coq_097.v | 2 +- test-suite/bugs/closed/HoTT_coq_107.v | 2 +- test-suite/bugs/closed/HoTT_coq_108.v | 2 +- test-suite/bugs/closed/HoTT_coq_123.v | 2 +- test-suite/bugs/opened/4803.v | 2 +- test-suite/output-modulo-time/ltacprof.v | 2 +- test-suite/output-modulo-time/ltacprof_cutoff.v | 2 +- test-suite/output/ErrorInModule.v | 2 +- test-suite/output/ErrorInSection.v | 2 +- test-suite/success/Compat84.v | 2 +- 50 files changed, 49 insertions(+), 50 deletions(-) diff --git a/test-suite/bugs/closed/3080.v b/test-suite/bugs/closed/3080.v index 7d0dc090e..36ab7ff59 100644 --- a/test-suite/bugs/closed/3080.v +++ b/test-suite/bugs/closed/3080.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-nois") -*- *) +(* -*- coq-prog-args: ("-nois") -*- *) Delimit Scope type_scope with type. Delimit Scope function_scope with function. diff --git a/test-suite/bugs/closed/3323.v b/test-suite/bugs/closed/3323.v index 22b1603b6..4622634ea 100644 --- a/test-suite/bugs/closed/3323.v +++ b/test-suite/bugs/closed/3323.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 5302 lines to 4649 lines, then from 4660 lines to 355 lines, then from 360 lines to 269 lines, then from 269 lines to 175 lines, then from 144 lines to 119 lines, then from 297 lines to 117 lines, then from 95 lines to 79 lines, then from 82 lines to 68 lines *) Set Universe Polymorphism. diff --git a/test-suite/bugs/closed/3332.v b/test-suite/bugs/closed/3332.v index d86470cde..a3564bfcc 100644 --- a/test-suite/bugs/closed/3332.v +++ b/test-suite/bugs/closed/3332.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-time") -*- *) +(* -*- coq-prog-args: ("-time") -*- *) Definition foo : True. Proof. Abort. (* Toplevel input, characters 15-21: diff --git a/test-suite/bugs/closed/3346.v b/test-suite/bugs/closed/3346.v index 638404f2c..09bd78934 100644 --- a/test-suite/bugs/closed/3346.v +++ b/test-suite/bugs/closed/3346.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Monomorphic Inductive paths (A : Type) (a : A) : A -> Type := idpath : paths A a a. (* This should fail with -indices-matter *) Fail Check paths nat O O : Prop. diff --git a/test-suite/bugs/closed/3348.v b/test-suite/bugs/closed/3348.v index d9ac09d8d..904de6896 100644 --- a/test-suite/bugs/closed/3348.v +++ b/test-suite/bugs/closed/3348.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Set Universe Polymorphism. Set Printing Universes. Inductive Empty : Set := . diff --git a/test-suite/bugs/closed/3352.v b/test-suite/bugs/closed/3352.v index f8113e4c7..555accfd5 100644 --- a/test-suite/bugs/closed/3352.v +++ b/test-suite/bugs/closed/3352.v @@ -20,7 +20,7 @@ while it is expected to have type "IsHProp (* Top.17 *) Empty" Defined. Module B. -(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 7725 lines to 78 lines, then from 51 lines to 13 lines *) Set Universe Polymorphism. Inductive paths {A} (a : A) : A -> Type := idpath : paths a a where "x = y" := (@paths _ x y) : type_scope. diff --git a/test-suite/bugs/closed/3375.v b/test-suite/bugs/closed/3375.v index d7ce02ea8..1e0c8e61f 100644 --- a/test-suite/bugs/closed/3375.v +++ b/test-suite/bugs/closed/3375.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-impredicative-set") -*- *) +(* -*- mode: coq; coq-prog-args: ("-impredicative-set") -*- *) (* File reduced by coq-bug-finder from original input, then from 5968 lines to 11933 lines, then from 11239 lines to 11231 lines, then from 10365 lines to 446 lines, then from 456 lines to 379 lines, then from 391 lines to 373 lines, then from 369 lines to 351 lines, then from 350 lines to 340 lines, then from 348 lines to 320 lines, then from 328 lines to 302 lines, then from 330 lines to 45 lines *) Set Universe Polymorphism. diff --git a/test-suite/bugs/closed/3393.v b/test-suite/bugs/closed/3393.v index f7ab5f76a..ae8e41e29 100644 --- a/test-suite/bugs/closed/3393.v +++ b/test-suite/bugs/closed/3393.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 8760 lines to 7519 lines, then from 7050 lines to 909 lines, then from 846 lines to 578 lines, then from 497 lines to 392 lines, then from 365 lines to 322 lines, then from 252 lines to 205 lines, then from 215 lines to 204 lines, then from 210 lines to 182 lines, then from 175 lines to 157 lines *) Set Universe Polymorphism. Axiom admit : forall {T}, T. diff --git a/test-suite/bugs/closed/3427.v b/test-suite/bugs/closed/3427.v index 374a53928..9a57ca770 100644 --- a/test-suite/bugs/closed/3427.v +++ b/test-suite/bugs/closed/3427.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 0 lines to 7171 lines, then from 7184 lines to 558 lines, then from 556 lines to 209 lines *) Generalizable All Variables. Set Universe Polymorphism. diff --git a/test-suite/bugs/closed/3539.v b/test-suite/bugs/closed/3539.v index d258bb31f..b0c4b2370 100644 --- a/test-suite/bugs/closed/3539.v +++ b/test-suite/bugs/closed/3539.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 11678 lines to 11330 lines, then from 10721 lines to 9544 lines, then from 9549 lines to 794 lines, then from 810 lines to 785 lines, then from 628 lines to 246 lines, then from 220 lines to 89 lines, then from 80 lines to 47 lines *) (* coqc version trunk (August 2014) compiled on Aug 22 2014 4:17:28 with OCaml 4.01.0 coqtop version cagnode17:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (a67cc6941434124465f20b14a1256f2ede31a60e) *) diff --git a/test-suite/bugs/closed/3612.v b/test-suite/bugs/closed/3612.v index a54768507..90182a488 100644 --- a/test-suite/bugs/closed/3612.v +++ b/test-suite/bugs/closed/3612.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter" "-nois") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter" "-nois") -*- *) (* File reduced by coq-bug-finder from original input, then from 3595 lines to 3518 lines, then from 3133 lines to 2950 lines, then from 2911 lines to 415 lines, then from 431 lines to 407 \ lines, then from 421 lines to 428 lines, then from 444 lines to 429 lines, then from 434 lines to 66 lines, then from 163 lines to 48 lines *) (* coqc version trunk (September 2014) compiled on Sep 11 2014 14:48:8 with OCaml 4.01.0 diff --git a/test-suite/bugs/closed/3649.v b/test-suite/bugs/closed/3649.v index fc4c171e2..367d380ec 100644 --- a/test-suite/bugs/closed/3649.v +++ b/test-suite/bugs/closed/3649.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-nois") -*- *) +(* -*- coq-prog-args: ("-nois") -*- *) (* File reduced by coq-bug-finder from original input, then from 9518 lines to 404 lines, then from 410 lines to 208 lines, then from 162 lines to 77 lines *) (* coqc version trunk (September 2014) compiled on Sep 18 2014 21:0:5 with OCaml 4.01.0 coqtop version cagnode16:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (07e4438bd758c2ced8caf09a6961ccd77d84e42b) *) diff --git a/test-suite/bugs/closed/3881.v b/test-suite/bugs/closed/3881.v index a327bbf2a..bb6af6a66 100644 --- a/test-suite/bugs/closed/3881.v +++ b/test-suite/bugs/closed/3881.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-nois" "-R" "../theories" "Coq") -*- *) +(* -*- coq-prog-args: ("-nois" "-R" "../theories" "Coq") -*- *) (* File reduced by coq-bug-finder from original input, then from 2236 lines to 1877 lines, then from 1652 lines to 160 lines, then from 102 lines to 34 lines *) (* coqc version trunk (December 2014) compiled on Dec 23 2014 22:6:43 with OCaml 4.01.0 coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (90ed6636dea41486ddf2cc0daead83f9f0788163) *) diff --git a/test-suite/bugs/closed/3956.v b/test-suite/bugs/closed/3956.v index c19a2d4a0..66dee702a 100644 --- a/test-suite/bugs/closed/3956.v +++ b/test-suite/bugs/closed/3956.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter"); mode: visual-line -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter"); mode: visual-line -*- *) Set Universe Polymorphism. Set Primitive Projections. Close Scope nat_scope. diff --git a/test-suite/bugs/closed/4089.v b/test-suite/bugs/closed/4089.v index e4d76732a..fc1c504f1 100644 --- a/test-suite/bugs/closed/4089.v +++ b/test-suite/bugs/closed/4089.v @@ -1,6 +1,6 @@ Unset Strict Universe Declaration. Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from original input, then from 6522 lines to 318 lines, then from 1139 lines to 361 lines *) (* coqc version 8.5beta1 (February 2015) compiled on Feb 23 2015 18:32:3 with OCaml 4.01.0 coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (ebfc19d792492417b129063fb511aa423e9d9e08) *) diff --git a/test-suite/bugs/closed/4121.v b/test-suite/bugs/closed/4121.v index d34a2b8b1..a8bf950ff 100644 --- a/test-suite/bugs/closed/4121.v +++ b/test-suite/bugs/closed/4121.v @@ -1,5 +1,5 @@ Unset Strict Universe Declaration. -(* -*- coq-prog-args: ("-emacs" "-nois") -*- *) +(* -*- coq-prog-args: ("-nois") -*- *) (* File reduced by coq-bug-finder from original input, then from 830 lines to 47 lines, then from 25 lines to 11 lines *) (* coqc version 8.5beta1 (March 2015) compiled on Mar 11 2015 18:51:36 with OCaml 4.01.0 coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (8dbfee5c5f897af8186cb1bdfb04fd4f88eca677) *) diff --git a/test-suite/bugs/closed/4394.v b/test-suite/bugs/closed/4394.v index 60c935459..1ad81345d 100644 --- a/test-suite/bugs/closed/4394.v +++ b/test-suite/bugs/closed/4394.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) Require Import Equality List. Inductive Foo (I : Type -> Type) (A : Type) : Type := diff --git a/test-suite/bugs/closed/4400.v b/test-suite/bugs/closed/4400.v index 5c23f8404..a89cf0cbc 100644 --- a/test-suite/bugs/closed/4400.v +++ b/test-suite/bugs/closed/4400.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-require" "Coq.Compat.Coq84" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-require" "Coq.Compat.Coq84" "-compat" "8.4") -*- *) Require Import Coq.Lists.List Coq.Logic.JMeq Program.Equality. Set Printing Universes. Inductive Foo (I : Type -> Type) (A : Type) : Type := diff --git a/test-suite/bugs/closed/4456.v b/test-suite/bugs/closed/4456.v index a32acf789..56a7b4f6e 100644 --- a/test-suite/bugs/closed/4456.v +++ b/test-suite/bugs/closed/4456.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Fiat" "-top" "BooleanRecognizerMin" "-R" "." "Top") -*- *) +(* -*- mode: coq; coq-prog-args: ("-R" "." "Fiat" "-top" "BooleanRecognizerMin" "-R" "." "Top") -*- *) (* File reduced by coq-bug-finder from original input, then from 2475 lines to 729 lines, then from 746 lines to 658 lines, then from 675 lines to 658 lines *) (* coqc version 8.5beta3 (November 2015) compiled on Nov 11 2015 18:23:0 with OCaml 4.01.0 coqtop version 8.5beta3 (November 2015) *) diff --git a/test-suite/bugs/closed/4527.v b/test-suite/bugs/closed/4527.v index 08628377f..4fab9d44f 100644 --- a/test-suite/bugs/closed/4527.v +++ b/test-suite/bugs/closed/4527.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_bad_univ_length_01") -*- *) +(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_bad_univ_length_01") -*- *) (* File reduced by coq-bug-finder from original input, then from 1199 lines to 430 lines, then from 444 lines to 430 lines, then from 964 lines to 255 lines, then from 269 lines to 255 lines *) diff --git a/test-suite/bugs/closed/4533.v b/test-suite/bugs/closed/4533.v index ae17fb145..ccef1c304 100644 --- a/test-suite/bugs/closed/4533.v +++ b/test-suite/bugs/closed/4533.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_lex_wrong_rewrite_02") -*- *) +(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_lex_wrong_rewrite_02") -*- *) (* File reduced by coq-bug-finder from original input, then from 1125 lines to 346 lines, then from 360 lines to 346 lines, then from 822 lines to 271 lines, then from 285 lines to 271 lines *) diff --git a/test-suite/bugs/closed/4544.v b/test-suite/bugs/closed/4544.v index da140c931..82f1e3fe7 100644 --- a/test-suite/bugs/closed/4544.v +++ b/test-suite/bugs/closed/4544.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_oog_looping_rewrite_01") -*- *) +(* -*- mode: coq; coq-prog-args: ("-nois" "-indices-matter" "-R" "." "Top" "-top" "bug_oog_looping_rewrite_01") -*- *) (* File reduced by coq-bug-finder from original input, then from 2553 lines to 1932 lines, then from 1946 lines to 1932 lines, then from 2467 lines to 1002 lines, then from 1016 lines to 1002 lines *) (* coqc version 8.5 (January 2016) compiled on Jan 23 2016 16:15:22 with OCaml 4.01.0 coqtop version 8.5 (January 2016) *) diff --git a/test-suite/bugs/closed/4656.v b/test-suite/bugs/closed/4656.v index c89a86d63..a59eed2c8 100644 --- a/test-suite/bugs/closed/4656.v +++ b/test-suite/bugs/closed/4656.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) Goal True. constructor 1. Qed. diff --git a/test-suite/bugs/closed/4673.v b/test-suite/bugs/closed/4673.v index 1ae508185..10e48db6d 100644 --- a/test-suite/bugs/closed/4673.v +++ b/test-suite/bugs/closed/4673.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Fiat" "-top" "BooleanRecognizerOptimized" "-R" "." "Top") -*- *) +(* -*- mode: coq; coq-prog-args: ("-R" "." "Fiat" "-top" "BooleanRecognizerOptimized" "-R" "." "Top") -*- *) (* File reduced by coq-bug-finder from original input, then from 2407 lines to 22 lines, then from 528 lines to 35 lines, then from 331 lines to 42 lines, then from 56 lines to 42 lines, then from 63 lines to 46 lines, then from 60 lines to 46 lines *) (* coqc version 8.5 (February 2016) compiled on Feb 21 2016 15:26:16 with OCaml 4.02.3 coqtop version 8.5 (February 2016) *) Axiom proof_admitted : False. diff --git a/test-suite/bugs/closed/4722.v b/test-suite/bugs/closed/4722.v index f047624c8..2d41828f1 100644 --- a/test-suite/bugs/closed/4722.v +++ b/test-suite/bugs/closed/4722.v @@ -1 +1 @@ -(* -*- coq-prog-args: ("-emacs" "-R" "4722" "Foo") -*- *) +(* -*- coq-prog-args: ("-R" "4722" "Foo") -*- *) diff --git a/test-suite/bugs/closed/4727.v b/test-suite/bugs/closed/4727.v index 3854bbffd..cfb4548d2 100644 --- a/test-suite/bugs/closed/4727.v +++ b/test-suite/bugs/closed/4727.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) Goal forall (P : Set) (l : P) (P0 : Set) (w w0 : P0) (T : Type) (a : P * T) (o : P -> option P0), (forall (l1 l2 : P) (w1 : P0), o l1 = Some w1 -> o l2 = Some w1 -> l1 = l2) -> o l = Some w -> o (fst a) = Some w0 -> {w = w0} + {w <> w0} -> False. diff --git a/test-suite/bugs/closed/4733.v b/test-suite/bugs/closed/4733.v index a6ebda61c..a90abd71c 100644 --- a/test-suite/bugs/closed/4733.v +++ b/test-suite/bugs/closed/4733.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) (*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*) Require Import Coq.Lists.List. Require Import Coq.Vectors.Vector. diff --git a/test-suite/bugs/closed/4769.v b/test-suite/bugs/closed/4769.v index d87906f31..33a1d1a50 100644 --- a/test-suite/bugs/closed/4769.v +++ b/test-suite/bugs/closed/4769.v @@ -1,5 +1,5 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-R" "." "Top" "-top" "bug_hom_anom_10") -*- *) +(* -*- mode: coq; coq-prog-args: ("-nois" "-R" "." "Top" "-top" "bug_hom_anom_10") -*- *) (* File reduced by coq-bug-finder from original input, then from 156 lines to 41 lines, then from 237 lines to 45 lines, then from 163 lines to 66 lines, then from 342 lines to 121 lines, then from 353 lines to 184 lines, then from 343 lines to 255 lines, then from 435 lines to 322 lines, then from 475 lines to 351 lines, then from 442 lines to 377 lines, then from 505 lines to 410 lines, then from 591 lines to 481 lines, then from 596 lines to 535 lines, then from 647 lines to 570 lines, then from 669 lines to 596 lines, then from 687 lines to 620 lines, then from 728 lines to 652 lines, then from 1384 lines to 683 lines, then from 984 lines to 707 lines, then from 1124 lines to 734 lines, then from 775 lines to 738 lines, then from 950 lines to 763 lines, then from 857 lines to 798 lines, then from 983 lines to 752 lines, then from 1598 lines to 859 lines, then from 873 lines to 859 lines, then from 875 lines to 862 lines, then from 901 lines to 863 lines, then from 1047 lines to 865 lines, then from 929 lines to 871 lines, then from 989 lines to 884 lines, then from 900 lines to 884 lines, then from 884 lines to 751 lines, then from 763 lines to 593 lines, then from 482 lines to 232 lines, then from 416 lines to 227 lines, then from 290 lines to 231 lines, then from 348 lines to 235 lines, then from 249 lines to 235 lines, then from 249 lines to 172 lines, then from 186 lines to 172 lines, then from 140 lines to 113 lines, then from 127 lines to 113 lines *) (* coqc version trunk (June 2016) compiled on Jun 2 2016 10:16:20 with OCaml 4.02.3 coqtop version trunk (June 2016) *) diff --git a/test-suite/bugs/closed/4780.v b/test-suite/bugs/closed/4780.v index 4cec43184..71a51c631 100644 --- a/test-suite/bugs/closed/4780.v +++ b/test-suite/bugs/closed/4780.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Top" "-top" "bug_bad_induction_01") -*- *) +(* -*- mode: coq; coq-prog-args: ("-R" "." "Top" "-top" "bug_bad_induction_01") -*- *) (* File reduced by coq-bug-finder from original input, then from 1889 lines to 144 lines, then from 158 lines to 144 lines *) (* coqc version 8.5pl1 (April 2016) compiled on Apr 18 2016 14:48:5 with OCaml 4.02.3 coqtop version 8.5pl1 (April 2016) *) diff --git a/test-suite/bugs/closed/4785_compat_85.v b/test-suite/bugs/closed/4785_compat_85.v index 9d65840ac..bbb34f465 100644 --- a/test-suite/bugs/closed/4785_compat_85.v +++ b/test-suite/bugs/closed/4785_compat_85.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.5") -*- *) +(* -*- coq-prog-args: ("-compat" "8.5") -*- *) Require Coq.Lists.List Coq.Vectors.Vector. Require Coq.Compat.Coq85. diff --git a/test-suite/bugs/closed/4811.v b/test-suite/bugs/closed/4811.v index a91496262..fe6e65a0f 100644 --- a/test-suite/bugs/closed/4811.v +++ b/test-suite/bugs/closed/4811.v @@ -2,7 +2,7 @@ (* Submitted by Jason Gross *) -(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "src" "Crypto" "-R" "Bedrock" "Bedrock" "-R" "coqprime-8.5/Coqprime" "Coqprime" "-top" "GF255192") -*- *) +(* -*- mode: coq; coq-prog-args: ("-R" "src" "Crypto" "-R" "Bedrock" "Bedrock" "-R" "coqprime-8.5/Coqprime" "Coqprime" "-top" "GF255192") -*- *) (* File reduced by coq-bug-finder from original input, then from 162 lines to 23 lines, then from 245 lines to 95 lines, then from 198 lines to 101 lines, then from 654 lines to 452 lines, then from 591 lines to 505 lines, then from 1770 lines to 580 lines, then from 2238 lines to 1715 lines, then from 1776 lines to 1738 lines, then from 1750 lines to 1679 lines, then from 1693 lines to 1679 lines *) (* coqc version 8.5pl1 (April 2016) compiled on Apr 18 2016 14:48:5 with OCaml 4.02.3 coqtop version 8.5pl1 (April 2016) *) diff --git a/test-suite/bugs/closed/4818.v b/test-suite/bugs/closed/4818.v index 904abb228..e411ce62f 100644 --- a/test-suite/bugs/closed/4818.v +++ b/test-suite/bugs/closed/4818.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-R" "." "Prob" "-top" "Product") -*- *) +(* -*- mode: coq; coq-prog-args: ("-R" "." "Prob" "-top" "Product") -*- *) (* File reduced by coq-bug-finder from original input, then from 391 lines to 77 lines, then from 857 lines to 119 lines, then from 1584 lines to 126 lines, then from 362 lines to 135 lines, then from 149 lines to 135 lines *) (* coqc version 8.5pl1 (June 2016) compiled on Jun 9 2016 17:27:17 with OCaml 4.02.3 coqtop version 8.5pl1 (June 2016) *) diff --git a/test-suite/bugs/closed/5198.v b/test-suite/bugs/closed/5198.v index 7254afb42..72722f5f6 100644 --- a/test-suite/bugs/closed/5198.v +++ b/test-suite/bugs/closed/5198.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-boot" "-nois") -*- *) +(* -*- mode: coq; coq-prog-args: ("-boot" "-nois") -*- *) (* File reduced by coq-bug-finder from original input, then from 286 lines to 27 lines, then from 224 lines to 53 lines, then from 218 lines to 56 lines, then from 269 lines to 180 lines, then from 132 lines to 48 lines, then from diff --git a/test-suite/bugs/closed/HoTT_coq_012.v b/test-suite/bugs/closed/HoTT_coq_012.v index a3c697f8c..420aaf974 100644 --- a/test-suite/bugs/closed/HoTT_coq_012.v +++ b/test-suite/bugs/closed/HoTT_coq_012.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Definition UU := Type. Inductive toto (B : UU) : UU := c (x : B). diff --git a/test-suite/bugs/closed/HoTT_coq_032.v b/test-suite/bugs/closed/HoTT_coq_032.v index 3f5d7b021..39a7103d1 100644 --- a/test-suite/bugs/closed/HoTT_coq_032.v +++ b/test-suite/bugs/closed/HoTT_coq_032.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-xml") -*- *) +(* -*- mode: coq; coq-prog-args: ("-xml") -*- *) Set Implicit Arguments. Generalizable All Variables. Set Asymmetric Patterns. diff --git a/test-suite/bugs/closed/HoTT_coq_053.v b/test-suite/bugs/closed/HoTT_coq_053.v index e2bf1dbed..263dec485 100644 --- a/test-suite/bugs/closed/HoTT_coq_053.v +++ b/test-suite/bugs/closed/HoTT_coq_053.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Set Printing Universes. Set Implicit Arguments. Generalizable All Variables. diff --git a/test-suite/bugs/closed/HoTT_coq_054.v b/test-suite/bugs/closed/HoTT_coq_054.v index c68796593..635024e98 100644 --- a/test-suite/bugs/closed/HoTT_coq_054.v +++ b/test-suite/bugs/closed/HoTT_coq_054.v @@ -1,4 +1,3 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs") -*- *) Inductive Empty : Prop := . Inductive paths {A : Type} (a : A) : A -> Type := diff --git a/test-suite/bugs/closed/HoTT_coq_055.v b/test-suite/bugs/closed/HoTT_coq_055.v index a25098771..561b7e557 100644 --- a/test-suite/bugs/closed/HoTT_coq_055.v +++ b/test-suite/bugs/closed/HoTT_coq_055.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Set Universe Polymorphism. Inductive Empty : Prop := . diff --git a/test-suite/bugs/closed/HoTT_coq_059.v b/test-suite/bugs/closed/HoTT_coq_059.v index 9c7e394dc..2e6c735cf 100644 --- a/test-suite/bugs/closed/HoTT_coq_059.v +++ b/test-suite/bugs/closed/HoTT_coq_059.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Set Universe Polymorphism. Inductive eq {A} (x : A) : A -> Type := eq_refl : eq x x. diff --git a/test-suite/bugs/closed/HoTT_coq_062.v b/test-suite/bugs/closed/HoTT_coq_062.v index 90d1d1830..e01f73f1b 100644 --- a/test-suite/bugs/closed/HoTT_coq_062.v +++ b/test-suite/bugs/closed/HoTT_coq_062.v @@ -1,6 +1,6 @@ Unset Strict Universe Declaration. Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) (* File reduced by coq-bug-finder from 5012 lines to 4659 lines, then from 4220 lines to 501 lines, then from 513 lines to 495 lines, then from 513 lines to 495 lines, then from 412 lines to 79 lines, then from 412 lines to 79 lines. *) Set Universe Polymorphism. Definition admit {T} : T. diff --git a/test-suite/bugs/closed/HoTT_coq_097.v b/test-suite/bugs/closed/HoTT_coq_097.v index 38e8007b6..f2b2e57b9 100644 --- a/test-suite/bugs/closed/HoTT_coq_097.v +++ b/test-suite/bugs/closed/HoTT_coq_097.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) Set Universe Polymorphism. Set Printing Universes. Inductive Empty : Set := . diff --git a/test-suite/bugs/closed/HoTT_coq_107.v b/test-suite/bugs/closed/HoTT_coq_107.v index 7c1ab8dc2..fa4072a8f 100644 --- a/test-suite/bugs/closed/HoTT_coq_107.v +++ b/test-suite/bugs/closed/HoTT_coq_107.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-nois" "-R" "../theories" "Coq") -*- *) +(* -*- mode: coq; coq-prog-args: ("-nois" "-R" "../theories" "Coq") -*- *) (* File reduced by coq-bug-finder from 4897 lines to 2605 lines, then from 2297 lines to 236 lines, then from 239 lines to 137 lines, then from 118 lines to 67 lines, then from 520 lines to 76 lines. *) (** Note: The bug here is the same as the #113, that is, HoTT_coq_113.v *) Require Import Coq.Init.Logic. diff --git a/test-suite/bugs/closed/HoTT_coq_108.v b/test-suite/bugs/closed/HoTT_coq_108.v index b6c0da76b..ea4bcd8b4 100644 --- a/test-suite/bugs/closed/HoTT_coq_108.v +++ b/test-suite/bugs/closed/HoTT_coq_108.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") -*- *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") -*- *) (* NOTE: This bug is only triggered with -load-vernac-source / in interactive mode. *) (* File reduced by coq-bug-finder from 139 lines to 124 lines. *) Set Universe Polymorphism. diff --git a/test-suite/bugs/closed/HoTT_coq_123.v b/test-suite/bugs/closed/HoTT_coq_123.v index 6ee6e6532..cd9cad406 100644 --- a/test-suite/bugs/closed/HoTT_coq_123.v +++ b/test-suite/bugs/closed/HoTT_coq_123.v @@ -1,5 +1,5 @@ Require Import TestSuite.admit. -(* -*- mode: coq; coq-prog-args: ("-emacs" "-indices-matter") *) +(* -*- mode: coq; coq-prog-args: ("-indices-matter") *) (* File reduced by coq-bug-finder from original input, then from 4988 lines to 856 lines, then from 648 lines to 398 lines, then from 401 lines to 332 lines, then from 287 lines to 250 lines, then from 257 lines to 241 lines, then from 223 lines to 175 lines *) Set Universe Polymorphism. Set Asymmetric Patterns. diff --git a/test-suite/bugs/opened/4803.v b/test-suite/bugs/opened/4803.v index 4530548b8..4541f13d0 100644 --- a/test-suite/bugs/opened/4803.v +++ b/test-suite/bugs/opened/4803.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) (*Suppose a user wants to declare a new list-like notation with support for singletons in both 8.4 and 8.5. If they use*) Require Import Coq.Lists.List. Require Import Coq.Vectors.Vector. diff --git a/test-suite/output-modulo-time/ltacprof.v b/test-suite/output-modulo-time/ltacprof.v index 6611db70e..1e9e91979 100644 --- a/test-suite/output-modulo-time/ltacprof.v +++ b/test-suite/output-modulo-time/ltacprof.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-profile-ltac-cutoff" "0.0") -*- *) +(* -*- coq-prog-args: ("-profile-ltac-cutoff" "0.0") -*- *) Ltac sleep' := do 100 (do 100 (do 100 idtac)). Ltac sleep := sleep'. diff --git a/test-suite/output-modulo-time/ltacprof_cutoff.v b/test-suite/output-modulo-time/ltacprof_cutoff.v index 50131470e..3dad6271a 100644 --- a/test-suite/output-modulo-time/ltacprof_cutoff.v +++ b/test-suite/output-modulo-time/ltacprof_cutoff.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-profile-ltac") -*- *) +(* -*- coq-prog-args: ("-profile-ltac") -*- *) Require Coq.ZArith.BinInt. Ltac sleep := do 50 (idtac; let sleep := (eval vm_compute in Coq.ZArith.BinInt.Z.div_eucl) in idtac). diff --git a/test-suite/output/ErrorInModule.v b/test-suite/output/ErrorInModule.v index e69e23276..b2e3c3e92 100644 --- a/test-suite/output/ErrorInModule.v +++ b/test-suite/output/ErrorInModule.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-quick") -*- *) +(* -*- mode: coq; coq-prog-args: ("-quick") -*- *) Module M. Definition foo := nonexistent. End M. diff --git a/test-suite/output/ErrorInSection.v b/test-suite/output/ErrorInSection.v index 3036f8f05..505c5ce37 100644 --- a/test-suite/output/ErrorInSection.v +++ b/test-suite/output/ErrorInSection.v @@ -1,4 +1,4 @@ -(* -*- mode: coq; coq-prog-args: ("-emacs" "-quick") -*- *) +(* -*- mode: coq; coq-prog-args: ("-quick") -*- *) Section S. Definition foo := nonexistent. End S. diff --git a/test-suite/success/Compat84.v b/test-suite/success/Compat84.v index db6348fa1..732a024fc 100644 --- a/test-suite/success/Compat84.v +++ b/test-suite/success/Compat84.v @@ -1,4 +1,4 @@ -(* -*- coq-prog-args: ("-emacs" "-compat" "8.4") -*- *) +(* -*- coq-prog-args: ("-compat" "8.4") -*- *) Goal True. solve [ constructor 1 ]. Undo. -- cgit v1.2.3 From d0252cac3167ef1e5cd26c1b9b40aea06d343413 Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Mon, 1 May 2017 17:48:57 +0200 Subject: More consistent writing of de Bruijn. --- checker/term.ml | 4 ++-- dev/doc/cic.dtd | 2 +- doc/RecTutorial/coqartmacros.tex | 2 +- doc/refman/RefMan-tus.tex | 18 +++++++++--------- engine/evarutil.ml | 2 +- kernel/constr.ml | 2 +- kernel/constr.mli | 2 +- kernel/nativeconv.ml | 2 +- kernel/term.ml | 2 +- kernel/term.mli | 2 +- kernel/vars.ml | 2 +- kernel/vars.mli | 4 ++-- tactics/dnet.mli | 2 +- vernac/auto_ind_decl.ml | 4 ++-- vernac/obligations.ml | 6 +++--- 15 files changed, 28 insertions(+), 28 deletions(-) diff --git a/checker/term.ml b/checker/term.ml index 591348cb6..24e6008d3 100644 --- a/checker/term.ml +++ b/checker/term.ml @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* This module instantiates the structure of generic deBruijn terms to Coq *) +(* This module instantiates the structure of generic de Bruijn terms to Coq *) open CErrors open Util @@ -94,7 +94,7 @@ let closedn n c = in try closed_rec n c; true with LocalOccur -> false -(* [closed0 M] is true iff [M] is a (deBruijn) closed term *) +(* [closed0 M] is true iff [M] is a (de Bruijn) closed term *) let closed0 = closedn 0 diff --git a/dev/doc/cic.dtd b/dev/doc/cic.dtd index 7a9d7eb1e..cc33efd48 100644 --- a/dev/doc/cic.dtd +++ b/dev/doc/cic.dtd @@ -125,7 +125,7 @@ id ID #REQUIRED sort %sort; #REQUIRED> - + diff --git a/doc/RecTutorial/coqartmacros.tex b/doc/RecTutorial/coqartmacros.tex index 2a2c21196..72d749269 100644 --- a/doc/RecTutorial/coqartmacros.tex +++ b/doc/RecTutorial/coqartmacros.tex @@ -149,7 +149,7 @@ \newcommand{\PicAbst}[3]{\begin{bundle}{\bf abst}\chunk{#1}\chunk{#2}\chunk{#3}% \end{bundle}} -% the same in DeBruijn form +% the same in de Bruijn form \newcommand{\PicDbj}[2]{\begin{bundle}{\bf abst}\chunk{#1}\chunk{#2} \end{bundle}} diff --git a/doc/refman/RefMan-tus.tex b/doc/refman/RefMan-tus.tex index c55b3775c..017de6d48 100644 --- a/doc/refman/RefMan-tus.tex +++ b/doc/refman/RefMan-tus.tex @@ -288,8 +288,8 @@ constructors: \item $(\texttt{VAR}\;id)$, a reference to a global identifier called $id$; \item $(\texttt{Rel}\;n)$, a bound variable, whose binder is the $nth$ binder up in the term; -\item $\texttt{DLAM}\;(x,t)$, a deBruijn's binder on the term $t$; -\item $\texttt{DLAMV}\;(x,vt)$, a deBruijn's binder on all the terms of +\item $\texttt{DLAM}\;(x,t)$, a de Bruijn's binder on the term $t$; +\item $\texttt{DLAMV}\;(x,vt)$, a de Bruijn's binder on all the terms of the vector $vt$; \item $(\texttt{DOP0}\;op)$, a unary operator $op$; \item $\texttt{DOP2}\;(op,t_1,t_2)$, the application of a binary @@ -299,7 +299,7 @@ vector of terms $vt$. \end{itemize} In this meta-language, bound variables are represented using the -so-called deBruijn's indexes. In this representation, an occurrence of +so-called de Bruijn's indexes. In this representation, an occurrence of a bound variable is denoted by an integer, meaning the number of binders that must be traversed to reach its own binder\footnote{Actually, $(\texttt{Rel}\;n)$ means that $(n-1)$ binders @@ -339,7 +339,7 @@ on the terms of the meta-language: \fun{val Generic.dependent : 'op term -> 'op term -> bool} {Returns true if the first term is a sub-term of the second.} %\fun{val Generic.subst\_var : identifier -> 'op term -> 'op term} -% { $(\texttt{subst\_var}\;id\;t)$ substitutes the deBruijn's index +% { $(\texttt{subst\_var}\;id\;t)$ substitutes the de Bruijn's index % associated to $id$ to every occurrence of the term % $(\texttt{VAR}\;id)$ in $t$.} \end{description} @@ -482,7 +482,7 @@ following constructor functions: \begin{description} \fun{val Term.mkRel : int -> constr} - {$(\texttt{mkRel}\;n)$ represents deBruijn's index $n$.} + {$(\texttt{mkRel}\;n)$ represents de Bruijn's index $n$.} \fun{val Term.mkVar : identifier -> constr} {$(\texttt{mkVar}\;id)$ @@ -545,7 +545,7 @@ following constructor functions: \fun{val Term.mkProd : name ->constr ->constr -> constr} {$(\texttt{mkProd}\;x\;A\;B)$ represents the product $(x:A)B$. - The free ocurrences of $x$ in $B$ are represented by deBruijn's + The free ocurrences of $x$ in $B$ are represented by de Bruijn's indexes.} \fun{val Term.mkNamedProd : identifier -> constr -> constr -> constr} @@ -553,14 +553,14 @@ following constructor functions: but the bound occurrences of $x$ in $B$ are denoted by the identifier $(\texttt{mkVar}\;x)$. The function automatically changes each occurrences of this identifier into the corresponding - deBruijn's index.} + de Bruijn's index.} \fun{val Term.mkArrow : constr -> constr -> constr} {$(\texttt{arrow}\;A\;B)$ represents the type $(A\rightarrow B)$.} \fun{val Term.mkLambda : name -> constr -> constr -> constr} {$(\texttt{mkLambda}\;x\;A\;b)$ represents the lambda abstraction - $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by deBruijn's + $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by de Bruijn's indexes.} \fun{val Term.mkNamedLambda : identifier -> constr -> constr -> constr} @@ -666,7 +666,7 @@ use the primitive \textsl{Case} described in Chapter \ref{Cic} \item Restoring type coercions and synthesizing the implicit arguments (the one denoted by question marks in {\Coq} syntax: see Section~\ref{Coercions}). -\item Transforming the named bound variables into deBruijn's indexes. +\item Transforming the named bound variables into de Bruijn's indexes. \item Classifying the global names into the different classes of constants (defined constants, constructors, inductive types, etc). \end{enumerate} diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 1624dc93e..22f93faa6 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -462,7 +462,7 @@ let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?nami ev (* This assumes an evar with identity instance and generalizes it over only - the De Bruijn part of the context *) + the de Bruijn part of the context *) let generalize_evar_over_rels sigma (ev,args) = let open EConstr in let evi = Evd.find sigma ev in diff --git a/kernel/constr.ml b/kernel/constr.ml index 49e17d74d..ae58fd068 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -124,7 +124,7 @@ type types = constr (* Term constructors *) (*********************) -(* Constructs a DeBruijn index with number n *) +(* Constructs a de Bruijn index with number n *) let rels = [|Rel 1;Rel 2;Rel 3;Rel 4;Rel 5;Rel 6;Rel 7; Rel 8; Rel 9;Rel 10;Rel 11;Rel 12;Rel 13;Rel 14;Rel 15; Rel 16|] diff --git a/kernel/constr.mli b/kernel/constr.mli index d90afea8f..e0954160f 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -70,7 +70,7 @@ type types = constr (** {6 Term constructors. } *) -(** Constructs a DeBruijn index (DB indices begin at 1) *) +(** Constructs a de Bruijn index (DB indices begin at 1) *) val mkRel : int -> constr (** Constructs a Variable *) diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml index 3c0afe380..3593d94c2 100644 --- a/kernel/nativeconv.ml +++ b/kernel/nativeconv.ml @@ -141,7 +141,7 @@ let native_conv_gen pb sigma env univs t1 t2 = let t1 = Sys.time () in let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in if !Flags.debug then Feedback.msg_debug (Pp.str time_info); - (* TODO change 0 when we can have deBruijn *) + (* TODO change 0 when we can have de Bruijn *) fst (conv_val env pb 0 !rt1 !rt2 univs) end | _ -> anomaly (Pp.str "Compilation failure") diff --git a/kernel/term.ml b/kernel/term.ml index 7c6136f08..03562d9f3 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -169,7 +169,7 @@ let hcons_types = Constr.hcons exception DestKO -(* Destructs a DeBruijn index *) +(* Destructs a de Bruijn index *) let destRel c = match kind_of_term c with | Rel n -> n | _ -> raise DestKO diff --git a/kernel/term.mli b/kernel/term.mli index 623d2c8a6..241ef322f 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -127,7 +127,7 @@ val is_small : sorts -> bool exception DestKO -(** Destructs a DeBruijn index *) +(** Destructs a de Bruijn index *) val destRel : constr -> int (** Destructs an existential variable *) diff --git a/kernel/vars.ml b/kernel/vars.ml index 4affb5f9f..f1c0a4f08 100644 --- a/kernel/vars.ml +++ b/kernel/vars.ml @@ -27,7 +27,7 @@ let closedn n c = in try closed_rec n c; true with LocalOccur -> false -(* [closed0 M] is true iff [M] is a (deBruijn) closed term *) +(* [closed0 M] is true iff [M] is a (de Bruijn) closed term *) let closed0 c = closedn 0 c diff --git a/kernel/vars.mli b/kernel/vars.mli index adeac422e..df5c55118 100644 --- a/kernel/vars.mli +++ b/kernel/vars.mli @@ -11,10 +11,10 @@ open Constr (** {6 Occur checks } *) -(** [closedn n M] is true iff [M] is a (deBruijn) closed term under n binders *) +(** [closedn n M] is true iff [M] is a (de Bruijn) closed term under n binders *) val closedn : int -> constr -> bool -(** [closed0 M] is true iff [M] is a (deBruijn) closed term *) +(** [closed0 M] is true iff [M] is a (de Bruijn) closed term *) val closed0 : constr -> bool (** [noccurn n M] returns true iff [Rel n] does NOT occur in term [M] *) diff --git a/tactics/dnet.mli b/tactics/dnet.mli index 9f29c60b6..565a916f8 100644 --- a/tactics/dnet.mli +++ b/tactics/dnet.mli @@ -26,7 +26,7 @@ distincts, or you'll get unexpected behaviours. Warning 2: This structure is perfect, i.e. the set of candidates - returned is equal to the set of solutions. Beware of DeBruijn + returned is equal to the set of solutions. Beware of de Bruijn shifts and sorts subtyping though (which makes the comparison not symmetric, see term_dnet.ml). diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index b9c4c6cc5..c91dcc505 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -95,7 +95,7 @@ let destruct_on_using c id = let destruct_on_as c l = destruct false None c (Some (dl,l)) None -(* reconstruct the inductive with the correct deBruijn indexes *) +(* reconstruct the inductive with the correct de Bruijn indexes *) let mkFullInd (ind,u) n = let mib = Global.lookup_mind (fst ind) in let nparams = mib.mind_nparams in @@ -174,7 +174,7 @@ let build_beq_scheme mode kn = (* give a type A, this function tries to find the equality on A declared previously *) (* nlist = the number of args (A , B , ... ) - eqA = the deBruijn index of the first eq param + eqA = the de Bruijn index of the first eq param ndx = how much to translate due to the 2nd Case *) let compute_A_equality rel_list nlist eqA ndx t = diff --git a/vernac/obligations.ml b/vernac/obligations.ml index ea2401b5c..e0520216b 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -8,7 +8,7 @@ open Declare (** - Get types of existentials ; - Flatten dependency tree (prefix order) ; - - Replace existentials by De Bruijn indices in term, applied to the right arguments ; + - Replace existentials by de Bruijn indices in term, applied to the right arguments ; - Apply term prefixed by quantification on "existentials". *) @@ -51,7 +51,7 @@ type oblinfo = ev_tac: unit Proofview.tactic option; ev_deps: Int.Set.t } -(** Substitute evar references in t using De Bruijn indices, +(** Substitute evar references in t using de Bruijn indices, where n binders were passed through. *) let subst_evar_constr evs n idf t = @@ -102,7 +102,7 @@ let subst_evar_constr evs n idf t = t', !seen, !transparent -(** Substitute variable references in t using De Bruijn indices, +(** Substitute variable references in t using de Bruijn indices, where n binders were passed through. *) let subst_vars acc n t = let var_index id = Util.List.index Id.equal id acc in -- cgit v1.2.3 From cea40f37ab638031b9d5c6434ee5651a16ea1f3e Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 09:04:17 +0200 Subject: Fixing Set Rewriting Schemes bugs introduced in v8.5. - Fixing a typo introduced in 31dbba5f. - Adapting to computation of universe constraints in pretyping. - Adding a regression test. --- tactics/eqschemes.ml | 8 +++++++- test-suite/success/Scheme.v | 23 +++++++++++++++++++++++ 2 files changed, 30 insertions(+), 1 deletion(-) diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 1a45217a4..e39159fb8 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -89,6 +89,11 @@ let get_coq_eq ctx = with Not_found -> error "eq not found." +let univ_of_eq env eq = + match kind_of_term (Retyping.get_type_of env Evd.empty eq) with + | Prod (_,t,_) -> (match kind_of_term t with Sort (Type u) -> u | _ -> assert false) + | _ -> assert false + (**********************************************************************) (* Check if an inductive type [ind] has the form *) (* *) @@ -744,7 +749,7 @@ let build_congr env (eq,refl,ctx) ind = let ty = get_type (lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity) in let constrsign,ccl = decompose_prod_assum mip.mind_nf_lc.(0) in let _,constrargs = decompose_app ccl in - if Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt) then + if not (Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt)) then error "Constructor must have no arguments"; let b = List.nth constrargs (i + mib.mind_nparams - 1) in let varB = fresh env (Id.of_string "B") in @@ -752,6 +757,7 @@ let build_congr env (eq,refl,ctx) ind = let varf = fresh env (Id.of_string "f") in let ci = make_case_info (Global.env()) ind RegularStyle in let uni, ctx = Universes.extend_context (Universes.new_global_univ ()) ctx in + let ctx = (fst ctx, Univ.enforce_leq uni (univ_of_eq env eq) (snd ctx)) in let c = my_it_mkLambda_or_LetIn paramsctxt (mkNamedLambda varB (mkSort (Type uni)) diff --git a/test-suite/success/Scheme.v b/test-suite/success/Scheme.v index dd5aa81d1..855f26698 100644 --- a/test-suite/success/Scheme.v +++ b/test-suite/success/Scheme.v @@ -2,3 +2,26 @@ Scheme Induction for eq Sort Prop. Check eq_ind_dep. + +(* This was broken in v8.5 *) + +Set Rewriting Schemes. +Inductive myeq A (a:A) : A -> Prop := myrefl : myeq A a a. +Unset Rewriting Schemes. + +Check myeq_rect. +Check myeq_ind. +Check myeq_rec. +Check myeq_congr. +Check myeq_sym_internal. +Check myeq_rew. +Check myeq_rew_dep. +Check myeq_rew_fwd_dep. +Check myeq_rew_r. +Check internal_myeq_sym_involutive. +Check myeq_rew_r_dep. +Check myeq_rew_fwd_r_dep. + +Set Rewriting Schemes. +Inductive myeq_true : bool -> Prop := myrefl_true : myeq_true true. +Unset Rewriting Schemes. -- cgit v1.2.3 From 8dc2adfd102185c6c3c4b61709f2b55aefab2641 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Sat, 29 Apr 2017 13:07:45 -0400 Subject: Add bmsherman/topology to the ci This development of @bmsherman tests universe polymorphism and setoid rewriting in type, and should build with v8.6 and trunk. --- .travis.yml | 1 + Makefile.ci | 2 +- dev/ci/ci-basic-overlay.sh | 6 ++++++ dev/ci/ci-formal-topology.sh | 28 ++++++++++++++++++++++++++++ 4 files changed, 36 insertions(+), 1 deletion(-) create mode 100755 dev/ci/ci-formal-topology.sh diff --git a/.travis.yml b/.travis.yml index 77aac23b7..3ed547bb1 100644 --- a/.travis.yml +++ b/.travis.yml @@ -41,6 +41,7 @@ env: - TEST_TARGET="ci-math-classes" - TEST_TARGET="ci-math-comp" - TEST_TARGET="ci-sf" + - TEST_TARGET="ci-formal-topology" - TEST_TARGET="ci-unimath" - TEST_TARGET="ci-vst" # Not ready yet for 8.7 diff --git a/Makefile.ci b/Makefile.ci index 4c4606aff..013685218 100644 --- a/Makefile.ci +++ b/Makefile.ci @@ -1,7 +1,7 @@ CI_TARGETS=ci-all ci-hott ci-math-comp ci-compcert ci-sf ci-cpdt \ ci-color ci-math-classes ci-tlc ci-fiat-crypto ci-fiat-parsers \ ci-coquelicot ci-flocq ci-iris-coq ci-metacoq ci-geocoq \ - ci-unimath ci-vst ci-bedrock-src ci-bedrock-facade + ci-unimath ci-vst ci-bedrock-src ci-bedrock-facade ci-formal-topology .PHONY: $(CI_TARGETS) diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh index e0851816c..5da13c100 100644 --- a/dev/ci/ci-basic-overlay.sh +++ b/dev/ci/ci-basic-overlay.sh @@ -106,6 +106,12 @@ : ${bedrock_facade_CI_BRANCH:=master} : ${bedrock_facade_CI_GITURL:=https://github.com/JasonGross/bedrock.git} +######################################################################## +# formal-topology +######################################################################## +: ${formal_topology_CI_BRANCH:=ci} +: ${formal_topology_CI_GITURL:=https://github.com/bmsherman/topology.git} + ######################################################################## # CoLoR ######################################################################## diff --git a/dev/ci/ci-formal-topology.sh b/dev/ci/ci-formal-topology.sh new file mode 100755 index 000000000..ecb36349f --- /dev/null +++ b/dev/ci/ci-formal-topology.sh @@ -0,0 +1,28 @@ +#!/usr/bin/env bash + +ci_dir="$(dirname "$0")" +source ${ci_dir}/ci-common.sh + +math_classes_CI_DIR=${CI_BUILD_DIR}/math-classes + +Corn_CI_DIR=${CI_BUILD_DIR}/corn + +formal_topology_CI_DIR=${CI_BUILD_DIR}/formal-topology + +# Setup Math-Classes + +git_checkout ${math_classes_CI_BRANCH} ${math_classes_CI_GITURL} ${math_classes_CI_DIR} + +( cd ${math_classes_CI_DIR} && make -j ${NJOBS} && make install ) + +# Setup Corn + +git_checkout ${Corn_CI_BRANCH} ${Corn_CI_GITURL} ${Corn_CI_DIR} + +( cd ${Corn_CI_DIR} && make -j ${NJOBS} && make install ) + +# Setup formal-topology + +git_checkout ${formal_topology_CI_BRANCH} ${formal_topology_CI_GITURL} ${formal_topology_CI_DIR} + +( cd ${formal_topology_CI_DIR} && make -j ${NJOBS} ) -- cgit v1.2.3 From 824caa1f93563ab9437fb238459d757447a0aa12 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 14:15:55 +0200 Subject: Avoiding registering files from _build_ci when not calling Makefile.ci. --- Makefile | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Makefile b/Makefile index e84d5e377..e50a1b18f 100644 --- a/Makefile +++ b/Makefile @@ -52,7 +52,8 @@ FIND_VCS_CLAUSE:='(' \ -name '.bzr' -o \ -name 'debian' -o \ -name "$${GIT_DIR}" -o \ - -name '_build' \ + -name '_build' -o \ + -name '_build_ci' \ ')' -prune -o define find -- cgit v1.2.3 From 874ff85e2d52b33010007bb3a1f1add9391b030f Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Tue, 2 May 2017 16:33:59 +0200 Subject: Remove unused module definition after merging PR#592. --- tactics/hints.ml | 2 -- 1 file changed, 2 deletions(-) diff --git a/tactics/hints.ml b/tactics/hints.ml index d57e4875c..c5bacc5a2 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -module CVars = Vars - open Pp open Util open CErrors -- cgit v1.2.3 From 6156fae91e69d379cfa0263a4c4cafb48da56f85 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Tue, 2 May 2017 16:09:56 +0200 Subject: Fix two new unused opens. --- vernac/lemmas.ml | 2 -- 1 file changed, 2 deletions(-) diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 1344701ff..b79795aeb 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -28,10 +28,8 @@ open Pretyping open Termops open Namegen open Reductionops -open Constrexpr open Constrintern open Impargs -open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -- cgit v1.2.3 From ba9806a242c7ece41a93181d1fb91279c89794ed Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Tue, 2 May 2017 16:10:19 +0200 Subject: Remove dead code in native compiler. --- kernel/nativecode.ml | 12 ------- kernel/nativelambda.ml | 93 -------------------------------------------------- 2 files changed, 105 deletions(-) diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index ba80ff590..5130aa9a4 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -42,8 +42,6 @@ module LNset = Set.Make(LNord) let lname_ctr = ref (-1) -let reset_lname = lname_ctr := -1 - let fresh_lname n = incr lname_ctr; { lname = n; luid = !lname_ctr } @@ -112,40 +110,30 @@ let gname_hash gn = match gn with let case_ctr = ref (-1) -let reset_gcase () = case_ctr := -1 - let fresh_gcase l = incr case_ctr; Gcase (l,!case_ctr) let pred_ctr = ref (-1) -let reset_gpred () = pred_ctr := -1 - let fresh_gpred l = incr pred_ctr; Gpred (l,!pred_ctr) let fixtype_ctr = ref (-1) -let reset_gfixtype () = fixtype_ctr := -1 - let fresh_gfixtype l = incr fixtype_ctr; Gfixtype (l,!fixtype_ctr) let norm_ctr = ref (-1) -let reset_norm () = norm_ctr := -1 - let fresh_gnorm l = incr norm_ctr; Gnorm (l,!norm_ctr) let normtbl_ctr = ref (-1) -let reset_normtbl () = normtbl_ctr := -1 - let fresh_gnormtbl l = incr normtbl_ctr; Gnormtbl (l,!normtbl_ctr) diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml index fcb75c661..72d9c4851 100644 --- a/kernel/nativelambda.ml +++ b/kernel/nativelambda.ml @@ -16,10 +16,6 @@ open Nativeinstr module RelDecl = Context.Rel.Declaration -(* I'm not messing with this stuff. *) -[@@@ocaml.warning "-32"] - -(** This file defines the lambda code generation phase of the native compiler *) exception NotClosed @@ -164,10 +160,6 @@ let rec lam_exsubst subst lam = | Lrel(id,i) -> lam_subst_rel lam id i subst | _ -> map_lam_with_binders liftn lam_exsubst subst lam -let lam_subst subst lam = - if is_subs_id subst then lam - else lam_exsubst subst lam - let lam_subst_args subst args = if is_subs_id subst then args else Array.smartmap (lam_exsubst subst) args @@ -281,71 +273,6 @@ and reduce_lapp substf lids body substa largs = Llam(Array.of_list lids, simplify (liftn (List.length lids) substf) body) | [], _::_ -> simplify_app substf body substa (Array.of_list largs) - -(* [occurrence kind k lam]: - If [kind] is [true] return [true] if the variable [k] does not appear in - [lam], return [false] if the variable appear one time and not - under a lambda, a fixpoint, a cofixpoint; else raise Not_found. - If [kind] is [false] return [false] if the variable does not appear in [lam] - else raise [Not_found] -*) - -let rec occurrence k kind lam = - match lam with - | Lrel (_,n) -> - if Int.equal n k then - if kind then false else raise Not_found - else kind - | Lvar _ | Lconst _ | Lproj _ | Luint _ | Lval _ | Lsort _ | Lind _ - | Lconstruct _ | Llazy | Lforce | Lmeta _ | Levar _ -> kind - | Lprod(dom, codom) -> - occurrence k (occurrence k kind dom) codom - | Llam(ids,body) -> - let _ = occurrence (k+Array.length ids) false body in kind - | Llet(_,def,body) -> - occurrence (k+1) (occurrence k kind def) body - | Lapp(f, args) -> - occurrence_args k (occurrence k kind f) args - | Lprim(_,_,_,args) | Lmakeblock(_,_,_,args) -> - occurrence_args k kind args - | Lcase(_,t,a,br) -> - let kind = occurrence k (occurrence k kind t) a in - let r = ref kind in - Array.iter (fun (_,ids,c) -> - r := occurrence (k+Array.length ids) kind c && !r) br; - !r - | Lif (t, bt, bf) -> - let kind = occurrence k kind t in - kind && occurrence k kind bt && occurrence k kind bf - | Lfix(_,(ids,ltypes,lbodies)) - | Lcofix(_,(ids,ltypes,lbodies)) -> - let kind = occurrence_args k kind ltypes in - let _ = occurrence_args (k+Array.length ids) false lbodies in - kind - -and occurrence_args k kind args = - Array.fold_left (occurrence k) kind args - -let occur_once lam = - try let _ = occurrence 1 true lam in true - with Not_found -> false - -(* [remove_let lam] remove let expression in [lam] if the variable is *) -(* used at most once time in the body, and does not appear under *) -(* a lambda or a fix or a cofix *) - -let rec remove_let subst lam = - match lam with - | Lrel(id,i) -> lam_subst_rel lam id i subst - | Llet(id,def,body) -> - let def' = remove_let subst def in - if occur_once body then remove_let (cons def' subst) body - else - let body' = remove_let (lift subst) body in - if def == def' && body == body' then lam else Llet(id,def',body') - | _ -> map_lam_with_binders liftn remove_let subst lam - - (*s Translation from [constr] to [lambda] *) (* Translation of constructor *) @@ -410,8 +337,6 @@ module Vect = size = 0; } - let length v = v.size - let extend v = if Int.equal v.size (Array.length v.elems) then let new_size = min (2*v.size) Sys.max_array_length in @@ -425,33 +350,15 @@ module Vect = v.elems.(v.size) <- a; v.size <- v.size + 1 - let push_pos v a = - let pos = v.size in - push v a; - pos - let popn v n = v.size <- max 0 (v.size - n) let pop v = popn v 1 - let get v n = - if v.size <= n then invalid_arg "Vect.get:index out of bounds"; - v.elems.(n) - let get_last v n = if v.size <= n then invalid_arg "Vect.get:index out of bounds"; v.elems.(v.size - n - 1) - - let last v = - if Int.equal v.size 0 then invalid_arg "Vect.last:index out of bounds"; - v.elems.(v.size - 1) - - let clear v = v.size <- 0 - - let to_array v = Array.sub v.elems 0 v.size - end let empty_args = [||] -- cgit v1.2.3 From dba083b267a2aede3654dcebaab0013d284d32c7 Mon Sep 17 00:00:00 2001 From: "Abhishek Anand (@brixpro-home)" Date: Sun, 22 Jan 2017 11:59:22 -0500 Subject: applied the patch for printing types of let bindings by @herbelin --- pretyping/detyping.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 483e2b432..1eec85f45 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -695,7 +695,7 @@ and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = let c = detype (lax,false) avoid env sigma (Option.get body) in (* Heuristic: we display the type if in Prop *) let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in - let t = if s != InProp then None else Some (detype (lax,false) avoid env sigma ty) in + let t = if s != InProp && not !Flags.raw_print then None else Some (detype (lax,false) avoid env sigma ty) in GLetIn (dl, na', c, t, r) let detype_rel_context ?(lax=false) where avoid env sigma sign = -- cgit v1.2.3 From 4a84961049f4f00897ae72a13954edbcc9aaba5e Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Wed, 3 May 2017 10:23:46 +0200 Subject: Fix outdated description in RefMan. --- doc/refman/RefMan-pro.tex | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex index c37367de5..4c333379b 100644 --- a/doc/refman/RefMan-pro.tex +++ b/doc/refman/RefMan-pro.tex @@ -118,7 +118,7 @@ the current proof and declare the initial goal as an axiom. \subsection[\tt Proof {\term}.]{\tt Proof {\term}.\comindex{Proof} \label{BeginProof}} This command applies in proof editing mode. It is equivalent to {\tt - exact {\term}; Save.} That is, you have to give the full proof in + exact {\term}. Qed.} That is, you have to give the full proof in one gulp, as a proof term (see Section~\ref{exact}). \variant {\tt Proof.} -- cgit v1.2.3 From e9b745af47ba3386724b874e3fd74b6dad33b015 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Thu, 6 Apr 2017 22:48:32 +0200 Subject: Allow flexible anonymous universes in instances and sorts. The addition to the test suite showcases the usage. --- intf/misctypes.mli | 4 +-- library/declare.ml | 5 ++-- parsing/g_constr.ml4 | 8 ++--- pretyping/detyping.ml | 7 +++-- pretyping/miscops.ml | 2 +- pretyping/pretyping.ml | 61 +++++++++++++++++++++------------------ printing/ppconstr.ml | 8 ++--- test-suite/success/polymorphism.v | 32 +++++++++++++++++++- 8 files changed, 82 insertions(+), 45 deletions(-) diff --git a/intf/misctypes.mli b/intf/misctypes.mli index 33dc2a335..7c2dc5177 100644 --- a/intf/misctypes.mli +++ b/intf/misctypes.mli @@ -48,8 +48,8 @@ type 'a glob_sort_gen = | GProp (** representation of [Prop] literal *) | GSet (** representation of [Set] literal *) | GType of 'a (** representation of [Type] literal *) -type sort_info = string Loc.located list -type level_info = string Loc.located option +type sort_info = Name.t Loc.located list +type level_info = Name.t Loc.located option type glob_sort = sort_info glob_sort_gen type glob_level = level_info glob_sort_gen diff --git a/library/declare.ml b/library/declare.ml index 31c9c24bc..91e0cb44b 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -514,11 +514,10 @@ let do_constraint poly l = match x with | GProp -> Loc.dummy_loc, (false, Univ.Level.prop) | GSet -> Loc.dummy_loc, (false, Univ.Level.set) - | GType None -> + | GType None | GType (Some (_, Anonymous)) -> user_err ~hdr:"Constraint" (str "Cannot declare constraints on anonymous universes") - | GType (Some (loc, id)) -> - let id = Id.of_string id in + | GType (Some (loc, Name id)) -> let names, _ = Global.global_universe_names () in try loc, Idmap.find id names with Not_found -> diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4 index 0f2ed88fe..15f100c3b 100644 --- a/parsing/g_constr.ml4 +++ b/parsing/g_constr.ml4 @@ -146,12 +146,12 @@ GEXTEND Gram [ [ "Set" -> GSet | "Prop" -> GProp | "Type" -> GType [] - | "Type"; "@{"; u = universe; "}" -> GType (List.map (fun (loc,x) -> (loc, Id.to_string x)) u) + | "Type"; "@{"; u = universe; "}" -> GType u ] ] ; universe: - [ [ IDENT "max"; "("; ids = LIST1 identref SEP ","; ")" -> ids - | id = identref -> [id] + [ [ IDENT "max"; "("; ids = LIST1 name SEP ","; ")" -> ids + | id = name -> [id] ] ] ; lconstr: @@ -298,7 +298,7 @@ GEXTEND Gram [ [ "Set" -> GSet | "Prop" -> GProp | "Type" -> GType None - | id = identref -> GType (Some (fst id, Id.to_string (snd id))) + | id = name -> GType (Some id) ] ] ; fix_constr: diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 483e2b432..8a90a3f9b 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -422,7 +422,9 @@ let detype_sort sigma = function | Type u -> GType (if !print_universes - then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)] + then + let u = Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u) in + [dl, Name.mk_name (Id.of_string_soft u)] else []) type binder_kind = BProd | BLambda | BLetIn @@ -434,7 +436,8 @@ let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index let set_detype_anonymous f = detype_anonymous := f let detype_level sigma l = - GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l))) + let l = Pp.string_of_ppcmds (Termops.pr_evd_level sigma l) in + GType (Some (dl, Name.mk_name (Id.of_string_soft l))) let detype_instance sigma l = let l = EInstance.kind sigma l in diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml index 7fe81c9a4..1669f8334 100644 --- a/pretyping/miscops.ml +++ b/pretyping/miscops.ml @@ -30,7 +30,7 @@ let smartmap_cast_type f c = let glob_sort_eq g1 g2 = match g1, g2 with | GProp, GProp -> true | GSet, GSet -> true -| GType l1, GType l2 -> List.equal (fun x y -> CString.equal (snd x) (snd y)) l1 l2 +| GType l1, GType l2 -> List.equal (fun x y -> Names.Name.equal (snd x) (snd y)) l1 l2 | _ -> false let intro_pattern_naming_eq nam1 nam2 = match nam1, nam2 with diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 68ef97659..767e4be35 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -192,45 +192,50 @@ let _ = optwrite = (:=) Universes.set_minimization }) (** Miscellaneous interpretation functions *) -let interp_universe_level_name evd (loc,s) = - let names, _ = Global.global_universe_names () in - if CString.string_contains ~where:s ~what:"." then - match List.rev (CString.split '.' s) with - | [] -> anomaly (str"Invalid universe name " ++ str s) - | n :: dp -> - let num = int_of_string n in - let dp = DirPath.make (List.map Id.of_string dp) in - let level = Univ.Level.make dp num in - let evd = - try Evd.add_global_univ evd level - with UGraph.AlreadyDeclared -> evd - in evd, level - else - try - let level = Evd.universe_of_name evd s in - evd, level - with Not_found -> - try - let id = try Id.of_string s with _ -> raise Not_found in - evd, snd (Idmap.find id names) - with Not_found -> - if not (is_strict_universe_declarations ()) then - new_univ_level_variable ~loc ~name:s univ_rigid evd - else user_err ~loc ~hdr:"interp_universe_level_name" - (Pp.(str "Undeclared universe: " ++ str s)) +let interp_universe_level_name ~anon_rigidity evd (loc,s) = + match s with + | Anonymous -> + new_univ_level_variable ~loc anon_rigidity evd + | Name s -> + let s = Id.to_string s in + let names, _ = Global.global_universe_names () in + if CString.string_contains ~where:s ~what:"." then + match List.rev (CString.split '.' s) with + | [] -> anomaly (str"Invalid universe name " ++ str s) + | n :: dp -> + let num = int_of_string n in + let dp = DirPath.make (List.map Id.of_string dp) in + let level = Univ.Level.make dp num in + let evd = + try Evd.add_global_univ evd level + with UGraph.AlreadyDeclared -> evd + in evd, level + else + try + let level = Evd.universe_of_name evd s in + evd, level + with Not_found -> + try + let id = try Id.of_string s with _ -> raise Not_found in + evd, snd (Idmap.find id names) + with Not_found -> + if not (is_strict_universe_declarations ()) then + new_univ_level_variable ~loc ~name:s univ_rigid evd + else user_err ~loc ~hdr:"interp_universe_level_name" + (Pp.(str "Undeclared universe: " ++ str s)) let interp_universe ?loc evd = function | [] -> let evd, l = new_univ_level_variable ?loc univ_rigid evd in evd, Univ.Universe.make l | l -> List.fold_left (fun (evd, u) l -> - let evd', l = interp_universe_level_name evd l in + let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible_alg evd l in (evd', Univ.sup u (Univ.Universe.make l))) (evd, Univ.Universe.type0m) l let interp_level_info loc evd : Misctypes.level_info -> _ = function | None -> new_univ_level_variable ~loc univ_rigid evd - | Some (loc,s) -> interp_universe_level_name evd (loc,s) + | Some (loc,s) -> interp_universe_level_name ~anon_rigidity:univ_flexible evd (loc,s) let interp_sort ?loc evd = function | GProp -> evd, Prop Null diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 3041ac259..b546c39ae 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -151,8 +151,8 @@ let tag_var = tag Tag.variable let pr_univ l = match l with - | [_,x] -> str x - | l -> str"max(" ++ prlist_with_sep (fun () -> str",") (fun x -> str (snd x)) l ++ str")" + | [_,x] -> pr_name x + | l -> str"max(" ++ prlist_with_sep (fun () -> str",") (fun x -> pr_name (snd x)) l ++ str")" let pr_univ_annot pr x = str "@{" ++ pr x ++ str "}" @@ -166,7 +166,7 @@ let tag_var = tag Tag.variable | GProp -> tag_type (str "Prop") | GSet -> tag_type (str "Set") | GType None -> tag_type (str "Type") - | GType (Some (_, u)) -> tag_type (str u) + | GType (Some (_, u)) -> tag_type (pr_name u) let pr_qualid sp = let (sl, id) = repr_qualid sp in @@ -191,7 +191,7 @@ let tag_var = tag Tag.variable tag_type (str "Set") | GType u -> (match u with - | Some (_,u) -> str u + | Some (_,u) -> pr_name u | None -> tag_type (str "Type")) let pr_universe_instance l = diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v index 878875bd9..0a58fe89a 100644 --- a/test-suite/success/polymorphism.v +++ b/test-suite/success/polymorphism.v @@ -321,4 +321,34 @@ Definition unwrap' := fun (X : Type) (b : box X) => let (unw) := b in unw. Fail Definition bad : False := TypeNeqSmallType.paradox (unwrap' Type (wrap _ Type)) eq_refl. -End Hurkens'. \ No newline at end of file +End Hurkens'. + +Module Anonymous. + Set Universe Polymorphism. + + Definition defaultid := (fun x => x) : Type -> Type. + Definition collapseid := defaultid@{_ _}. + Check collapseid@{_}. + + Definition anonid := (fun x => x) : Type -> Type@{_}. + Check anonid@{_}. + + Definition defaultalg := Type : Type. + Definition usedefaultalg := defaultalg@{_ _}. + Check usedefaultalg@{_ _}. + + Definition anonalg := (fun x => x) (Type : Type@{_}). + Check anonalg@{_}. + + Definition unrelated@{i j} := nat. + Definition useunrelated := unrelated@{_ _}. + Check useunrelated@{_ _}. + + Definition inthemiddle@{i j k} := + let _ := defaultid@{i j} in + defaultalg@{k j}. + (* i <= j < k *) + Definition collapsethemiddle := inthemiddle@{i _ j}. + Check collapsethemiddle@{_ _}. + +End Anonymous. -- cgit v1.2.3 From 4361c1ed9ac5646055f9f0eecc4a003d720c1994 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Wed, 12 Apr 2017 13:29:16 +0200 Subject: Type@{_} should not produce a flexible algebraic universe. Otherwise [(fun x => x) (Type : Type@{_})] becomes [(fun x : Type@{i+1} => x) (Type@{i} : Type@{i+1})] breaking the invariant that terms do not contain algebraic universes (at the lambda abstraction). --- pretyping/pretyping.ml | 3 ++- test-suite/success/polymorphism.v | 10 +++++----- 2 files changed, 7 insertions(+), 6 deletions(-) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 767e4be35..4886423bd 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -229,7 +229,8 @@ let interp_universe ?loc evd = function evd, Univ.Universe.make l | l -> List.fold_left (fun (evd, u) l -> - let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible_alg evd l in + (* [univ_flexible_alg] can produce algebraic universes in terms *) + let evd', l = interp_universe_level_name ~anon_rigidity:univ_flexible evd l in (evd', Univ.sup u (Univ.Universe.make l))) (evd, Univ.Universe.type0m) l diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v index 0a58fe89a..66ff55edc 100644 --- a/test-suite/success/polymorphism.v +++ b/test-suite/success/polymorphism.v @@ -333,12 +333,12 @@ Module Anonymous. Definition anonid := (fun x => x) : Type -> Type@{_}. Check anonid@{_}. - Definition defaultalg := Type : Type. - Definition usedefaultalg := defaultalg@{_ _}. + Definition defaultalg := (fun x : Type => x) (Type : Type). + Definition usedefaultalg := defaultalg@{_ _ _}. Check usedefaultalg@{_ _}. - Definition anonalg := (fun x => x) (Type : Type@{_}). - Check anonalg@{_}. + Definition anonalg := (fun x : Type@{_} => x) (Type : Type). + Check anonalg@{_ _}. Definition unrelated@{i j} := nat. Definition useunrelated := unrelated@{_ _}. @@ -346,7 +346,7 @@ Module Anonymous. Definition inthemiddle@{i j k} := let _ := defaultid@{i j} in - defaultalg@{k j}. + anonalg@{k j}. (* i <= j < k *) Definition collapsethemiddle := inthemiddle@{i _ j}. Check collapsethemiddle@{_ _}. -- cgit v1.2.3 From 8adfa0e5290056b7683a3a8b778ca16182a1eb3d Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Tue, 2 May 2017 14:43:32 +0200 Subject: Reorganize comment documentation of ChoiceFacts.v Shortnames and natural language descriptions of principles are moved next to each principle. The table of contents is moved to after the principle definitions. Extra definitions are moved to the definition section (eg DependentFunctionalChoice) Compatibility notations have been moved to the end of the file. Details: The following used to be announced but were neither defined or used, and have been removed: - OAC! - Ext_pred = extensionality of predicates - Ext_fun_prop_repr = choice of a representative among extensional functions to Prop GuardedFunctionalRelReification was announced with shortname GAC! but shortname GFR_fun was used next to it. Only the former has been retained. Shortnames and descriptions have been invented for InhabitedForallCommute DependentFunctionalRelReification ExtensionalPropositionRepresentative ExtensionalFunctionRepresentative Some modification of headlines --- theories/Logic/ChoiceFacts.v | 283 ++++++++++++++++++++----------------------- 1 file changed, 131 insertions(+), 152 deletions(-) diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v index f1f20606b..116897f4c 100644 --- a/theories/Logic/ChoiceFacts.v +++ b/theories/Logic/ChoiceFacts.v @@ -8,94 +8,9 @@ (************************************************************************) (** Some facts and definitions concerning choice and description in - intuitionistic logic. - -We investigate the relations between the following choice and -description principles - -- AC_rel = relational form of the (non extensional) axiom of choice - (a "set-theoretic" axiom of choice) -- AC_fun = functional form of the (non extensional) axiom of choice - (a "type-theoretic" axiom of choice) -- DC_fun = functional form of the dependent axiom of choice -- ACw_fun = functional form of the countable axiom of choice -- AC! = functional relation reification - (known as axiom of unique choice in topos theory, - sometimes called principle of definite description in - the context of constructive type theory, sometimes - called axiom of no choice) - -- AC_fun_repr = functional choice of a representative in an equivalence class -- AC_fun_setoid_gen = functional form of the general form of the (so-called - extensional) axiom of choice over setoids -- AC_fun_setoid = functional form of the (so-called extensional) axiom of - choice from setoids -- AC_fun_setoid_simple = functional form of the (so-called extensional) axiom of - choice from setoids on locally compatible relations - -- GAC_rel = guarded relational form of the (non extensional) axiom of choice -- GAC_fun = guarded functional form of the (non extensional) axiom of choice -- GAC! = guarded functional relation reification - -- OAC_rel = "omniscient" relational form of the (non extensional) axiom of choice -- OAC_fun = "omniscient" functional form of the (non extensional) axiom of choice - (called AC* in Bell [[Bell]]) -- OAC! - -- ID_iota = intuitionistic definite description -- ID_epsilon = intuitionistic indefinite description - -- D_iota = (weakly classical) definite description principle -- D_epsilon = (weakly classical) indefinite description principle - -- PI = proof irrelevance -- IGP = independence of general premises - (an unconstrained generalisation of the constructive principle of - independence of premises) -- Drinker = drinker's paradox (small form) - (called Ex in Bell [[Bell]]) -- EM = excluded-middle - -- Ext_pred_repr = choice of a representative among extensional predicates -- Ext_pred = extensionality of predicates -- Ext_fun_prop_repr = choice of a representative among extensional functions to Prop - -We let also - -- IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.) -- IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.) -- IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.) - -with no prerequisite on the non-emptiness of domains - -Table of contents - -1. Definitions - -2. IPL_2^2 |- AC_rel + AC! = AC_fun - -3.1. typed IPL_2 + Sigma-types + PI |- AC_rel = GAC_rel and IPL_2 |- AC_rel + IGP -> GAC_rel and IPL_2 |- GAC_rel = OAC_rel - -3.2. IPL^2 |- AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker - -3.3. D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker - -4. Derivability of choice for decidable relations with well-ordered codomain - -5. Equivalence of choices on dependent or non dependent functional types - -6. Non contradiction of constructive descriptions wrt functional choices - -7. Definite description transports classical logic to the computational world - -8. Choice -> Dependent choice -> Countable choice - -9.1. AC_fun_ext = AC_fun + Ext_fun_repr + EM - -9.2. AC_fun_ext = AC_fun + Ext_prop_fun_repr + PI - -References: - + intuitionistic logic. *) +(** * References: *) +(** [[Bell]] John L. Bell, Choice principles in intuitionistic set theory, unpublished. @@ -133,47 +48,75 @@ Variable P:A->Prop. (** ** Constructive choice and description *) -(** AC_rel *) +(** AC_rel = relational form of the (non extensional) axiom of choice + (a "set-theoretic" axiom of choice) *) Definition RelationalChoice_on := forall R:A->B->Prop, (forall x : A, exists y : B, R x y) -> (exists R' : A->B->Prop, subrelation R' R /\ forall x, exists! y, R' x y). -(** AC_fun *) +(** AC_fun = functional form of the (non extensional) axiom of choice + (a "type-theoretic" axiom of choice) *) (* Note: This is called Type-Theoretic Description Axiom (TTDA) in [[Werner97]] (using a non-standard meaning of "description"). This is called intensional axiom of choice (AC_int) in [[Carlström04]] *) +Definition FunctionalChoice_on_rel (R:A->B->Prop) := + (forall x:A, exists y : B, R x y) -> + exists f : A -> B, (forall x:A, R x (f x)). + Definition FunctionalChoice_on := forall R:A->B->Prop, (forall x : A, exists y : B, R x y) -> (exists f : A->B, forall x : A, R x (f x)). -(** DC_fun *) +(** AC_fun_dep = functional form of the (non extensional) axiom of + choice, with dependent functions *) +Definition DependentFunctionalChoice_on (A:Type) (B:A -> Type) := + forall R:forall x:A, B x -> Prop, + (forall x:A, exists y : B x, R x y) -> + (exists f : (forall x:A, B x), forall x:A, R x (f x)). + +(** AC_trunc = axiom of choice for propositional truncations + (truncation and quantification commute) *) +Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) := + (forall x, inhabited (B x)) -> inhabited (forall x, B x). + +(** DC_fun = functional form of the dependent axiom of choice *) Definition FunctionalDependentChoice_on := forall (R:A->A->Prop), (forall x, exists y, R x y) -> forall x0, (exists f : nat -> A, f 0 = x0 /\ forall n, R (f n) (f (S n))). -(** ACw_fun *) +(** ACw_fun = functional form of the countable axiom of choice *) Definition FunctionalCountableChoice_on := forall (R:nat->A->Prop), (forall n, exists y, R n y) -> (exists f : nat -> A, forall n, R n (f n)). -(** AC! or Functional Relation Reification (known as Axiom of Unique Choice - in topos theory; also called principle of definite description *) +(** AC! = functional relation reification + (known as axiom of unique choice in topos theory, + sometimes called principle of definite description in + the context of constructive type theory, sometimes + called axiom of no choice) *) Definition FunctionalRelReification_on := forall R:A->B->Prop, (forall x : A, exists! y : B, R x y) -> (exists f : A->B, forall x : A, R x (f x)). -(** AC_fun_repr *) +(** AC_dep! = functional relation reification, with dependent functions + see AC! *) +Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) := + forall (R:forall x:A, B x -> Prop), + (forall x:A, exists! y : B x, R x y) -> + (exists f : (forall x:A, B x), forall x:A, R x (f x)). + +(** AC_fun_repr = functional choice of a representative in an equivalence class *) (* Note: This is called Type-Theoretic Choice Axiom (TTCA) in [[Werner97]] (by reference to the extensional set-theoretic @@ -187,7 +130,8 @@ Definition RepresentativeFunctionalChoice_on := (Equivalence R) -> (exists f : A->A, forall x : A, (R x (f x)) /\ forall x', R x x' -> f x = f x'). -(** AC_fun_setoid *) +(** AC_fun_setoid = functional form of the (so-called extensional) axiom of + choice from setoids *) Definition SetoidFunctionalChoice_on := forall R : A -> A -> Prop, @@ -197,7 +141,8 @@ Definition SetoidFunctionalChoice_on := (forall x, exists y, T x y) -> exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x'). -(** AC_fun_setoid_gen *) +(** AC_fun_setoid_gen = functional form of the general form of the (so-called + extensional) axiom of choice over setoids *) (* Note: This is called extensional axiom of choice (AC_ext) in [[Carlström04]]. *) @@ -213,7 +158,8 @@ Definition GeneralizedSetoidFunctionalChoice_on := exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x')). -(** AC_fun_setoid_simple *) +(** AC_fun_setoid_simple = functional form of the (so-called extensional) axiom of + choice from setoids on locally compatible relations *) Definition SimpleSetoidFunctionalChoice_on A B := forall R : A -> A -> Prop, @@ -222,19 +168,19 @@ Definition SimpleSetoidFunctionalChoice_on A B := (forall x, exists y, forall x', R x x' -> T x' y) -> exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x'). -(** ID_epsilon (constructive version of indefinite description; - combined with proof-irrelevance, it may be connected to - Carlström's type theory with a constructive indefinite description - operator) *) +(** ID_epsilon = constructive version of indefinite description; + combined with proof-irrelevance, it may be connected to + Carlström's type theory with a constructive indefinite description + operator *) Definition ConstructiveIndefiniteDescription_on := forall P:A->Prop, (exists x, P x) -> { x:A | P x }. -(** ID_iota (constructive version of definite description; combined - with proof-irrelevance, it may be connected to Carlström's and - Stenlund's type theory with a constructive definite description - operator) *) +(** ID_iota = constructive version of definite description; + combined with proof-irrelevance, it may be connected to + Carlström's and Stenlund's type theory with a + constructive definite description operator) *) Definition ConstructiveDefiniteDescription_on := forall P:A->Prop, @@ -242,7 +188,7 @@ Definition ConstructiveDefiniteDescription_on := (** ** Weakly classical choice and description *) -(** GAC_rel *) +(** GAC_rel = guarded relational form of the (non extensional) axiom of choice *) Definition GuardedRelationalChoice_on := forall P : A->Prop, forall R : A->B->Prop, @@ -250,7 +196,7 @@ Definition GuardedRelationalChoice_on := (exists R' : A->B->Prop, subrelation R' R /\ forall x, P x -> exists! y, R' x y). -(** GAC_fun *) +(** GAC_fun = guarded functional form of the (non extensional) axiom of choice *) Definition GuardedFunctionalChoice_on := forall P : A->Prop, forall R : A->B->Prop, @@ -258,7 +204,7 @@ Definition GuardedFunctionalChoice_on := (forall x : A, P x -> exists y : B, R x y) -> (exists f : A->B, forall x, P x -> R x (f x)). -(** GFR_fun *) +(** GAC! = guarded functional relation reification *) Definition GuardedFunctionalRelReification_on := forall P : A->Prop, forall R : A->B->Prop, @@ -266,27 +212,28 @@ Definition GuardedFunctionalRelReification_on := (forall x : A, P x -> exists! y : B, R x y) -> (exists f : A->B, forall x : A, P x -> R x (f x)). -(** OAC_rel *) +(** OAC_rel = "omniscient" relational form of the (non extensional) axiom of choice *) Definition OmniscientRelationalChoice_on := forall R : A->B->Prop, exists R' : A->B->Prop, subrelation R' R /\ forall x : A, (exists y : B, R x y) -> exists! y, R' x y. -(** OAC_fun *) +(** OAC_fun = "omniscient" functional form of the (non extensional) axiom of choice + (called AC* in Bell [[Bell]]) *) Definition OmniscientFunctionalChoice_on := forall R : A->B->Prop, inhabited B -> exists f : A->B, forall x : A, (exists y : B, R x y) -> R x (f x). -(** D_epsilon *) +(** D_epsilon = (weakly classical) indefinite description principle *) Definition EpsilonStatement_on := forall P:A->Prop, inhabited A -> { x:A | (exists x, P x) -> P x }. -(** D_iota *) +(** D_iota = (weakly classical) definite description principle *) Definition IotaStatement_on := forall P:A->Prop, @@ -300,14 +247,20 @@ Notation RelationalChoice := (forall A B : Type, RelationalChoice_on A B). Notation FunctionalChoice := (forall A B : Type, FunctionalChoice_on A B). -Definition FunctionalDependentChoice := +Notation DependentFunctionalChoice := + (forall A (B:A->Type), DependentFunctionalChoice_on B). +Notation InhabitedForallCommute := + (forall A (B : A -> Type), InhabitedForallCommute_on B). +Notation FunctionalDependentChoice := (forall A : Type, FunctionalDependentChoice_on A). -Definition FunctionalCountableChoice := +Notation FunctionalCountableChoice := (forall A : Type, FunctionalCountableChoice_on A). Notation FunctionalChoiceOnInhabitedSet := (forall A B : Type, inhabited B -> FunctionalChoice_on A B). Notation FunctionalRelReification := (forall A B : Type, FunctionalRelReification_on A B). +Notation DependentFunctionalRelReification := + (forall A (B:A->Type), DependentFunctionalRelReification_on B). Notation RepresentativeFunctionalChoice := (forall A : Type, RepresentativeFunctionalChoice_on A). Notation SetoidFunctionalChoice := @@ -341,38 +294,87 @@ Notation EpsilonStatement := (** Subclassical schemes *) +(** PI = proof irrelevance *) Definition ProofIrrelevance := forall (A:Prop) (a1 a2:A), a1 = a2. +(** IGP = independence of general premises + (an unconstrained generalisation of the constructive principle of + independence of premises) *) Definition IndependenceOfGeneralPremises := forall (A:Type) (P:A -> Prop) (Q:Prop), inhabited A -> (Q -> exists x, P x) -> exists x, Q -> P x. +(** Drinker = drinker's paradox (small form) + (called Ex in Bell [[Bell]]) *) Definition SmallDrinker'sParadox := forall (A:Type) (P:A -> Prop), inhabited A -> exists x, (exists x, P x) -> P x. +(** EM = excluded-middle *) Definition ExcludedMiddle := forall P:Prop, P \/ ~ P. (** Extensional schemes *) +(** Ext_prop_repr = choice of a representative among extensional propositions *) Local Notation ExtensionalPropositionRepresentative := (forall (A:Type), exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ forall Q, (P <-> Q) -> h P = h Q). +(** Ext_pred_repr = choice of a representative among extensional predicates *) Local Notation ExtensionalPredicateRepresentative := (forall (A:Type), exists h : (A->Prop) -> (A->Prop), forall (P : A -> Prop), (forall x, P x <-> h P x) /\ forall Q, (forall x, P x <-> Q x) -> h P = h Q). +(** Ext_fun_repr = choice of a representative among extensional functions *) Local Notation ExtensionalFunctionRepresentative := (forall (A B:Type), exists h : (A->B) -> (A->B), forall (f : A -> B), (forall x, f x = h f x) /\ forall g, (forall x, f x = g x) -> h f = h g). +(** We let also + +- IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.) +- IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.) +- IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.) + +with no prerequisite on the non-emptiness of domains +*) + +(**********************************************************************) +(** * Table of contents *) + +(* This is very fragile. *) +(** +1. Definitions + +2. IPL_2^2 |- AC_rel + AC! = AC_fun + +3.1. typed IPL_2 + Sigma-types + PI |- AC_rel = GAC_rel and IPL_2 |- AC_rel + IGP -> GAC_rel and IPL_2 |- GAC_rel = OAC_rel + +3.2. IPL^2 |- AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker + +3.3. D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker + +4. Derivability of choice for decidable relations with well-ordered codomain + +5. AC_fun = AC_fun_dep = AC_trunc + +6. Non contradiction of constructive descriptions wrt functional choices + +7. Definite description transports classical logic to the computational world + +8. Choice -> Dependent choice -> Countable choice + +9.1. AC_fun_setoid = AC_fun + Ext_fun_repr + EM + +9.2. AC_fun_setoid = AC_fun + Ext_pred_repr + PI + *) + (**********************************************************************) (** * AC_rel + AC! = AC_fun @@ -400,9 +402,6 @@ Proof. apply HR'R; assumption. Qed. -Notation description_rel_choice_imp_funct_choice := - functional_rel_reification_and_rel_choice_imp_fun_choice (compat "8.6"). - Lemma fun_choice_imp_rel_choice : forall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A B. Proof. @@ -416,8 +415,6 @@ Proof. trivial. Qed. -Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (compat "8.6"). - Lemma fun_choice_imp_functional_rel_reification : forall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A B. Proof. @@ -431,8 +428,6 @@ Proof. exists f; exact H0. Qed. -Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (compat "8.6"). - Corollary fun_choice_iff_rel_choice_and_functional_rel_reification : forall A B : Type, FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A B. @@ -444,8 +439,6 @@ Proof. intros [H H0]; exact (functional_rel_reification_and_rel_choice_imp_fun_choice H0 H). Qed. -Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr := - fun_choice_iff_rel_choice_and_functional_rel_reification (compat "8.6"). (**********************************************************************) (** * Connection between the guarded, non guarded and omniscient choices *) @@ -687,10 +680,6 @@ Qed. Require Import Wf_nat. Require Import Decidable. -Definition FunctionalChoice_on_rel (A B:Type) (R:A->B->Prop) := - (forall x:A, exists y : B, R x y) -> - exists f : A -> B, (forall x:A, R x (f x)). - Lemma classical_denumerable_description_imp_fun_choice : forall A:Type, FunctionalRelReification_on A nat -> @@ -712,18 +701,10 @@ Proof. Qed. (**********************************************************************) -(** * Choice on dependent and non dependent function types are equivalent *) +(** * AC_fun = AC_fun_dep = AC_trunc *) (** ** Choice on dependent and non dependent function types are equivalent *) -Definition DependentFunctionalChoice_on (A:Type) (B:A -> Type) := - forall R:forall x:A, B x -> Prop, - (forall x:A, exists y : B x, R x y) -> - (exists f : (forall x:A, B x), forall x:A, R x (f x)). - -Notation DependentFunctionalChoice := - (forall A (B:A->Type), DependentFunctionalChoice_on B). - (** The easy part *) Theorem dep_non_dep_functional_choice : @@ -760,13 +741,7 @@ Proof. destruct Heq using eq_indd; trivial. Qed. -(** Functional choice can be reformulated as a property on [inhabited] *) - -Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) := - (forall x, inhabited (B x)) -> inhabited (forall x, B x). - -Notation InhabitedForallCommute := - (forall A (B : A -> Type), InhabitedForallCommute_on B). +(** ** Functional choice and truncation choice are equivalent *) Theorem functional_choice_to_inhabited_forall_commute : FunctionalChoice -> InhabitedForallCommute. @@ -795,14 +770,6 @@ Qed. (** ** Reification of dependent and non dependent functional relation are equivalent *) -Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) := - forall (R:forall x:A, B x -> Prop), - (forall x:A, exists! y : B x, R x y) -> - (exists f : (forall x:A, B x), forall x:A, R x (f x)). - -Notation DependentFunctionalRelReification := - (forall A (B:A->Type), DependentFunctionalRelReification_on B). - (** The easy part *) Theorem dep_non_dep_functional_rel_reification : @@ -1337,3 +1304,15 @@ Proof. apply repr_fun_choice_imp_excluded_middle. now apply setoid_fun_choice_imp_repr_fun_choice. Qed. + +(**********************************************************************) +(** * Compatibility notations *) +Notation description_rel_choice_imp_funct_choice := + functional_rel_reification_and_rel_choice_imp_fun_choice (compat "8.6"). + +Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (compat "8.6"). + +Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr := + fun_choice_iff_rel_choice_and_functional_rel_reification (compat "8.6"). + +Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (compat "8.6"). -- cgit v1.2.3 From ca4aee0fcf1b54363a6a1eb837cd05cd7ffcc0d9 Mon Sep 17 00:00:00 2001 From: Tej Chajed Date: Wed, 3 May 2017 07:47:51 -0400 Subject: Report a useful error for dependent induction The dependent induction tactic notation is in the standard library but not loaded by default, leading to a parser error message that is confusing and unhelpful. This commit adds a notation for dependent induction to Init that fails and reports [Require Import Coq.Program.Equality.] is required to use [dependent induction]. --- theories/Init/Tactics.v | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v index 5d1e87ae0..7a846cd1b 100644 --- a/theories/Init/Tactics.v +++ b/theories/Init/Tactics.v @@ -236,3 +236,10 @@ Tactic Notation "clear" "dependent" hyp(h) := Tactic Notation "revert" "dependent" hyp(h) := generalize dependent h. + +(** Provide an error message for dependent induction that reports an import is +required to use it. Importing Coq.Program.Equality will shadow this notation +with the actual [dependent induction] tactic. *) + +Tactic Notation "dependent" "induction" ident(H) := + fail "To use dependent induction, first [Require Import Coq.Program.Equality.]". -- cgit v1.2.3 From 844bffb7d6c84a02dcef300dda9099487b23c09a Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Tue, 11 Apr 2017 21:17:03 +0200 Subject: Added an option Set Debug Cbv. --- pretyping/cbv.ml | 27 ++++++++++++++++++++++++--- 1 file changed, 24 insertions(+), 3 deletions(-) diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index e18625c42..bd7350dc4 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -175,6 +175,19 @@ let cofixp_reducible flgs _ stk = else false +let debug_cbv = ref false +let _ = Goptions.declare_bool_option { + Goptions.optsync = true; Goptions.optdepr = false; + Goptions.optname = "cbv visited constants display"; + Goptions.optkey = ["Debug";"Cbv"]; + Goptions.optread = (fun () -> !debug_cbv); + Goptions.optwrite = (fun a -> debug_cbv:=a); +} + +let pr_key = function + | ConstKey (sp,_) -> Names.Constant.print sp + | VarKey id -> Names.Id.print id + | RelKey n -> Pp.(str "REL_" ++ int n) (* The main recursive functions * @@ -254,9 +267,17 @@ let rec norm_head info env t stack = and norm_head_ref k info env stack normt = if red_set_ref (info_flags info) normt then match ref_value_cache info normt with - | Some body -> strip_appl (shift_value k body) stack - | None -> (VAL(0,make_constr_ref k normt),stack) - else (VAL(0,make_constr_ref k normt),stack) + | Some body -> + if !debug_cbv then Feedback.msg_debug Pp.(str "Unfolding " ++ pr_key normt); + strip_appl (shift_value k body) stack + | None -> + if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt); + (VAL(0,make_constr_ref k normt),stack) + else + begin + if !debug_cbv then Feedback.msg_debug Pp.(str "Not unfolding " ++ pr_key normt); + (VAL(0,make_constr_ref k normt),stack) + end (* cbv_stack_term performs weak reduction on constr t under the subs * env, with context stack, i.e. ([env]t stack). First computes weak -- cgit v1.2.3 From 952e9aea47d3fad2d0f488d968ff0e90fa403ebc Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Wed, 3 May 2017 11:37:08 +0200 Subject: Adding an even more compact mode for goal display. We want to print variables in vertical boxes as much as possible. The criterion to distinguish "variable" from "hypothesis" is not obvious. We chose this one but may change it in the future: X:T is a variable if T is not of type Prop AND T is "simple" which means T is either id or (t T1 .. Tn) Ti's are sort-typed (typicall Ti:Type, but not Ti:nat). --- printing/printer.ml | 60 +++++++++++++++++++++++++++++++++++++++++++------ printing/printer.mli | 5 ++++- vernac/vernacentries.ml | 18 +++++++++++++++ 3 files changed, 75 insertions(+), 8 deletions(-) diff --git a/printing/printer.ml b/printing/printer.ml index 93d221f03..0e31a4a04 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -262,6 +262,13 @@ let pr_pattern t = pr_pattern_env (Global.env()) empty_names_context t*) (**********************************************************************) (* Contexts and declarations *) + +(* Flag for compact display of goals *) + +let get_compact_context,set_compact_context = + let compact_context = ref false in + (fun () -> !compact_context),(fun b -> compact_context := b) + let pr_compacted_decl env sigma decl = let ids, pbody, typ = match decl with | CompactedDecl.LocalAssum (ids, typ) -> @@ -342,15 +349,55 @@ let pr_ne_context_of header env sigma = List.is_empty (Environ.named_context env) then (mt ()) else let penv = pr_context_unlimited env sigma in (header ++ penv ++ fnl ()) +(* Heuristic for horizontalizing hypothesis: + Detecting variable which type is a simple id or of the form (t x y ...) + where t is a product or only sorts (typically [Type -> Type -> ...] + and not [nat -> nat -> ...] ). + + Special case for non-Prop dependent terms. *) +let rec should_compact env sigma typ = + get_compact_context() && + match kind_of_term typ with + | Rel _ | Var _ | Sort _ | Const _ | Ind _ -> true + | App (c,args) -> + let _,type_of_c = Typing.type_of env sigma c in + let _,type_of_typ = Typing.type_of env sigma typ in + not (is_Prop type_of_typ) + && (* These two more tests detect rare cases of non-Prop-sorted + dependent hypothesis: *) + let lnamedtyp , _ = decompose_prod type_of_c in + (* c has a non dependent type *) + List.for_all (fun (_,typarg) -> isSort typarg) lnamedtyp + && (* and real arguments are recursively elligible to compaction. *) + Array.for_all (should_compact env sigma) args + | _ -> false + + +(* If option Compact Contexts is set, we pack "simple" hypothesis in a + hov box (with three sapaces as a separator), the global box being a + v box *) let rec bld_sign_env env sigma ctxt pps = match ctxt with | [] -> pps + | CompactedDecl.LocalAssum (ids,typ)::ctxt' when should_compact env sigma typ -> + let pps',ctxt' = bld_sign_env_id env sigma ctxt (mt ()) true in + (* putting simple hyps in a more horizontal flavor *) + bld_sign_env env sigma ctxt' (pps ++ brk (0,0) ++ hov 0 pps') | d:: ctxt' -> - let pidt = pr_var_list_decl env sigma d in - let pps' = pps ++ brk (0,0) ++ pidt in - bld_sign_env env sigma ctxt' pps' + let pidt = pr_var_list_decl env sigma d in + let pps' = pps ++ brk (0,0) ++ pidt in + bld_sign_env env sigma ctxt' pps' +and bld_sign_env_id env sigma ctxt pps is_start = + match ctxt with + | [] -> pps,ctxt + | CompactedDecl.LocalAssum(ids,typ) as d :: ctxt' when should_compact env sigma typ -> + let pidt = pr_var_list_decl env sigma d in + let pps' = pps ++ (if not is_start then brk (3,0) else (mt ())) ++ pidt in + bld_sign_env_id env sigma ctxt' pps' false + | _ -> pps,ctxt +(* compact printing an env (variables and de Bruijn). Separator: three + spaces between simple hyps, and newline otherwise *) let pr_context_limit_compact ?n env sigma = let ctxt = Termops.compact_named_context (named_context env) in let lgth = List.length ctxt in @@ -360,15 +407,14 @@ let pr_context_limit_compact ?n env sigma = | Some n when n > lgth -> lgth | Some n -> n in let ctxt_chopped,ctxt_hidden = Util.List.chop n_capped ctxt in - (* a dot line hinting the number of hiden hyps. *) + (* a dot line hinting the number of hidden hyps. *) let hidden_dots = String.make (List.length ctxt_hidden) '.' in let sign_env = v 0 (str hidden_dots ++ (mt ()) ++ bld_sign_env env sigma (List.rev ctxt_chopped) (mt ())) in let db_env = - fold_rel_context - (fun env d pps -> pps ++ fnl () ++ pr_rel_decl env sigma d) + fold_rel_context (fun env d pps -> pps ++ fnl () ++ pr_rel_decl env sigma d) env ~init:(mt ()) in - (sign_env ++ db_env) + sign_env ++ db_env (* The number of printed hypothesis in a goal *) (* If [None], no limit *) diff --git a/printing/printer.mli b/printing/printer.mli index 504392e35..9c9f6e766 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -111,6 +111,9 @@ val pr_pconstructor : env -> pconstructor -> std_ppcmds (** Contexts *) +(** Display compact contexts of goals (simple hyps on the same line) *) +val set_compact_context : bool -> unit +val get_compact_context : unit -> bool val pr_context_unlimited : env -> evar_map -> std_ppcmds val pr_ne_context_of : std_ppcmds -> env -> evar_map -> std_ppcmds @@ -132,7 +135,7 @@ val pr_cpred : Cpred.t -> std_ppcmds val pr_idpred : Id.Pred.t -> std_ppcmds val pr_transparent_state : transparent_state -> std_ppcmds -(** Proofs *) +(** Proofs, these functions obey [Hyps Limit] and [Compact contexts]. *) val pr_goal : goal sigma -> std_ppcmds val pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> std_ppcmds diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 2cb6f3918..40cd1a29c 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1411,6 +1411,24 @@ let _ = optread = (fun _ -> None); optwrite = (fun _ -> ()) } +let _ = + declare_int_option + { optsync = false; + optdepr = false; + optname = "the hypotheses limit"; + optkey = ["Hyps";"Limit"]; + optread = Flags.print_hyps_limit; + optwrite = Flags.set_print_hyps_limit } + +let _ = + declare_bool_option + { optsync = true; + optdepr = false; + optname = "display compact goal contexts"; + optkey = ["Printing";"Compact";"Contexts"]; + optread = (fun () -> Printer.get_compact_context()); + optwrite = (fun b -> Printer.set_compact_context b) } + let _ = declare_int_option { optsync = true; -- cgit v1.2.3 From b12af1535c0ba8cab23e7f9ff18f75688c0e523c Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Wed, 3 May 2017 19:16:48 +0200 Subject: Make congruence reuse discriminate instead of rolling its own. This changes the produced terms a bit, eg Axiom T : Type. Lemma foo : true = false -> T. Proof. congruence. Qed. used to produce fun H : true = false => let Heq := H : true = false in @eq_rect Type True (fun X : Type => X) I T (@f_equal bool Type (fun t : bool => if t then True else T) true false Heq) now produces fun H : true = false => let Heq : true = false := H in let H0 : False := @eq_ind bool true (fun e : bool => if e then True else False) I false Heq in False_rect T H0 i.e. instead of proving [True = T] by [f_equal] then transporting [I] across this identity, it now proves [False] by [eq_ind] then uses exfalso. --- plugins/cc/cctac.ml | 27 ++++++--------------------- 1 file changed, 6 insertions(+), 21 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 7c5efaea3..1cb417bf4 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -397,33 +397,18 @@ let convert_to_hyp_tac c1 t1 c2 t2 p = simplest_elim false_t] end } -let discriminate_tac (cstr,u as cstru) p = +(* Essentially [assert (Heq : lhs = rhs) by proof_tac p; discriminate Heq] *) +let discriminate_tac cstru p = Proofview.Goal.enter { enter = begin fun gl -> - let t1=constr_of_term p.p_lhs and t2=constr_of_term p.p_rhs in + let lhs=constr_of_term p.p_lhs and rhs=constr_of_term p.p_rhs in let env = Proofview.Goal.env gl in - let concl = Proofview.Goal.concl gl in - let xid = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in - let identity = Universes.constr_of_global (Lazy.force _I) in - let identity = EConstr.of_constr identity in - let trivial = Universes.constr_of_global (Lazy.force _True) in - let trivial = EConstr.of_constr trivial in let evm = Tacmach.New.project gl in - let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl t1) in - let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in - let outtype = mkSort outtype in - let pred = mkLambda(Name xid,outtype,mkRel 1) in + let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl lhs) in let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in - let proj = build_projection intype cstru trivial concl gl in - let injt = app_global _f_equal - [|intype;outtype;proj;t1;t2;mkVar hid|] in - let endt k = - injt (fun injt -> - app_global _eq_rect - [|outtype;trivial;pred;identity;concl;injt|] k) in - let neweq = app_global _eq [|intype;t1;t2|] in + let neweq=app_global _eq [|intype;lhs;rhs|] in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm) (Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) - [proof_tac p; endt refine_exact_check]) + [proof_tac p; Equality.discrHyp hid]) end } (* wrap everything *) -- cgit v1.2.3 From 9b54bda96f51cc5387f195614620fae53dec5bd1 Mon Sep 17 00:00:00 2001 From: Cyprien Mangin Date: Wed, 3 May 2017 10:47:53 +0200 Subject: FIx bug #5300: uncaught exception in "Print Assumptions". --- test-suite/bugs/closed/5300.v | 39 +++++++++++++++++++++++++++++++++++++++ toplevel/assumptions.ml | 21 +++++++++++++++++++-- 2 files changed, 58 insertions(+), 2 deletions(-) create mode 100644 test-suite/bugs/closed/5300.v diff --git a/test-suite/bugs/closed/5300.v b/test-suite/bugs/closed/5300.v new file mode 100644 index 000000000..18202df50 --- /dev/null +++ b/test-suite/bugs/closed/5300.v @@ -0,0 +1,39 @@ +Module Test1. + + Module Type Foo. + Parameter t : unit. + End Foo. + + Module Bar : Foo. + Module Type Rnd. Definition t' : unit := tt. End Rnd. + Module Rnd_inst : Rnd. Definition t' : unit := tt. End Rnd_inst. + Definition t : unit. + exact Rnd_inst.t'. + Qed. + End Bar. + + Print Assumptions Bar.t. +End Test1. + +Module Test2. + Module Type Foo. + Parameter t1 : unit. + Parameter t2 : unit. + End Foo. + + Module Bar : Foo. + Inductive ind := . + Definition t' : unit := tt. + Definition t1 : unit. + Proof. + exact ((fun (_ : ind -> False) => tt) (fun H => match H with end)). + Qed. + Definition t2 : unit. + Proof. + exact t'. + Qed. + End Bar. + + Print Assumptions Bar.t1. + Print Assumptions Bar.t1. +End Test2. diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml index 45c539e22..a865ee987 100644 --- a/toplevel/assumptions.ml +++ b/toplevel/assumptions.ml @@ -42,6 +42,11 @@ let rec search_cst_label lab = function | (l, SFBconst cb) :: _ when Label.equal l lab -> cb | _ :: fields -> search_cst_label lab fields +let rec search_mind_label lab = function + | [] -> raise Not_found + | (l, SFBmind mind) :: _ when Label.equal l lab -> mind + | _ :: fields -> search_mind_label lab fields + (* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But: a) I don't see currently what should be used instead b) this shouldn't be critical for Print Assumption. At worse some @@ -133,6 +138,18 @@ let lookup_constant cst = else lookup_constant_in_impl cst (Some cb) with Not_found -> lookup_constant_in_impl cst None +let lookup_mind_in_impl mind = + try + let mp,dp,lab = repr_kn (canonical_mind mind) in + let fields = memoize_fields_of_mp mp in + search_mind_label lab fields + with Not_found -> + anomaly (str "Print Assumption: unknown inductive " ++ MutInd.print mind) + +let lookup_mind mind = + try Global.lookup_mind mind + with Not_found -> lookup_mind_in_impl mind + (** Graph traversal of an object, collecting on the way the dependencies of traversed objects *) @@ -204,7 +221,7 @@ and traverse_inductive (curr, data, ax2ty) mind obj = where I_0, I_1, ... are in the same mutual definition and c_ij are all their constructors. *) if Refmap_env.mem firstind_ref data then data, ax2ty else - let mib = Global.lookup_mind mind in + let mib = lookup_mind mind in (* Collects references of parameters *) let param_ctx = mib.mind_params_ctxt in let nparam = List.length param_ctx in @@ -308,7 +325,7 @@ let assumptions ?(add_opaque=false) ?(add_transparent=false) st gr t = else accu | IndRef (m,_) | ConstructRef ((m,_),_) -> - let mind = Global.lookup_mind m in + let mind = lookup_mind m in if mind.mind_typing_flags.check_guarded then accu else -- cgit v1.2.3 From 1fe73e6af81759fa8b78c8660745492ed886d477 Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Wed, 26 Apr 2017 11:57:58 +0200 Subject: Adding an option "Printing Unfocused". Off by default. + small refactoring of emacs hacks in printer.ml. --- ide/coq.ml | 6 ++- ide/coq.mli | 2 + ide/coqide_ui.ml | 1 + ide/ide_slave.ml | 3 +- ide/wg_ProofView.ml | 35 +++++++++---- printing/printer.ml | 139 ++++++++++++++++++++++++++++++++++----------------- printing/printer.mli | 21 +++++++- toplevel/coqtop.ml | 1 + 8 files changed, 147 insertions(+), 61 deletions(-) diff --git a/ide/coq.ml b/ide/coq.ml index 3a1d87787..cd45e2fcd 100644 --- a/ide/coq.ml +++ b/ide/coq.ml @@ -519,6 +519,7 @@ struct let all_basic = ["Printing"; "All"] let existential = ["Printing"; "Existential"; "Instances"] let universes = ["Printing"; "Universes"] + let unfocused = ["Printing"; "Unfocused"] type bool_descr = { opts : t list; init : bool; label : string } @@ -534,7 +535,8 @@ struct label = "Display _existential variable instances" }; { opts = [universes]; init = false; label = "Display _universe levels" }; { opts = [all_basic;existential;universes]; init = false; - label = "Display all _low-level contents" } + label = "Display all _low-level contents" }; + { opts = [unfocused]; init = false; label = "Display _unfocused goals" } ] (** The current status of the boolean options *) @@ -549,6 +551,8 @@ struct let _ = reset () + let printing_unfocused () = Hashtbl.find current_state unfocused + (** Transmitting options to coqtop *) let enforce h k = diff --git a/ide/coq.mli b/ide/coq.mli index ab8c12a6f..e8e2f5239 100644 --- a/ide/coq.mli +++ b/ide/coq.mli @@ -140,6 +140,8 @@ sig val set : t -> bool -> unit + val printing_unfocused: unit -> bool + (** [enforce] transmits to coq the current option values. It is also called by [goals] and [evars] above. *) diff --git a/ide/coqide_ui.ml b/ide/coqide_ui.ml index 91c281c8d..717c4000f 100644 --- a/ide/coqide_ui.ml +++ b/ide/coqide_ui.ml @@ -85,6 +85,7 @@ let init () = \n \ \n \ \n \ +\n \ \n \ \n \ \n \ diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index 966a94da9..56ada9d13 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -54,7 +54,8 @@ let coqide_known_option table = List.mem table [ ["Printing";"All"]; ["Printing";"Records"]; ["Printing";"Existential";"Instances"]; - ["Printing";"Universes"]] + ["Printing";"Universes"]; + ["Printing";"Unfocused"]] let is_known_option cmd = match cmd with | VernacSetOption (o,BoolValue true) diff --git a/ide/wg_ProofView.ml b/ide/wg_ProofView.ml index f801091cf..60b2d9e60 100644 --- a/ide/wg_ProofView.ml +++ b/ide/wg_ProofView.ml @@ -47,7 +47,7 @@ let hook_tag_cb tag menu_content sel_cb hover_cb = hover_cb start stop; false | _ -> false)) -let mode_tactic sel_cb (proof : #GText.view_skel) goals hints = match goals with +let mode_tactic sel_cb (proof : #GText.view_skel) goals unfoc_goals hints = match goals with | [] -> assert false | { Interface.goal_hyp = hyps; Interface.goal_ccl = cur_goal; } :: rem_goals -> let on_hover sel_start sel_stop = @@ -65,8 +65,11 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals hints = match goals with let head_str = Printf.sprintf "%d subgoal%s\n" goals_cnt (if 1 < goals_cnt then "s" else "") in - let goal_str index total = Printf.sprintf - "______________________________________(%d/%d)\n" index total + let goal_str shownum index total = + if shownum then Printf.sprintf + "______________________________________(%d/%d)\n" index total + else Printf.sprintf + "______________________________________\n" in (* Insert current goal and its hypotheses *) let hyps_hints, goal_hints = match hints with @@ -97,18 +100,29 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals hints = match goals with [tag] else [] in - proof#buffer#insert (goal_str 1 goals_cnt); + proof#buffer#insert (goal_str true 1 goals_cnt); insert_xml proof#buffer (Richpp.richpp_of_pp width cur_goal); proof#buffer#insert "\n" in (* Insert remaining goals (no hypotheses) *) - let fold_goal i _ { Interface.goal_ccl = g } = - proof#buffer#insert (goal_str i goals_cnt); + let fold_goal shownum i _ { Interface.goal_ccl = g } = + proof#buffer#insert (goal_str shownum i goals_cnt); insert_xml proof#buffer (Richpp.richpp_of_pp width g); proof#buffer#insert "\n" in - let () = Util.List.fold_left_i fold_goal 2 () rem_goals in - + let () = Util.List.fold_left_i (fold_goal true) 2 () rem_goals in + (* show unfocused goal if option set *) + (* Insert remaining goals (no hypotheses) *) + if Coq.PrintOpt.printing_unfocused () then + begin + ignore(proof#buffer#place_cursor ~where:(proof#buffer#end_iter)); + let unfoc = List.flatten (List.rev (List.map (fun (x,y) -> x@y) unfoc_goals)) in + if unfoc<>[] then + begin + proof#buffer#insert "\nUnfocused Goals:\n"; + Util.List.fold_left_i (fold_goal false) 0 () unfoc + end + end; ignore(proof#buffer#place_cursor ~where:(proof#buffer#end_iter#backward_to_tag_toggle (Some Tags.Proof.goal))); @@ -172,8 +186,9 @@ let display mode (view : #GText.view_skel) goals hints evars = in List.iteri iter bg end - | Some { Interface.fg_goals = fg } -> - mode view fg hints + | Some { Interface.fg_goals = fg; bg_goals = bg } -> + mode view fg bg hints + let proof_view () = let buffer = GSourceView2.source_buffer diff --git a/printing/printer.ml b/printing/printer.ml index 93d221f03..3d67e89f5 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -34,6 +34,49 @@ let delayed_emacs_cmd s = let get_current_context () = Pfedit.get_current_context () +let enable_unfocused_goal_printing = ref false +let enable_goal_tags_printing = ref false +let enable_goal_names_printing = ref false + +let should_tag() = !enable_goal_tags_printing +let should_unfoc() = !enable_unfocused_goal_printing +let should_gname() = !enable_goal_names_printing + + +let _ = + let open Goptions in + declare_bool_option + { optsync = true; + optdepr = false; + optname = "printing of unfocused goal"; + optkey = ["Printing";"Unfocused"]; + optread = (fun () -> !enable_unfocused_goal_printing); + optwrite = (fun b -> enable_unfocused_goal_printing:=b) } + +(* This is set on by proofgeneral proof-tree mode. But may be used for + other purposes *) +let _ = + let open Goptions in + declare_bool_option + { optsync = true; + optdepr = false; + optname = "printing of goal tags"; + optkey = ["Printing";"Goal";"Tags"]; + optread = (fun () -> !enable_goal_tags_printing); + optwrite = (fun b -> enable_goal_tags_printing:=b) } + + +let _ = + let open Goptions in + declare_bool_option + { optsync = true; + optdepr = false; + optname = "printing of goal names"; + optkey = ["Printing";"Goal";"Names"]; + optread = (fun () -> !enable_goal_names_printing); + optwrite = (fun b -> enable_goal_names_printing:=b) } + + (**********************************************************************) (** Terms *) @@ -419,23 +462,25 @@ let default_pr_goal gs = (* display a goal tag *) let pr_goal_tag g = let s = " (ID " ^ Goal.uid g ^ ")" in - str (emacs_str s) - -let display_name = false + str s (* display a goal name *) let pr_goal_name sigma g = - if display_name then str " " ++ Pp.surround (pr_existential_key sigma g) + if should_gname() then str " " ++ Pp.surround (pr_existential_key sigma g) else mt () +let pr_goal_header nme sigma g = + let (g,sigma) = Goal.V82.nf_evar sigma g in + str "subgoal " ++ nme ++ (if should_tag() then pr_goal_tag g else str"") + ++ (if should_gname() then str " " ++ Pp.surround (pr_existential_key sigma g) else mt ()) + (* display the conclusion of a goal *) let pr_concl n sigma g = let (g,sigma) = Goal.V82.nf_evar sigma g in let env = Goal.V82.env sigma g in let pc = pr_goal_concl_style_env env sigma (Goal.V82.concl sigma g) in - str (emacs_str "") ++ - str "subgoal " ++ int n ++ pr_goal_tag g ++ pr_goal_name sigma g ++ - str " is:" ++ cut () ++ str" " ++ pc + let header = pr_goal_header (int n) sigma g in + header ++ str " is:" ++ cut () ++ str" " ++ pc (* display evar type: a context and a type *) let pr_evgl_sign sigma evi = @@ -491,8 +536,8 @@ let pr_ne_evar_set hd tl sigma l = let pr_selected_subgoal name sigma g = let pg = default_pr_goal { sigma=sigma ; it=g; } in - v 0 (str "subgoal " ++ name ++ pr_goal_tag g ++ pr_goal_name sigma g - ++ str " is:" ++ cut () ++ pg) + let header = pr_goal_header name sigma g in + v 0 (header ++ str " is:" ++ cut () ++ pg) let default_pr_subgoal n sigma = let rec prrec p = function @@ -585,27 +630,27 @@ let print_dependent_evars gl sigma seeds = end i (str ",") end evars (str "") in - fnl () ++ - str "(dependent evars:" ++ evars ++ str ")" ++ fnl () - else - fnl () ++ - str "(dependent evars: (printing disabled) )" ++ fnl () + cut () ++ cut () ++ + str "(dependent evars:" ++ evars ++ str ")" + else if !Flags.print_emacs then + (* IDEs prefer something dummy instead of nothing *) + cut () ++ cut () ++ str "(dependent evars: (printing disabled) )" + else mt () in - constraints ++ delayed_emacs_cmd evars + constraints ++ evars () (* Print open subgoals. Checks for uninstantiated existential variables *) (* spiwack: [seeds] is for printing dependent evars in emacs mode. *) (* spiwack: [pr_first] is true when the first goal must be singled out and printed in its entirety. *) -(* courtieu: in emacs mode, even less cases where the first goal is printed - in its entirety *) -let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds shelf stack goals = +let default_pr_subgoals ?(pr_first=true) + close_cmd sigma seeds shelf stack unfocused goals = (** Printing functions for the extra informations. *) let rec print_stack a = function | [] -> Pp.int a | b::l -> Pp.int a ++ str"-" ++ print_stack b l in - let print_unfocused l = + let print_unfocused_nums l = match l with | [] -> None | a::l -> Some (str"unfocused: " ++ print_stack a l) @@ -625,7 +670,7 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds shelf stack goals | [] -> Pp.mt () | a::l -> Pp.spc () ++ str"(" ++ print_comma_separated_list a l ++ str")" in - let extra = Option.List.flatten [ print_unfocused stack ; print_shelf shelf ] in + let extra = Option.List.flatten [ print_unfocused_nums stack ; print_shelf shelf ] in let print_extra = print_extra_list extra in let focused_if_needed = let needed = not (CList.is_empty extra) && pr_first in @@ -642,8 +687,9 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds shelf stack goals in let print_multiple_goals g l = if pr_first then - default_pr_goal { it = g ; sigma = sigma; } ++ fnl () ++ - pr_rec 2 l + default_pr_goal { it = g ; sigma = sigma; } + ++ (if l=[] then str"" else cut ()) + ++ pr_rec 2 l else pr_rec 1 (g::l) in @@ -658,32 +704,27 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds shelf stack goals begin let exl = Evarutil.non_instantiated sigma in if Evar.Map.is_empty exl then - (str"No more subgoals." - ++ print_dependent_evars None sigma seeds) + (str"No more subgoals." ++ print_dependent_evars None sigma seeds) else let pei = pr_evars_int sigma 1 exl in - (str "No more subgoals, but there are non-instantiated existential variables:" - ++ fnl () ++ (hov 0 pei) - ++ print_dependent_evars None sigma seeds ++ fnl () ++ - str "You can use Grab Existential Variables.") + v 0 ((str "No more subgoals," + ++ str " but there are non-instantiated existential variables:" + ++ cut () ++ (hov 0 pei) + ++ print_dependent_evars None sigma seeds + ++ cut () ++ str "You can use Grab Existential Variables.")) end - | [g] when not !Flags.print_emacs && pr_first -> - let pg = default_pr_goal { it = g ; sigma = sigma; } in - v 0 ( - str "1" ++ focused_if_needed ++ str"subgoal" ++ print_extra - ++ pr_goal_tag g ++ pr_goal_name sigma g ++ cut () ++ pg - ++ print_dependent_evars (Some g) sigma seeds - ) | g1::rest -> let goals = print_multiple_goals g1 rest in let ngoals = List.length rest+1 in v 0 ( - int ngoals ++ focused_if_needed ++ str(String.plural ngoals "subgoal") ++ - print_extra ++ - str ((if display_name then (fun x -> x) else emacs_str) ", subgoal 1") - ++ pr_goal_tag g1 - ++ pr_goal_name sigma g1 ++ cut () - ++ goals + int ngoals ++ focused_if_needed ++ str(String.plural ngoals "subgoal") + ++ print_extra + ++ str (if (should_gname()) then ", subgoal 1" else "") + ++ (if should_tag() then pr_goal_tag g1 else str"") + ++ pr_goal_name sigma g1 ++ cut () ++ goals + ++ (if unfocused=[] then str "" + else (cut() ++ cut() ++ str "*** Unfocused goals:" ++ cut() + ++ pr_rec (List.length rest + 2) unfocused)) ++ print_dependent_evars (Some g1) sigma seeds ) @@ -692,7 +733,7 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds shelf stack goals type printer_pr = { - pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> std_ppcmds; + pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> goal list -> std_ppcmds; pr_subgoal : int -> evar_map -> goal list -> std_ppcmds; pr_goal : goal sigma -> std_ppcmds; } @@ -726,16 +767,16 @@ let pr_open_subgoals ?(proof=Proof_global.give_me_the_proof ()) () = begin match goals with | [] -> let { Evd.it = bgoals ; sigma = bsigma } = Proof.V82.background_subgoals p in begin match bgoals,shelf,given_up with - | [] , [] , [] -> pr_subgoals None sigma seeds shelf stack goals + | [] , [] , [] -> pr_subgoals None sigma seeds shelf stack [] goals | [] , [] , _ -> Feedback.msg_info (str "No more subgoals, but there are some goals you gave up:"); fnl () - ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] given_up + ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] [] given_up ++ fnl () ++ str "You need to go back and solve them." | [] , _ , _ -> Feedback.msg_info (str "All the remaining goals are on the shelf."); fnl () - ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] shelf + ++ pr_subgoals ~pr_first:false None bsigma seeds [] [] [] shelf | _ , _, _ -> let end_cmd = str "This subproof is complete, but there are some unfocused goals." ++ @@ -743,9 +784,13 @@ let pr_open_subgoals ?(proof=Proof_global.give_me_the_proof ()) () = if Pp.ismt s then s else fnl () ++ s) ++ fnl () in - pr_subgoals ~pr_first:false (Some end_cmd) bsigma seeds shelf [] bgoals + pr_subgoals ~pr_first:false (Some end_cmd) bsigma seeds shelf [] [] bgoals end - | _ -> pr_subgoals None sigma seeds shelf stack goals + | _ -> + let { Evd.it = bgoals ; sigma = bsigma } = Proof.V82.background_subgoals p in + let bgoals_focused, bgoals_unfocused = List.partition (fun x -> List.mem x goals) bgoals in + let unfocused_if_needed = if should_unfoc() then bgoals_unfocused else [] in + pr_subgoals ~pr_first:true None bsigma seeds shelf [] unfocused_if_needed bgoals_focused end let pr_nth_open_subgoal n = diff --git a/printing/printer.mli b/printing/printer.mli index 504392e35..c28295054 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -18,6 +18,11 @@ open Glob_term (** These are the entry points for printing terms, context, tac, ... *) + +val enable_unfocused_goal_printing: bool ref +val enable_goal_tags_printing : bool ref +val enable_goal_names_printing : bool ref + (** Terms *) val pr_lconstr_env : env -> evar_map -> constr -> std_ppcmds @@ -135,7 +140,19 @@ val pr_transparent_state : transparent_state -> std_ppcmds (** Proofs *) val pr_goal : goal sigma -> std_ppcmds -val pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> std_ppcmds + +(** [pr_subgoals ~pr_first pp sigma seeds shelf focus_stack unfocused goals] + prints the goals of the list [goals] followed by the goals in + [unfocused], in a short way (typically only the conclusion) except + for the first goal if [pr_first] is true. This function can be + replaced by another one by calling [set_printer_pr] (see below), + typically by plugin writers. The default printer prints only the + focused goals unless the conrresponding option + [enable_unfocused_goal_printing] is set. [seeds] is for printing + dependent evars (mainly for emacs proof tree mode). *) +val pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list + -> goal list -> goal list -> std_ppcmds + val pr_subgoal : int -> evar_map -> goal list -> std_ppcmds val pr_concl : int -> evar_map -> goal -> std_ppcmds @@ -190,7 +207,7 @@ val pr_goal_by_id : Id.t -> std_ppcmds val pr_goal_by_uid : string -> std_ppcmds type printer_pr = { - pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> std_ppcmds; + pr_subgoals : ?pr_first:bool -> std_ppcmds option -> evar_map -> evar list -> Goal.goal list -> int list -> goal list -> goal list -> std_ppcmds; pr_subgoal : int -> evar_map -> goal list -> std_ppcmds; pr_goal : goal sigma -> std_ppcmds; };; diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index c3a755860..5687418f2 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -247,6 +247,7 @@ let set_emacs () = if not (Option.is_empty !toploop) then error "Flag -emacs is incompatible with a custom toplevel loop"; Flags.print_emacs := true; + Printer.enable_goal_tags_printing := true; color := `OFF (** Options for CoqIDE *) -- cgit v1.2.3 From 8bcf71360261d789ac3ab919bc49309df678628f Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Tue, 2 May 2017 21:32:26 +0200 Subject: labelizing arguments --- ide/wg_ProofView.ml | 16 ++++++++-------- printing/printer.ml | 2 +- 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/ide/wg_ProofView.ml b/ide/wg_ProofView.ml index 60b2d9e60..0bf5afbfd 100644 --- a/ide/wg_ProofView.ml +++ b/ide/wg_ProofView.ml @@ -47,7 +47,7 @@ let hook_tag_cb tag menu_content sel_cb hover_cb = hover_cb start stop; false | _ -> false)) -let mode_tactic sel_cb (proof : #GText.view_skel) goals unfoc_goals hints = match goals with +let mode_tactic sel_cb (proof : #GText.view_skel) goals ~unfoc_goals hints = match goals with | [] -> assert false | { Interface.goal_hyp = hyps; Interface.goal_ccl = cur_goal; } :: rem_goals -> let on_hover sel_start sel_stop = @@ -65,7 +65,7 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals unfoc_goals hints = matc let head_str = Printf.sprintf "%d subgoal%s\n" goals_cnt (if 1 < goals_cnt then "s" else "") in - let goal_str shownum index total = + let goal_str ?(shownum=false) index total = if shownum then Printf.sprintf "______________________________________(%d/%d)\n" index total else Printf.sprintf @@ -100,17 +100,17 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals unfoc_goals hints = matc [tag] else [] in - proof#buffer#insert (goal_str true 1 goals_cnt); + proof#buffer#insert (goal_str ~shownum:true 1 goals_cnt); insert_xml proof#buffer (Richpp.richpp_of_pp width cur_goal); proof#buffer#insert "\n" in (* Insert remaining goals (no hypotheses) *) - let fold_goal shownum i _ { Interface.goal_ccl = g } = - proof#buffer#insert (goal_str shownum i goals_cnt); + let fold_goal ?(shownum=false) i _ { Interface.goal_ccl = g } = + proof#buffer#insert (goal_str ~shownum i goals_cnt); insert_xml proof#buffer (Richpp.richpp_of_pp width g); proof#buffer#insert "\n" in - let () = Util.List.fold_left_i (fold_goal true) 2 () rem_goals in + let () = Util.List.fold_left_i (fold_goal ~shownum:true) 2 () rem_goals in (* show unfocused goal if option set *) (* Insert remaining goals (no hypotheses) *) if Coq.PrintOpt.printing_unfocused () then @@ -120,7 +120,7 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals unfoc_goals hints = matc if unfoc<>[] then begin proof#buffer#insert "\nUnfocused Goals:\n"; - Util.List.fold_left_i (fold_goal false) 0 () unfoc + Util.List.fold_left_i (fold_goal ~shownum:false) 0 () unfoc end end; ignore(proof#buffer#place_cursor @@ -187,7 +187,7 @@ let display mode (view : #GText.view_skel) goals hints evars = List.iteri iter bg end | Some { Interface.fg_goals = fg; bg_goals = bg } -> - mode view fg bg hints + mode view fg ~unfoc_goals:bg hints let proof_view () = diff --git a/printing/printer.ml b/printing/printer.ml index 3d67e89f5..1d3e67d9d 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -688,7 +688,7 @@ let default_pr_subgoals ?(pr_first=true) let print_multiple_goals g l = if pr_first then default_pr_goal { it = g ; sigma = sigma; } - ++ (if l=[] then str"" else cut ()) + ++ (if l=[] then mt () else cut ()) ++ pr_rec 2 l else pr_rec 1 (g::l) -- cgit v1.2.3 From 8bdb00bbea9dce95f7e0bc18250a41545660299d Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Thu, 4 May 2017 11:25:22 +0200 Subject: Warning removed. --- printing/printer.ml | 2 -- 1 file changed, 2 deletions(-) diff --git a/printing/printer.ml b/printing/printer.ml index 1d3e67d9d..91a7d2289 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -28,8 +28,6 @@ module CompactedDecl = Context.Compacted.Declaration let emacs_str s = if !Flags.print_emacs then s else "" -let delayed_emacs_cmd s = - if !Flags.print_emacs then s () else str "" let get_current_context () = Pfedit.get_current_context () -- cgit v1.2.3 From 4cc655377e3c73fd3066cd8136d17605f167ef56 Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Thu, 4 May 2017 10:28:56 +0200 Subject: Improve documentation of assert / pose proof / specialize. This commits documents the as clause of specialize and that the as clause of pose proof is optional. It also mentions a feature of assert ( := ) that was available since 8.5 and was mentionned by @herbelin in: https://github.com/coq/coq/pull/248#issuecomment-297970503 --- doc/refman/RefMan-tac.tex | 28 ++++++++++++++++++---------- 1 file changed, 18 insertions(+), 10 deletions(-) diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex index 0edc66f83..87b9e4914 100644 --- a/doc/refman/RefMan-tac.tex +++ b/doc/refman/RefMan-tac.tex @@ -1275,15 +1275,18 @@ in the list of subgoals remaining to prove. \item{\tt assert ( {\ident} := {\term} )} - This behaves as {\tt assert ({\ident} :\ {\type});[exact - {\term}|idtac]} where {\type} is the type of {\term}. This is - deprecated in favor of {\tt pose proof}. + This behaves as {\tt assert ({\ident} :\ {\type}) by exact {\term}} + where {\type} is the type of {\term}. This is deprecated in favor of + {\tt pose proof}. + + If the head of {\term} is {\ident}, the tactic behaves as + {\tt specialize \term}. \ErrMsg \errindex{Variable {\ident} is already declared} -\item \texttt{pose proof {\term} as {\intropattern}\tacindex{pose proof}} +\item \texttt{pose proof {\term} \zeroone{as {\intropattern}}\tacindex{pose proof}} - This tactic behaves like \texttt{assert T as {\intropattern} by + This tactic behaves like \texttt{assert T \zeroone{as {\intropattern}} by exact {\term}} where \texttt{T} is the type of {\term}. In particular, \texttt{pose proof {\term} as {\ident}} behaves as @@ -1326,8 +1329,8 @@ in the list of subgoals remaining to prove. following subgoals: {\tt U -> T} and \texttt{U}. The subgoal {\tt U -> T} comes first in the list of remaining subgoal to prove. -\item {\tt specialize ({\ident} \term$_1$ \dots\ \term$_n$)\tacindex{specialize}} \\ - {\tt specialize {\ident} with \bindinglist} +\item {\tt specialize ({\ident} \term$_1$ \dots\ \term$_n$)\tacindex{specialize} \zeroone{as \intropattern}}\\ + {\tt specialize {\ident} with {\bindinglist} \zeroone{as \intropattern}} The tactic {\tt specialize} works on local hypothesis \ident. The premises of this hypothesis (either universal @@ -1338,14 +1341,19 @@ in the list of subgoals remaining to prove. second form, all instantiation elements must be given, whereas in the first form the application to \term$_1$ {\ldots} \term$_n$ can be partial. The first form is equivalent to - {\tt assert (\ident' := {\ident} {\term$_1$} \dots\ \term$_n$); - clear \ident; rename \ident' into \ident}. + {\tt assert ({\ident} := {\ident} {\term$_1$} \dots\ \term$_n$)}. + + With the {\tt as} clause, the local hypothesis {\ident} is left + unchanged and instead, the modified hypothesis is introduced as + specified by the {\intropattern}. The name {\ident} can also refer to a global lemma or hypothesis. In this case, for compatibility reasons, the behavior of {\tt specialize} is close to that of {\tt generalize}: the instantiated statement becomes an additional - premise of the goal. + premise of the goal. The {\tt as} clause is especially useful + in this case to immediately introduce the instantiated statement + as a local hypothesis. \begin{ErrMsgs} \item \errindexbis{{\ident} is used in hypothesis \ident'}{is used in hypothesis} -- cgit v1.2.3 From 8f4d79115c3686de41e20c41ef6dbce8b8546366 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 25 Feb 2017 18:51:36 +0100 Subject: Adding an option "Set Ltac Batch Debug" to additionally run Ltac debug in batch mode. --- plugins/ltac/tactic_debug.ml | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml index dffeade29..dac15ff79 100644 --- a/plugins/ltac/tactic_debug.ml +++ b/plugins/ltac/tactic_debug.ml @@ -85,6 +85,19 @@ let skipped = Proofview.NonLogical.run (Proofview.NonLogical.ref 0) let skip = Proofview.NonLogical.run (Proofview.NonLogical.ref 0) let breakpoint = Proofview.NonLogical.run (Proofview.NonLogical.ref None) +let batch = ref false + +open Goptions + +let _ = + declare_bool_option + { optsync = false; + optdepr = false; + optname = "Ltac batch debug"; + optkey = ["Ltac";"Batch";"Debug"]; + optread = (fun () -> !batch); + optwrite = (fun x -> batch := x) } + let rec drop_spaces inst i = if String.length inst > i && inst.[i] == ' ' then drop_spaces inst (i+1) else i @@ -150,6 +163,7 @@ let rec prompt level = begin let open Proofview.NonLogical in Proofview.NonLogical.print_notice (fnl () ++ str "TcDebug (" ++ int level ++ str ") > ") >> + if Pervasives.(!batch) then return (DebugOn (level+1)) else let exit = (skip:=0) >> (skipped:=0) >> raise Sys.Break in Proofview.NonLogical.catch Proofview.NonLogical.read_line begin function (e, info) -> match e with -- cgit v1.2.3 From 91d78c590b35c9edcf9f68c408ba82090f68e156 Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Fri, 5 May 2017 12:59:41 +0200 Subject: Adapted to EConstr. --- printing/printer.ml | 12 ++++++------ vernac/vernacentries.ml | 9 --------- 2 files changed, 6 insertions(+), 15 deletions(-) diff --git a/printing/printer.ml b/printing/printer.ml index 0e31a4a04..997d866f9 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -359,14 +359,14 @@ let rec should_compact env sigma typ = match kind_of_term typ with | Rel _ | Var _ | Sort _ | Const _ | Ind _ -> true | App (c,args) -> - let _,type_of_c = Typing.type_of env sigma c in - let _,type_of_typ = Typing.type_of env sigma typ in - not (is_Prop type_of_typ) + let _,type_of_c = Typing.type_of env sigma (EConstr.of_constr c) in + let _,type_of_typ = Typing.type_of env sigma (EConstr.of_constr typ) in + not (is_Prop (EConstr.to_constr sigma type_of_typ)) && (* These two more tests detect rare cases of non-Prop-sorted dependent hypothesis: *) - let lnamedtyp , _ = decompose_prod type_of_c in + let lnamedtyp , _ = EConstr.decompose_prod sigma type_of_c in (* c has a non dependent type *) - List.for_all (fun (_,typarg) -> isSort typarg) lnamedtyp + List.for_all (fun (_,typarg) -> EConstr.isSort sigma typarg) lnamedtyp && (* and real arguments are recursively elligible to compaction. *) Array.for_all (should_compact env sigma) args | _ -> false @@ -389,7 +389,7 @@ let rec bld_sign_env env sigma ctxt pps = and bld_sign_env_id env sigma ctxt pps is_start = match ctxt with | [] -> pps,ctxt - | CompactedDecl.LocalAssum(ids,typ) as d :: ctxt' when should_compact env sigma typ -> + | CompactedDecl.LocalAssum(ids,typ) as d :: ctxt' when should_compact env sigma typ -> let pidt = pr_var_list_decl env sigma d in let pps' = pps ++ (if not is_start then brk (3,0) else (mt ())) ++ pidt in bld_sign_env_id env sigma ctxt' pps' false diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 40cd1a29c..d4e6af995 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1411,15 +1411,6 @@ let _ = optread = (fun _ -> None); optwrite = (fun _ -> ()) } -let _ = - declare_int_option - { optsync = false; - optdepr = false; - optname = "the hypotheses limit"; - optkey = ["Hyps";"Limit"]; - optread = Flags.print_hyps_limit; - optwrite = Flags.set_print_hyps_limit } - let _ = declare_bool_option { optsync = true; -- cgit v1.2.3 From 5548e5f6bc5446f7541cfc7d93b0b47e4b751e86 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 5 May 2017 15:39:17 +0200 Subject: Remove unused open. --- tactics/eqdecide.ml | 1 - 1 file changed, 1 deletion(-) diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index 472cd8f22..641929a77 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -67,7 +67,6 @@ let choose_noteq eqonleft = left_with_bindings false Misctypes.NoBindings open Sigma.Notations -open Context.Rel.Declaration (* A surgical generalize which selects the right occurrences by hand *) (* This prevents issues where c2 is also a subterm of c1 (see e.g. #5449) *) -- cgit v1.2.3 From cf0e030f3aebecb316852fc4a152fb212f9e7ef5 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 5 May 2017 15:58:52 +0200 Subject: coqtop -help: don't die if coqlib can't be found --- toplevel/coqtop.ml | 10 +++++++++- toplevel/usage.ml | 1 + 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index 5687418f2..8f50bfb3d 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -292,9 +292,17 @@ let init_gc () = We no longer use [Arg.parse], in order to use share [Usage.print_usage] between coqtop and coqc. *) +let usage_no_coqlib = CWarnings.create ~name:"usage-no-coqlib" ~category:"filesystem" + (fun () -> Pp.str "cannot guess a path for Coq libraries; dynaminally loaded flags will not be mentioned") + +exception NoCoqLib let usage () = - Envars.set_coqlib ~fail:CErrors.error; + begin + try + Envars.set_coqlib ~fail:(fun x -> raise NoCoqLib); init_load_path (); + with NoCoqLib -> usage_no_coqlib () + end; if !batch_mode then Usage.print_usage_coqc () else begin Mltop.load_ml_objects_raw_rex diff --git a/toplevel/usage.ml b/toplevel/usage.ml index e457ca61d..e29048035 100644 --- a/toplevel/usage.ml +++ b/toplevel/usage.ml @@ -30,6 +30,7 @@ let print_usage_channel co command = \n -R dir coqdir recursively map physical dir to logical coqdir\ \n -Q dir coqdir map physical dir to logical coqdir\ \n -top coqdir set the toplevel name to be coqdir instead of Top\ +\n -coqlib dir set the coq standard library directory\ \n -exclude-dir f exclude subdirectories named f for option -R\ \n\ \n -noinit start without loading the Init library\ -- cgit v1.2.3 From cff6f53cbef53ce3902e59853f7a7dc9b7150f45 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 1 May 2017 13:25:10 +0200 Subject: Adding a test-suite pattern-unification example that Econstr fixed. --- test-suite/success/unification.v | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/test-suite/success/unification.v b/test-suite/success/unification.v index 296686e16..6f7498d65 100644 --- a/test-suite/success/unification.v +++ b/test-suite/success/unification.v @@ -188,3 +188,14 @@ Proof. apply idpath. apply idpath. Defined. + +(* An example where it is necessary to evar-normalize the instance of + an evar to evaluate if it is a pattern *) + +Check + let a := ?[P] in + fun (H : forall y (P : nat -> Prop), y = 0 -> P y) + x (p:x=0) => + H ?[y] a p : x = 0. +(* We have to solve "?P ?y[x] == x = 0" knowing from + "p : (x=0) == (?y[x] = 0)" that "?y := x" *) -- cgit v1.2.3 From 7e28feadd6394483b6f527d5aed7d663e189596e Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Tue, 12 Jul 2016 23:26:44 +0200 Subject: Upgrading some local function as a general-purpose combinator Option.List.map. --- lib/option.ml | 10 ++++++++++ lib/option.mli | 6 ++++++ pretyping/evarsolve.ml | 18 ++++-------------- 3 files changed, 20 insertions(+), 14 deletions(-) diff --git a/lib/option.ml b/lib/option.ml index fbb883d30..f3047a3ca 100644 --- a/lib/option.ml +++ b/lib/option.ml @@ -188,4 +188,14 @@ module List = |None -> find f t |x -> x + let map f l = + let rec aux f l = match l with + | [] -> [] + | x :: l -> + match f x with + | None -> raise Exit + | Some y -> y :: aux f l + in + try Some (aux f l) with Exit -> None + end diff --git a/lib/option.mli b/lib/option.mli index 5e085620e..a336c895c 100644 --- a/lib/option.mli +++ b/lib/option.mli @@ -123,4 +123,10 @@ module List : sig val flatten : 'a option list -> 'a list val find : ('a -> 'b option) -> 'a list -> 'b option + + (** [List.map f [a1;...;an]] is the list [Some [b1;...;bn]] if + for all i, there is a [bi] such that [f ai] is [Some bi]; it is + [None] if, for at least one i, [f ai] is [None]. *) + val map : ('a -> 'b option) -> 'a list -> 'b list option + end diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index f0d011477..4ada91eb5 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -470,23 +470,13 @@ let free_vars_and_rels_up_alias_expansion sigma aliases c = (* Managing pattern-unification *) (********************************) -let map_all f l = - let rec map_aux f l = match l with - | [] -> [] - | x :: l -> - match f x with - | None -> raise Exit - | Some y -> y :: map_aux f l - in - try Some (map_aux f l) with Exit -> None - let expand_and_check_vars sigma aliases l = let map a = match get_alias_chain_of sigma aliases a with | None, [] -> Some a | None, a :: _ -> Some a | Some _, _ -> None in - map_all map l + Option.List.map map l let alias_distinct l = let rec check (rels, vars) = function @@ -540,7 +530,7 @@ let is_unification_pattern_meta env evd nb m l t = | Rel n -> if n <= nb then Some (RelAlias n) else None | _ -> None in - match map_all map l with + match Option.List.map map l with | Some l -> begin match find_unification_pattern_args env evd l t with | Some _ as x when not (dependent evd (mkMeta m) t) -> x @@ -550,10 +540,10 @@ let is_unification_pattern_meta env evd nb m l t = None let is_unification_pattern_evar env evd (evk,args) l t = - match map_all (fun c -> to_alias evd c) l with + match Option.List.map (fun c -> to_alias evd c) l with | Some l when noccur_evar env evd evk t -> let args = remove_instance_local_defs evd evk args in - let args = map_all (fun c -> to_alias evd c) args in + let args = Option.List.map (fun c -> to_alias evd c) args in begin match args with | None -> None | Some args -> -- cgit v1.2.3 From 3f4c29f8258d627c7dbacbf1157825d96b146a6d Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 5 May 2017 17:34:43 +0200 Subject: Cosmetic: unifying style within option.ml. --- lib/option.ml | 34 ++++++++++++++++------------------ 1 file changed, 16 insertions(+), 18 deletions(-) diff --git a/lib/option.ml b/lib/option.ml index f3047a3ca..50fdd079d 100644 --- a/lib/option.ml +++ b/lib/option.ml @@ -20,24 +20,24 @@ let has_some = function | _ -> true let is_empty = function -| None -> true -| Some _ -> false + | None -> true + | Some _ -> false (** Lifting equality onto option types. *) let equal f x y = match x, y with -| None, None -> true -| Some x, Some y -> f x y -| _, _ -> false + | None, None -> true + | Some x, Some y -> f x y + | _, _ -> false let compare f x y = match x, y with -| None, None -> 0 -| Some x, Some y -> f x y -| None, Some _ -> -1 -| Some _, None -> 1 + | None, None -> 0 + | Some x, Some y -> f x y + | None, Some _ -> -1 + | Some _, None -> 1 let hash f = function -| None -> 0 -| Some x -> f x + | None -> 0 + | Some x -> f x exception IsNone @@ -57,13 +57,11 @@ let init b x = else None - (** [flatten x] is [Some y] if [x] is [Some (Some y)] and [None] otherwise. *) let flatten = function | Some (Some y) -> Some y | _ -> None - (** [append x y] is the first element of the concatenation of [x] and [y] seen as lists. *) let append o1 o2 = @@ -134,6 +132,7 @@ let cata f a = function | Some c -> f c | None -> a + (** {6 More Specific operations} ***) (** [default a x] is [y] if [x] is [Some y] and [a] otherwise. *) @@ -165,7 +164,6 @@ let lift2 f x y = | _,_ -> None - (** {6 Operations with Lists} *) module List = @@ -183,10 +181,10 @@ module List = | [] -> [] let rec find f = function - |[] -> None - |h :: t -> match f h with - |None -> find f t - |x -> x + | [] -> None + | h :: t -> match f h with + | None -> find f t + | x -> x let map f l = let rec aux f l = match l with -- cgit v1.2.3 From f24d8876837e2f9121064f496d89803f60ec2c71 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 5 May 2017 17:41:00 +0200 Subject: Documenting Option.List.find. --- lib/option.mli | 3 +++ 1 file changed, 3 insertions(+) diff --git a/lib/option.mli b/lib/option.mli index a336c895c..f06ad9f1d 100644 --- a/lib/option.mli +++ b/lib/option.mli @@ -122,6 +122,9 @@ module List : sig [Some y] (in the same order). *) val flatten : 'a option list -> 'a list + (** [List.find f l] is the first [f a] different from [None], + scrolling through elements [a] of [l] in left-to-right order; + it is [None] if no such element exists. *) val find : ('a -> 'b option) -> 'a list -> 'b option (** [List.map f [a1;...;an]] is the list [Some [b1;...;bn]] if -- cgit v1.2.3 From 787a2c911b88bff399b37bfa8e11c0f1acecc6cd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 4 May 2017 08:50:27 +0200 Subject: Fix bug #3659: -time should understand multibyte encodings. We assume Coq always outputs UTF-8 (is it really the case?) and cut strings after 30 UTF-8 characters instead of using the standard String function. --- toplevel/vernac.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index bfdae85d5..2a599444c 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -160,7 +160,9 @@ let pr_new_syntax po loc chan_beautify ocom = and a glimpse of the executed command *) let pp_cmd_header loc com = - let shorten s = try (String.sub s 0 30)^"..." with _ -> s in + let shorten s = + if Unicode.utf8_length s > 33 then (Unicode.utf8_sub s 0 30) ^ "..." else s + in let noblank s = for i = 0 to String.length s - 1 do match s.[i] with -- cgit v1.2.3 From a76a71eda10b054b04de4ff56f0637a32077edd4 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 5 May 2017 18:09:30 +0200 Subject: Remove two unused opens. --- tactics/autorewrite.ml | 2 -- 1 file changed, 2 deletions(-) diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 3c430cb17..2d54b61c7 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -9,8 +9,6 @@ open Equality open Names open Pp -open Tacticals -open Tactics open Term open Termops open CErrors -- cgit v1.2.3 From e5bf991cd1094ff1d5bc2f121bb6e85c8b1320c0 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 5 May 2017 18:12:55 +0200 Subject: Remove dead code and unused open. --- interp/implicit_quantifiers.ml | 6 ------ pretyping/glob_ops.ml | 5 ----- 2 files changed, 11 deletions(-) diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml index d6749e918..19c872b31 100644 --- a/interp/implicit_quantifiers.ml +++ b/interp/implicit_quantifiers.ml @@ -19,7 +19,6 @@ open Typeclasses_errors open Pp open Libobject open Nameops -open Misctypes open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration @@ -119,11 +118,6 @@ let free_vars_of_binders ?(bound=Id.Set.empty) l (binders : local_binder_expr li | [] -> bdvars, l in aux bound l binders -let add_name_to_ids set na = - match na with - | Anonymous -> set - | Name id -> Id.Set.add id set - let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.empty) = let rec vars bound vs = function | GVar (loc,id) -> diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 080ec5ed1..6509aaac3 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -266,11 +266,6 @@ let occur_glob_constr id = fold_glob_constr_with_binders g f barred acc c in occur false false -let add_name_to_ids set na = - match na with - | Anonymous -> set - | Name id -> Id.Set.add id set - let free_glob_vars = let rec vars bound vs = function | GVar (loc,id') -> if Id.Set.mem id' bound then vs else Id.Set.add id' vs -- cgit v1.2.3 From d411a796341a138cacd72350715871f48f82920b Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Fri, 5 May 2017 21:31:14 +0200 Subject: [toplevel] Fix a couple of logical errors in error printing. In 4abb41d008fc754f21916dcac9cce49f2d04dd6d we switched back to use exceptions for error printing. However a couple of mistakes were present in that commit: - We wrongly absorbed the exception on `Vernac.compile`. However, it should be propagated as the caller will correctly print the error now. This introduced a critical bug as now `coqc` return the wrong exit status on error, breaking all sort of things. - We printed parsing exceptions twice; now it is not necessary to print the exception in the parsing handler as it will be propagated. I've checked this commit versus all previously reported bugs and it seems to work; we should definitively add a test-suite case to check that the exit code of `coqc` is correct, plus several other cases such as bugs https://coq.inria.fr/bugs/show_bug.cgi?id=5467 https://coq.inria.fr/bugs/show_bug.cgi?id=5485 https://coq.inria.fr/bugs/show_bug.cgi?id=5484 etc... See also PR #583 --- toplevel/coqloop.ml | 8 ++++++-- toplevel/vernac.ml | 5 +---- 2 files changed, 7 insertions(+), 6 deletions(-) diff --git a/toplevel/coqloop.ml b/toplevel/coqloop.ml index 9a4f476bd..a80599cd5 100644 --- a/toplevel/coqloop.ml +++ b/toplevel/coqloop.ml @@ -173,12 +173,13 @@ let print_error_for_buffer ?loc lvl msg buf = then Topfmt.emacs_logger ?pre_hdr lvl msg else Topfmt.std_logger ?pre_hdr lvl msg +(* let print_toplevel_parse_error (e, info) buf = let loc = Loc.get_loc info in let lvl = Feedback.Error in let msg = CErrors.iprint (e, info) in print_error_for_buffer ?loc lvl msg buf - +*) end (*s The Coq prompt is the name of the focused proof, if any, and "Coq" @@ -260,7 +261,10 @@ let read_sentence sid input = with reraise -> let reraise = CErrors.push reraise in discard_to_dot (); - TopErr.print_toplevel_parse_error reraise top_buffer; + (* The caller of read_sentence does the error printing now, this + should be re-enabled once we rely on the feedback error + printer again *) + (* TopErr.print_toplevel_parse_error reraise top_buffer; *) Exninfo.iraise reraise (** Coqloop Console feedback handler *) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 9f6f77ef1..4fca4ea18 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -323,8 +323,5 @@ let compile verbosely f = let compile v f = ignore(CoqworkmgrApi.get 1); - begin - try compile v f - with any -> Topfmt.print_err_exn any - end; + compile v f; CoqworkmgrApi.giveback 1 -- cgit v1.2.3 From 75a808f893a68eaf82296535e0d168e0f09f8193 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Mon, 8 May 2017 15:26:05 +0200 Subject: Fix warnings in top_printers Note that [@@@ocaml.warning "-32"] caused an error with Drop. It was put there because I didn't realise the warning was about a real issue. --- dev/top_printers.ml | 11 ----------- 1 file changed, 11 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 7caaf2d9d..f8498c402 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -7,7 +7,6 @@ (************************************************************************) (* Printers for the ocaml toplevel. *) -[@@@ocaml.warning "-32"] open Util open Pp @@ -60,14 +59,6 @@ let pprecarg = function str "Imbr[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]" let ppwf_paths x = pp (Rtree.pp_tree pprecarg x) -let pprecarg = function - | Declarations.Norec -> str "Norec" - | Declarations.Mrec (mind,i) -> - str "Mrec[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]" - | Declarations.Imbr (mind,i) -> - str "Imbr[" ++ MutInd.print mind ++ pr_comma () ++ int i ++ str "]" -let ppwf_paths x = pp (Rtree.pp_tree pprecarg x) - (* term printers *) let rawdebug = ref false let ppevar evk = pp (str (Evd.string_of_existential evk)) @@ -458,8 +449,6 @@ let print_pure_constr csr = print_string (Printexc.to_string e);print_flush (); raise e -let ppfconstr c = ppconstr (CClosure.term_of_fconstr c) - let pploc x = let (l,r) = Loc.unloc x in print_string"(";print_int l;print_string",";print_int r;print_string")" -- cgit v1.2.3 From 6532ee2d63c6e96b930895b0d42c8d9c7cc9e911 Mon Sep 17 00:00:00 2001 From: Guillaume Melquiond Date: Tue, 9 May 2017 00:26:56 +0200 Subject: Prevent user-defined ring morphisms from ever being evaluated. --- plugins/setoid_ring/newring.ml | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index d59ffee54..6b8ef630a 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -333,12 +333,12 @@ let _ = add_map "ring" my_reference "gen_phiZ", (function _->Eval); (* Pphi_dev: evaluate polynomial and coef operations, protect ring operations and make recursive call on the var map *) - pol_cst "Pphi_dev", (function -1|8|9|10|11|12|14->Eval|13->Rec|_->Prot); + pol_cst "Pphi_dev", (function -1|8|9|10|12|14->Eval|11|13->Rec|_->Prot); pol_cst "Pphi_pow", (function -1|8|9|10|13|15|17->Eval|11|16->Rec|_->Prot); - (* PEeval: evaluate morphism and polynomial, protect ring + (* PEeval: evaluate polynomial, protect ring operations and make recursive call on the var map *) - pol_cst "PEeval", (function -1|8|10|13->Eval|12->Rec|_->Prot)]) + pol_cst "PEeval", (function -1|10|13->Eval|8|12->Rec|_->Prot)]) (****************************************************************************) (* Ring database *) @@ -780,20 +780,20 @@ let _ = add_map "field" (* display_linear: evaluate polynomials and coef operations, protect field operations and make recursive call on the var map *) my_reference "display_linear", - (function -1|9|10|11|12|13|15|16->Eval|14->Rec|_->Prot); + (function -1|9|10|11|13|15|16->Eval|12|14->Rec|_->Prot); my_reference "display_pow_linear", (function -1|9|10|11|14|16|18|19->Eval|12|17->Rec|_->Prot); (* Pphi_dev: evaluate polynomial and coef operations, protect ring operations and make recursive call on the var map *) - pol_cst "Pphi_dev", (function -1|8|9|10|11|12|14->Eval|13->Rec|_->Prot); + pol_cst "Pphi_dev", (function -1|8|9|10|12|14->Eval|11|13->Rec|_->Prot); pol_cst "Pphi_pow", - (function -1|8|9|10|11|13|15|17->Eval|16->Rec|_->Prot); - (* PEeval: evaluate morphism and polynomial, protect ring + (function -1|8|9|10|13|15|17->Eval|11|16->Rec|_->Prot); + (* PEeval: evaluate polynomial, protect ring operations and make recursive call on the var map *) - pol_cst "PEeval", (function -1|8|10|13->Eval|12->Rec|_->Prot); - (* FEeval: evaluate morphism, protect field + pol_cst "PEeval", (function -1|10|13->Eval|8|12->Rec|_->Prot); + (* FEeval: evaluate polynomial, protect field operations and make recursive call on the var map *) - my_reference "FEeval", (function -1|10|12|15->Eval|14->Rec|_->Prot)]);; + my_reference "FEeval", (function -1|12|15->Eval|10|14->Rec|_->Prot)]);; let _ = add_map "field_cond" (map_without_eq -- cgit v1.2.3 From 73e09aa251a1bb3dea5b4497d3865a21b2ff3be6 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 9 May 2017 11:19:29 -0700 Subject: Put .travis.yml in alphabetical order --- .travis.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.travis.yml b/.travis.yml index 3ed547bb1..270bdcaed 100644 --- a/.travis.yml +++ b/.travis.yml @@ -36,12 +36,12 @@ env: - TEST_TARGET="ci-fiat-crypto" - TEST_TARGET="ci-fiat-parsers" - TEST_TARGET="ci-flocq" + - TEST_TARGET="ci-formal-topology" - TEST_TARGET="ci-hott" - TEST_TARGET="ci-iris-coq" - TEST_TARGET="ci-math-classes" - TEST_TARGET="ci-math-comp" - TEST_TARGET="ci-sf" - - TEST_TARGET="ci-formal-topology" - TEST_TARGET="ci-unimath" - TEST_TARGET="ci-vst" # Not ready yet for 8.7 -- cgit v1.2.3 From 5e690404233e6c772c1d5ddc52142edf474953ac Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Tue, 9 May 2017 21:47:12 +0200 Subject: Cleaning old untested not any more interesting testing files. --- test-suite/bench/lists-100.v | 107 ---------------------------- test-suite/bench/lists_100.v | 107 ---------------------------- test-suite/kernel/inds.mv | 3 - test-suite/misc/berardi_test.v | 153 ----------------------------------------- 4 files changed, 370 deletions(-) delete mode 100644 test-suite/bench/lists-100.v delete mode 100644 test-suite/bench/lists_100.v delete mode 100644 test-suite/kernel/inds.mv delete mode 100644 test-suite/misc/berardi_test.v diff --git a/test-suite/bench/lists-100.v b/test-suite/bench/lists-100.v deleted file mode 100644 index 5c64716c7..000000000 --- a/test-suite/bench/lists-100.v +++ /dev/null @@ -1,107 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* list0 -> list0. -Inductive list1 : Set := nil1 : list1 | cons1 : Set -> list1 -> list1. -Inductive list2 : Set := nil2 : list2 | cons2 : Set -> list2 -> list2. -Inductive list3 : Set := nil3 : list3 | cons3 : Set -> list3 -> list3. -Inductive list4 : Set := nil4 : list4 | cons4 : Set -> list4 -> list4. -Inductive list5 : Set := nil5 : list5 | cons5 : Set -> list5 -> list5. -Inductive list6 : Set := nil6 : list6 | cons6 : Set -> list6 -> list6. -Inductive list7 : Set := nil7 : list7 | cons7 : Set -> list7 -> list7. -Inductive list8 : Set := nil8 : list8 | cons8 : Set -> list8 -> list8. -Inductive list9 : Set := nil9 : list9 | cons9 : Set -> list9 -> list9. -Inductive list10 : Set := nil10 : list10 | cons10 : Set -> list10 -> list10. -Inductive list11 : Set := nil11 : list11 | cons11 : Set -> list11 -> list11. -Inductive list12 : Set := nil12 : list12 | cons12 : Set -> list12 -> list12. -Inductive list13 : Set := nil13 : list13 | cons13 : Set -> list13 -> list13. -Inductive list14 : Set := nil14 : list14 | cons14 : Set -> list14 -> list14. -Inductive list15 : Set := nil15 : list15 | cons15 : Set -> list15 -> list15. -Inductive list16 : Set := nil16 : list16 | cons16 : Set -> list16 -> list16. -Inductive list17 : Set := nil17 : list17 | cons17 : Set -> list17 -> list17. -Inductive list18 : Set := nil18 : list18 | cons18 : Set -> list18 -> list18. -Inductive list19 : Set := nil19 : list19 | cons19 : Set -> list19 -> list19. -Inductive list20 : Set := nil20 : list20 | cons20 : Set -> list20 -> list20. -Inductive list21 : Set := nil21 : list21 | cons21 : Set -> list21 -> list21. -Inductive list22 : Set := nil22 : list22 | cons22 : Set -> list22 -> list22. -Inductive list23 : Set := nil23 : list23 | cons23 : Set -> list23 -> list23. -Inductive list24 : Set := nil24 : list24 | cons24 : Set -> list24 -> list24. -Inductive list25 : Set := nil25 : list25 | cons25 : Set -> list25 -> list25. -Inductive list26 : Set := nil26 : list26 | cons26 : Set -> list26 -> list26. -Inductive list27 : Set := nil27 : list27 | cons27 : Set -> list27 -> list27. -Inductive list28 : Set := nil28 : list28 | cons28 : Set -> list28 -> list28. -Inductive list29 : Set := nil29 : list29 | cons29 : Set -> list29 -> list29. -Inductive list30 : Set := nil30 : list30 | cons30 : Set -> list30 -> list30. -Inductive list31 : Set := nil31 : list31 | cons31 : Set -> list31 -> list31. -Inductive list32 : Set := nil32 : list32 | cons32 : Set -> list32 -> list32. -Inductive list33 : Set := nil33 : list33 | cons33 : Set -> list33 -> list33. -Inductive list34 : Set := nil34 : list34 | cons34 : Set -> list34 -> list34. -Inductive list35 : Set := nil35 : list35 | cons35 : Set -> list35 -> list35. -Inductive list36 : Set := nil36 : list36 | cons36 : Set -> list36 -> list36. -Inductive list37 : Set := nil37 : list37 | cons37 : Set -> list37 -> list37. -Inductive list38 : Set := nil38 : list38 | cons38 : Set -> list38 -> list38. -Inductive list39 : Set := nil39 : list39 | cons39 : Set -> list39 -> list39. -Inductive list40 : Set := nil40 : list40 | cons40 : Set -> list40 -> list40. -Inductive list41 : Set := nil41 : list41 | cons41 : Set -> list41 -> list41. -Inductive list42 : Set := nil42 : list42 | cons42 : Set -> list42 -> list42. -Inductive list43 : Set := nil43 : list43 | cons43 : Set -> list43 -> list43. -Inductive list44 : Set := nil44 : list44 | cons44 : Set -> list44 -> list44. -Inductive list45 : Set := nil45 : list45 | cons45 : Set -> list45 -> list45. -Inductive list46 : Set := nil46 : list46 | cons46 : Set -> list46 -> list46. -Inductive list47 : Set := nil47 : list47 | cons47 : Set -> list47 -> list47. -Inductive list48 : Set := nil48 : list48 | cons48 : Set -> list48 -> list48. -Inductive list49 : Set := nil49 : list49 | cons49 : Set -> list49 -> list49. -Inductive list50 : Set := nil50 : list50 | cons50 : Set -> list50 -> list50. -Inductive list51 : Set := nil51 : list51 | cons51 : Set -> list51 -> list51. -Inductive list52 : Set := nil52 : list52 | cons52 : Set -> list52 -> list52. -Inductive list53 : Set := nil53 : list53 | cons53 : Set -> list53 -> list53. -Inductive list54 : Set := nil54 : list54 | cons54 : Set -> list54 -> list54. -Inductive list55 : Set := nil55 : list55 | cons55 : Set -> list55 -> list55. -Inductive list56 : Set := nil56 : list56 | cons56 : Set -> list56 -> list56. -Inductive list57 : Set := nil57 : list57 | cons57 : Set -> list57 -> list57. -Inductive list58 : Set := nil58 : list58 | cons58 : Set -> list58 -> list58. -Inductive list59 : Set := nil59 : list59 | cons59 : Set -> list59 -> list59. -Inductive list60 : Set := nil60 : list60 | cons60 : Set -> list60 -> list60. -Inductive list61 : Set := nil61 : list61 | cons61 : Set -> list61 -> list61. -Inductive list62 : Set := nil62 : list62 | cons62 : Set -> list62 -> list62. -Inductive list63 : Set := nil63 : list63 | cons63 : Set -> list63 -> list63. -Inductive list64 : Set := nil64 : list64 | cons64 : Set -> list64 -> list64. -Inductive list65 : Set := nil65 : list65 | cons65 : Set -> list65 -> list65. -Inductive list66 : Set := nil66 : list66 | cons66 : Set -> list66 -> list66. -Inductive list67 : Set := nil67 : list67 | cons67 : Set -> list67 -> list67. -Inductive list68 : Set := nil68 : list68 | cons68 : Set -> list68 -> list68. -Inductive list69 : Set := nil69 : list69 | cons69 : Set -> list69 -> list69. -Inductive list70 : Set := nil70 : list70 | cons70 : Set -> list70 -> list70. -Inductive list71 : Set := nil71 : list71 | cons71 : Set -> list71 -> list71. -Inductive list72 : Set := nil72 : list72 | cons72 : Set -> list72 -> list72. -Inductive list73 : Set := nil73 : list73 | cons73 : Set -> list73 -> list73. -Inductive list74 : Set := nil74 : list74 | cons74 : Set -> list74 -> list74. -Inductive list75 : Set := nil75 : list75 | cons75 : Set -> list75 -> list75. -Inductive list76 : Set := nil76 : list76 | cons76 : Set -> list76 -> list76. -Inductive list77 : Set := nil77 : list77 | cons77 : Set -> list77 -> list77. -Inductive list78 : Set := nil78 : list78 | cons78 : Set -> list78 -> list78. -Inductive list79 : Set := nil79 : list79 | cons79 : Set -> list79 -> list79. -Inductive list80 : Set := nil80 : list80 | cons80 : Set -> list80 -> list80. -Inductive list81 : Set := nil81 : list81 | cons81 : Set -> list81 -> list81. -Inductive list82 : Set := nil82 : list82 | cons82 : Set -> list82 -> list82. -Inductive list83 : Set := nil83 : list83 | cons83 : Set -> list83 -> list83. -Inductive list84 : Set := nil84 : list84 | cons84 : Set -> list84 -> list84. -Inductive list85 : Set := nil85 : list85 | cons85 : Set -> list85 -> list85. -Inductive list86 : Set := nil86 : list86 | cons86 : Set -> list86 -> list86. -Inductive list87 : Set := nil87 : list87 | cons87 : Set -> list87 -> list87. -Inductive list88 : Set := nil88 : list88 | cons88 : Set -> list88 -> list88. -Inductive list89 : Set := nil89 : list89 | cons89 : Set -> list89 -> list89. -Inductive list90 : Set := nil90 : list90 | cons90 : Set -> list90 -> list90. -Inductive list91 : Set := nil91 : list91 | cons91 : Set -> list91 -> list91. -Inductive list92 : Set := nil92 : list92 | cons92 : Set -> list92 -> list92. -Inductive list93 : Set := nil93 : list93 | cons93 : Set -> list93 -> list93. -Inductive list94 : Set := nil94 : list94 | cons94 : Set -> list94 -> list94. -Inductive list95 : Set := nil95 : list95 | cons95 : Set -> list95 -> list95. -Inductive list96 : Set := nil96 : list96 | cons96 : Set -> list96 -> list96. -Inductive list97 : Set := nil97 : list97 | cons97 : Set -> list97 -> list97. -Inductive list98 : Set := nil98 : list98 | cons98 : Set -> list98 -> list98. -Inductive list99 : Set := nil99 : list99 | cons99 : Set -> list99 -> list99. diff --git a/test-suite/bench/lists_100.v b/test-suite/bench/lists_100.v deleted file mode 100644 index 5c64716c7..000000000 --- a/test-suite/bench/lists_100.v +++ /dev/null @@ -1,107 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* list0 -> list0. -Inductive list1 : Set := nil1 : list1 | cons1 : Set -> list1 -> list1. -Inductive list2 : Set := nil2 : list2 | cons2 : Set -> list2 -> list2. -Inductive list3 : Set := nil3 : list3 | cons3 : Set -> list3 -> list3. -Inductive list4 : Set := nil4 : list4 | cons4 : Set -> list4 -> list4. -Inductive list5 : Set := nil5 : list5 | cons5 : Set -> list5 -> list5. -Inductive list6 : Set := nil6 : list6 | cons6 : Set -> list6 -> list6. -Inductive list7 : Set := nil7 : list7 | cons7 : Set -> list7 -> list7. -Inductive list8 : Set := nil8 : list8 | cons8 : Set -> list8 -> list8. -Inductive list9 : Set := nil9 : list9 | cons9 : Set -> list9 -> list9. -Inductive list10 : Set := nil10 : list10 | cons10 : Set -> list10 -> list10. -Inductive list11 : Set := nil11 : list11 | cons11 : Set -> list11 -> list11. -Inductive list12 : Set := nil12 : list12 | cons12 : Set -> list12 -> list12. -Inductive list13 : Set := nil13 : list13 | cons13 : Set -> list13 -> list13. -Inductive list14 : Set := nil14 : list14 | cons14 : Set -> list14 -> list14. -Inductive list15 : Set := nil15 : list15 | cons15 : Set -> list15 -> list15. -Inductive list16 : Set := nil16 : list16 | cons16 : Set -> list16 -> list16. -Inductive list17 : Set := nil17 : list17 | cons17 : Set -> list17 -> list17. -Inductive list18 : Set := nil18 : list18 | cons18 : Set -> list18 -> list18. -Inductive list19 : Set := nil19 : list19 | cons19 : Set -> list19 -> list19. -Inductive list20 : Set := nil20 : list20 | cons20 : Set -> list20 -> list20. -Inductive list21 : Set := nil21 : list21 | cons21 : Set -> list21 -> list21. -Inductive list22 : Set := nil22 : list22 | cons22 : Set -> list22 -> list22. -Inductive list23 : Set := nil23 : list23 | cons23 : Set -> list23 -> list23. -Inductive list24 : Set := nil24 : list24 | cons24 : Set -> list24 -> list24. -Inductive list25 : Set := nil25 : list25 | cons25 : Set -> list25 -> list25. -Inductive list26 : Set := nil26 : list26 | cons26 : Set -> list26 -> list26. -Inductive list27 : Set := nil27 : list27 | cons27 : Set -> list27 -> list27. -Inductive list28 : Set := nil28 : list28 | cons28 : Set -> list28 -> list28. -Inductive list29 : Set := nil29 : list29 | cons29 : Set -> list29 -> list29. -Inductive list30 : Set := nil30 : list30 | cons30 : Set -> list30 -> list30. -Inductive list31 : Set := nil31 : list31 | cons31 : Set -> list31 -> list31. -Inductive list32 : Set := nil32 : list32 | cons32 : Set -> list32 -> list32. -Inductive list33 : Set := nil33 : list33 | cons33 : Set -> list33 -> list33. -Inductive list34 : Set := nil34 : list34 | cons34 : Set -> list34 -> list34. -Inductive list35 : Set := nil35 : list35 | cons35 : Set -> list35 -> list35. -Inductive list36 : Set := nil36 : list36 | cons36 : Set -> list36 -> list36. -Inductive list37 : Set := nil37 : list37 | cons37 : Set -> list37 -> list37. -Inductive list38 : Set := nil38 : list38 | cons38 : Set -> list38 -> list38. -Inductive list39 : Set := nil39 : list39 | cons39 : Set -> list39 -> list39. -Inductive list40 : Set := nil40 : list40 | cons40 : Set -> list40 -> list40. -Inductive list41 : Set := nil41 : list41 | cons41 : Set -> list41 -> list41. -Inductive list42 : Set := nil42 : list42 | cons42 : Set -> list42 -> list42. -Inductive list43 : Set := nil43 : list43 | cons43 : Set -> list43 -> list43. -Inductive list44 : Set := nil44 : list44 | cons44 : Set -> list44 -> list44. -Inductive list45 : Set := nil45 : list45 | cons45 : Set -> list45 -> list45. -Inductive list46 : Set := nil46 : list46 | cons46 : Set -> list46 -> list46. -Inductive list47 : Set := nil47 : list47 | cons47 : Set -> list47 -> list47. -Inductive list48 : Set := nil48 : list48 | cons48 : Set -> list48 -> list48. -Inductive list49 : Set := nil49 : list49 | cons49 : Set -> list49 -> list49. -Inductive list50 : Set := nil50 : list50 | cons50 : Set -> list50 -> list50. -Inductive list51 : Set := nil51 : list51 | cons51 : Set -> list51 -> list51. -Inductive list52 : Set := nil52 : list52 | cons52 : Set -> list52 -> list52. -Inductive list53 : Set := nil53 : list53 | cons53 : Set -> list53 -> list53. -Inductive list54 : Set := nil54 : list54 | cons54 : Set -> list54 -> list54. -Inductive list55 : Set := nil55 : list55 | cons55 : Set -> list55 -> list55. -Inductive list56 : Set := nil56 : list56 | cons56 : Set -> list56 -> list56. -Inductive list57 : Set := nil57 : list57 | cons57 : Set -> list57 -> list57. -Inductive list58 : Set := nil58 : list58 | cons58 : Set -> list58 -> list58. -Inductive list59 : Set := nil59 : list59 | cons59 : Set -> list59 -> list59. -Inductive list60 : Set := nil60 : list60 | cons60 : Set -> list60 -> list60. -Inductive list61 : Set := nil61 : list61 | cons61 : Set -> list61 -> list61. -Inductive list62 : Set := nil62 : list62 | cons62 : Set -> list62 -> list62. -Inductive list63 : Set := nil63 : list63 | cons63 : Set -> list63 -> list63. -Inductive list64 : Set := nil64 : list64 | cons64 : Set -> list64 -> list64. -Inductive list65 : Set := nil65 : list65 | cons65 : Set -> list65 -> list65. -Inductive list66 : Set := nil66 : list66 | cons66 : Set -> list66 -> list66. -Inductive list67 : Set := nil67 : list67 | cons67 : Set -> list67 -> list67. -Inductive list68 : Set := nil68 : list68 | cons68 : Set -> list68 -> list68. -Inductive list69 : Set := nil69 : list69 | cons69 : Set -> list69 -> list69. -Inductive list70 : Set := nil70 : list70 | cons70 : Set -> list70 -> list70. -Inductive list71 : Set := nil71 : list71 | cons71 : Set -> list71 -> list71. -Inductive list72 : Set := nil72 : list72 | cons72 : Set -> list72 -> list72. -Inductive list73 : Set := nil73 : list73 | cons73 : Set -> list73 -> list73. -Inductive list74 : Set := nil74 : list74 | cons74 : Set -> list74 -> list74. -Inductive list75 : Set := nil75 : list75 | cons75 : Set -> list75 -> list75. -Inductive list76 : Set := nil76 : list76 | cons76 : Set -> list76 -> list76. -Inductive list77 : Set := nil77 : list77 | cons77 : Set -> list77 -> list77. -Inductive list78 : Set := nil78 : list78 | cons78 : Set -> list78 -> list78. -Inductive list79 : Set := nil79 : list79 | cons79 : Set -> list79 -> list79. -Inductive list80 : Set := nil80 : list80 | cons80 : Set -> list80 -> list80. -Inductive list81 : Set := nil81 : list81 | cons81 : Set -> list81 -> list81. -Inductive list82 : Set := nil82 : list82 | cons82 : Set -> list82 -> list82. -Inductive list83 : Set := nil83 : list83 | cons83 : Set -> list83 -> list83. -Inductive list84 : Set := nil84 : list84 | cons84 : Set -> list84 -> list84. -Inductive list85 : Set := nil85 : list85 | cons85 : Set -> list85 -> list85. -Inductive list86 : Set := nil86 : list86 | cons86 : Set -> list86 -> list86. -Inductive list87 : Set := nil87 : list87 | cons87 : Set -> list87 -> list87. -Inductive list88 : Set := nil88 : list88 | cons88 : Set -> list88 -> list88. -Inductive list89 : Set := nil89 : list89 | cons89 : Set -> list89 -> list89. -Inductive list90 : Set := nil90 : list90 | cons90 : Set -> list90 -> list90. -Inductive list91 : Set := nil91 : list91 | cons91 : Set -> list91 -> list91. -Inductive list92 : Set := nil92 : list92 | cons92 : Set -> list92 -> list92. -Inductive list93 : Set := nil93 : list93 | cons93 : Set -> list93 -> list93. -Inductive list94 : Set := nil94 : list94 | cons94 : Set -> list94 -> list94. -Inductive list95 : Set := nil95 : list95 | cons95 : Set -> list95 -> list95. -Inductive list96 : Set := nil96 : list96 | cons96 : Set -> list96 -> list96. -Inductive list97 : Set := nil97 : list97 | cons97 : Set -> list97 -> list97. -Inductive list98 : Set := nil98 : list98 | cons98 : Set -> list98 -> list98. -Inductive list99 : Set := nil99 : list99 | cons99 : Set -> list99 -> list99. diff --git a/test-suite/kernel/inds.mv b/test-suite/kernel/inds.mv deleted file mode 100644 index bd1cc49f8..000000000 --- a/test-suite/kernel/inds.mv +++ /dev/null @@ -1,3 +0,0 @@ -Inductive [] nat : Set := O : nat | S : nat->nat. -Check Construct nat 0 1. -Check Construct nat 0 2. diff --git a/test-suite/misc/berardi_test.v b/test-suite/misc/berardi_test.v deleted file mode 100644 index a64db4dab..000000000 --- a/test-suite/misc/berardi_test.v +++ /dev/null @@ -1,153 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* > *) - -Set Implicit Arguments. - -Section Berardis_paradox. - -(** Excluded middle *) -Hypothesis EM : forall P:Prop, P \/ ~ P. - -(** Conditional on any proposition. *) -Definition IFProp (P B:Prop) (e1 e2:P) := - match EM B with - | or_introl _ => e1 - | or_intror _ => e2 - end. - -(** Axiom of choice applied to disjunction. - Provable in Coq because of dependent elimination. *) -Lemma AC_IF : - forall (P B:Prop) (e1 e2:P) (Q:P -> Prop), - (B -> Q e1) -> (~ B -> Q e2) -> Q (IFProp B e1 e2). -Proof. -intros P B e1 e2 Q p1 p2. -unfold IFProp. -case (EM B); assumption. -Qed. - - -(** We assume a type with two elements. They play the role of booleans. - The main theorem under the current assumptions is that [T=F] *) -Variable Bool : Prop. -Variable T : Bool. -Variable F : Bool. - -(** The powerset operator *) -Definition pow (P:Prop) := P -> Bool. - - -(** A piece of theory about retracts *) -Section Retracts. - -Variables A B : Prop. - -Record retract : Prop := - {i : A -> B; j : B -> A; inv : forall a:A, j (i a) = a}. - -Record retract_cond : Prop := - {i2 : A -> B; j2 : B -> A; inv2 : retract -> forall a:A, j2 (i2 a) = a}. - - -(** The dependent elimination above implies the axiom of choice: *) -Lemma AC : forall r:retract_cond, retract -> forall a:A, j2 r (i2 r a) = a. -Proof. -intros r. -case r; simpl. -trivial. -Qed. - -End Retracts. - -(** This lemma is basically a commutation of implication and existential - quantification: (EX x | A -> P(x)) <=> (A -> EX x | P(x)) - which is provable in classical logic ( => is already provable in - intuitionnistic logic). *) - -Lemma L1 : forall A B:Prop, retract_cond (pow A) (pow B). -Proof. -intros A B. -destruct (EM (retract (pow A) (pow B))) as [(f0,g0,e) | hf]. - exists f0 g0; trivial. - exists (fun (x:pow A) (y:B) => F) (fun (x:pow B) (y:A) => F); intros; - destruct hf; auto. -Qed. - - -(** The paradoxical set *) -Definition U := forall P:Prop, pow P. - -(** Bijection between [U] and [(pow U)] *) -Definition f (u:U) : pow U := u U. - -Definition g (h:pow U) : U := - fun X => let lX := j2 (L1 X U) in let rU := i2 (L1 U U) in lX (rU h). - -(** We deduce that the powerset of [U] is a retract of [U]. - This lemma is stated in Berardi's article, but is not used - afterwards. *) -Lemma retract_pow_U_U : retract (pow U) U. -Proof. -exists g f. -intro a. -unfold f, g; simpl. -apply AC. -exists (fun x:pow U => x) (fun x:pow U => x). -trivial. -Qed. - -(** Encoding of Russel's paradox *) - -(** The boolean negation. *) -Definition Not_b (b:Bool) := IFProp (b = T) F T. - -(** the set of elements not belonging to itself *) -Definition R : U := g (fun u:U => Not_b (u U u)). - - -Lemma not_has_fixpoint : R R = Not_b (R R). -Proof. -unfold R at 1. -unfold g. -rewrite AC with (r := L1 U U) (a := fun u:U => Not_b (u U u)). -trivial. -exists (fun x:pow U => x) (fun x:pow U => x); trivial. -Qed. - - -Theorem classical_proof_irrelevence : T = F. -Proof. -generalize not_has_fixpoint. -unfold Not_b. -apply AC_IF. -intros is_true is_false. -elim is_true; elim is_false; trivial. - -intros not_true is_true. -elim not_true; trivial. -Qed. - -End Berardis_paradox. -- cgit v1.2.3 From 7f17e151c0a2666263d3854c064acdfea29edf53 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 8 May 2017 17:19:22 +0200 Subject: Adding tests for testing exit status and #use"include". --- test-suite/.csdp.cache | Bin 89077 -> 89077 bytes test-suite/Makefile | 30 +++++++++++++++++++++++++++++- test-suite/misc/exitstatus/illtyped.v | 1 + 3 files changed, 30 insertions(+), 1 deletion(-) create mode 100644 test-suite/misc/exitstatus/illtyped.v diff --git a/test-suite/.csdp.cache b/test-suite/.csdp.cache index ba85286dd..b99d80e95 100644 Binary files a/test-suite/.csdp.cache and b/test-suite/.csdp.cache differ diff --git a/test-suite/Makefile b/test-suite/Makefile index c10cd4ed4..779f5455a 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -33,6 +33,7 @@ LIB := $(shell cd ..; pwd) coqtop := $(BIN)coqtop -boot -q -batch -test-mode -R prerequisite TestSuite bincoqc := $(BIN)coqc -coqlib $(LIB) -R prerequisite TestSuite bincoqchk := $(BIN)coqchk -coqlib $(LIB) -R prerequisite TestSuite +coqtopbyte := $(BIN)coqtop.byte command := $(coqtop) -top Top -async-proofs-cache force -load-vernac-source coqc := $(coqtop) -compile @@ -405,7 +406,7 @@ modules/%.vo: modules/%.v # Miscellaneous tests ####################################################################### -misc: misc/deps-order.log misc/universes.log misc/deps-checksum.log +misc: misc/deps-order.log misc/universes.log misc/deps-checksum.log misc/printers.log misc/exitstatus.log # Check that both coqdep and coqtop/coqc supports -R # Check that both coqdep and coqtop/coqc takes the later -R @@ -477,6 +478,33 @@ misc/universes.log: misc/universes/all_stdlib.v misc/universes/all_stdlib.v: cd .. && $(MAKE) test-suite/$@ +misc/printers.log: + @echo "TEST printers" + $(HIDE){ \ + echo $(call log_intro,$<); \ + printf "Drop. #use\"include\";; #quit;;" | $(coqtopbyte) 2>&1 | grep Unbound; R=$$?; times; \ + if [ $$R != 0 ]; then \ + echo $(log_success); \ + echo " misc/printers...Ok"; \ + else \ + echo $(log_failure); \ + echo " misc/printers...Error!"; \ + fi; \ + } > "$@" + +misc/exitstatus.log: + @echo "TEST exitstatus" + $(HIDE){ \ + echo $(call log_intro,$<); \ + $(coqc) misc/exitstatus/illtyped.v 2>&1; R=$$?; times; \ + if [ $$R != 0 ]; then \ + echo $(log_success); \ + echo " misc/exitstatus...Ok"; \ + else \ + echo $(log_failure); \ + echo " misc/exitstatus...Error!"; \ + fi; \ + } > "$@" # IDE : some tests of backtracking for coqtop -ideslave diff --git a/test-suite/misc/exitstatus/illtyped.v b/test-suite/misc/exitstatus/illtyped.v new file mode 100644 index 000000000..df6bbb389 --- /dev/null +++ b/test-suite/misc/exitstatus/illtyped.v @@ -0,0 +1 @@ +Check S S. -- cgit v1.2.3 From 63510329fae28f7a9d18935f24e3ebf0485dabc8 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 8 May 2017 18:28:55 +0200 Subject: A more regular naming of variables in test-suite Makefile. --- test-suite/Makefile | 36 ++++++++++++++++++------------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/test-suite/Makefile b/test-suite/Makefile index 779f5455a..5d3c11844 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -31,12 +31,12 @@ BIN := ../bin/ LIB := $(shell cd ..; pwd) coqtop := $(BIN)coqtop -boot -q -batch -test-mode -R prerequisite TestSuite -bincoqc := $(BIN)coqc -coqlib $(LIB) -R prerequisite TestSuite -bincoqchk := $(BIN)coqchk -coqlib $(LIB) -R prerequisite TestSuite +coqc := $(BIN)coqc -coqlib $(LIB) -R prerequisite TestSuite +coqchk := $(BIN)coqchk -coqlib $(LIB) -R prerequisite TestSuite coqtopbyte := $(BIN)coqtop.byte -command := $(coqtop) -top Top -async-proofs-cache force -load-vernac-source -coqc := $(coqtop) -compile +coqtopload := $(coqtop) -top Top -async-proofs-cache force -load-vernac-source +coqtopcompile := $(coqtop) -compile coqdep := $(BIN)coqdep -coqlib $(LIB) SHOW := $(if $(VERBOSE),@true,@echo) @@ -53,7 +53,7 @@ get_coq_prog_args_in_parens = $(subst $(SINGLE_QUOTE),,$(if $(call get_coq_prog_ # get the command to use with this set of arguments; if there's -compile, use coqc, else use coqtop has_compile_flag = $(filter "-compile",$(call get_coq_prog_args,$(1))) has_profile_ltac_or_compile_flag = $(filter "-profile-ltac-cutoff" "-profile-ltac" "-compile",$(call get_coq_prog_args,$(1))) -get_command_based_on_flags = $(if $(call has_profile_ltac_or_compile_flag,$(1)),$(coqc),$(command)) +get_command_based_on_flags = $(if $(call has_profile_ltac_or_compile_flag,$(1)),$(coqtopcompile),$(coqtopload)) bogomips:= @@ -220,7 +220,7 @@ $(addsuffix .log,$(wildcard prerequisite/*.v)): %.v.log: %.v @echo "TEST $< $(call get_coq_prog_args_in_parens,"$<")" $(HIDE){ \ echo $(call log_intro,$<); \ - $(coqc) "$<" $(call get_coq_prog_args,"$<") 2>&1; R=$$?; times; \ + $(coqtopcompile) "$<" $(call get_coq_prog_args,"$<") 2>&1; R=$$?; times; \ if [ $$R != 0 ]; then \ echo $(log_failure); \ echo " $<...could not be prepared" ; \ @@ -250,7 +250,7 @@ $(addsuffix .log,$(wildcard stm/*.v)): %.v.log: %.v @echo "TEST $< $(call get_coq_prog_args_in_parens,"$<")" $(HIDE){ \ echo $(call log_intro,$<); \ - $(coqc) "$<" $(call get_coq_prog_args,"$<") -async-proofs on \ + $(coqtopcompile) "$<" $(call get_coq_prog_args,"$<") -async-proofs on \ $$opts 2>&1; R=$$?; times; \ if [ $$R = 0 ]; then \ echo $(log_success); \ @@ -421,8 +421,8 @@ misc/deps-order.log: $(coqdep) -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/bar.v 2>&1 \ | head -n 1 > $$tmpoutput; \ diff -u misc/deps/deps.out $$tmpoutput 2>&1; R=$$?; times; \ - $(bincoqc) -R misc/deps/lib lib misc/deps/lib/foo.v 2>&1; \ - $(bincoqc) -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/foo.v 2>&1; \ + $(coqc) -R misc/deps/lib lib misc/deps/lib/foo.v 2>&1; \ + $(coqc) -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/foo.v 2>&1; \ $(coqtop) -R misc/deps/lib lib -R misc/deps/client client -load-vernac-source misc/deps/client/bar.v 2>&1; \ S=$$?; times; \ if [ $$R = 0 -a $$S = 0 ]; then \ @@ -441,9 +441,9 @@ misc/deps-checksum.log: $(HIDE){ \ echo $(call log_intro,deps-checksum); \ rm -f misc/deps/*/*.vo; \ - $(bincoqc) -R misc/deps/A A misc/deps/A/A.v; \ - $(bincoqc) -R misc/deps/B A misc/deps/B/A.v; \ - $(bincoqc) -R misc/deps/B A misc/deps/B/B.v; \ + $(coqc) -R misc/deps/A A misc/deps/A/A.v; \ + $(coqc) -R misc/deps/B A misc/deps/B/A.v; \ + $(coqc) -R misc/deps/B A misc/deps/B/B.v; \ $(coqtop) -R misc/deps/B A -R misc/deps/A A -load-vernac-source misc/deps/checksum.v 2>&1; \ if [ $$? = 0 ]; then \ echo $(log_success); \ @@ -461,8 +461,8 @@ universes: misc/universes.log misc/universes.log: misc/universes/all_stdlib.v @echo "TEST misc/universes" $(HIDE){ \ - $(bincoqc) -R misc/universes Universes misc/universes/all_stdlib 2>&1; \ - $(bincoqc) -R misc/universes Universes misc/universes/universes 2>&1; \ + $(coqc) -R misc/universes Universes misc/universes/all_stdlib 2>&1; \ + $(coqc) -R misc/universes Universes misc/universes/universes 2>&1; \ mv universes.txt misc/universes; \ N=`awk '{print $$3}' misc/universes/universes.txt | sort -u | wc -l`; \ times; \ @@ -529,9 +529,9 @@ vio: $(patsubst %.v,%.vio.log,$(wildcard vio/*.v)) %.vio.log:%.v @echo "TEST $<" $(HIDE){ \ - $(bincoqc) -quick -R vio vio $* 2>&1 && \ + $(coqc) -quick -R vio vio $* 2>&1 && \ $(coqtop) -R vio vio -vio2vo $*.vio 2>&1 && \ - $(bincoqchk) -R vio vio -norec $(subst /,.,$*) 2>&1; \ + $(coqchk) -R vio vio -norec $(subst /,.,$*) 2>&1; \ if [ $$? = 0 ]; then \ echo $(log_success); \ echo " $<...Ok"; \ @@ -546,8 +546,8 @@ coqchk: $(patsubst %.v,%.chk.log,$(wildcard coqchk/*.v)) %.chk.log:%.v @echo "TEST $<" $(HIDE){ \ - $(bincoqc) -R coqchk coqchk $* 2>&1 && \ - $(bincoqchk) -R coqchk coqchk -norec $(subst /,.,$*) 2>&1; \ + $(coqc) -R coqchk coqchk $* 2>&1 && \ + $(coqchk) -R coqchk coqchk -norec $(subst /,.,$*) 2>&1; \ if [ $$? = 0 ]; then \ echo $(log_success); \ echo " $<...Ok"; \ -- cgit v1.2.3 From 98a927aefb6df05a957d94cf2fd22d88e9e6c6b6 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Tue, 9 May 2017 21:48:43 +0200 Subject: Moving code for miscellaneous tests to specific files. The rule is that any file misc/*.sh will now be automatically executed from the test-file directory with appropriate environment variables set. The test can use any auxiliary material which is put in a directory of the same name as the test. This avoids making the main Makefile more or more complex. We loose some customization though (no more custom messages such as printing of the number of universes in the test about the number of necessary universes). We could preserve such customization if needed but I did not consider it was worth the effort. Warning: specific targets universes, deps-order, deps-checksum are removed by consistency. Do instead "make misc/universes.log", etc., or even "make misc" for all. --- test-suite/Makefile | 102 +++++---------------------------------- test-suite/misc/deps-checksum.sh | 5 ++ test-suite/misc/deps-order.sh | 20 ++++++++ test-suite/misc/exitstatus.sh | 7 +++ test-suite/misc/printers.sh | 3 ++ test-suite/misc/universes.sh | 8 +++ 6 files changed, 54 insertions(+), 91 deletions(-) create mode 100755 test-suite/misc/deps-checksum.sh create mode 100755 test-suite/misc/deps-order.sh create mode 100755 test-suite/misc/exitstatus.sh create mode 100755 test-suite/misc/printers.sh create mode 100755 test-suite/misc/universes.sh diff --git a/test-suite/Makefile b/test-suite/Makefile index 5d3c11844..39ef05269 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -406,103 +406,23 @@ modules/%.vo: modules/%.v # Miscellaneous tests ####################################################################### -misc: misc/deps-order.log misc/universes.log misc/deps-checksum.log misc/printers.log misc/exitstatus.log - -# Check that both coqdep and coqtop/coqc supports -R -# Check that both coqdep and coqtop/coqc takes the later -R -# See bugs 2242, 2337, 2339 -deps-order: misc/deps-order.log -misc/deps-order.log: - @echo "TEST misc/deps-order" - $(HIDE){ \ - echo $(call log_intro,deps-order); \ - rm -f misc/deps/*/*.vo; \ - tmpoutput=`mktemp /tmp/coqcheck.XXXXXX`; \ - $(coqdep) -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/bar.v 2>&1 \ - | head -n 1 > $$tmpoutput; \ - diff -u misc/deps/deps.out $$tmpoutput 2>&1; R=$$?; times; \ - $(coqc) -R misc/deps/lib lib misc/deps/lib/foo.v 2>&1; \ - $(coqc) -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/foo.v 2>&1; \ - $(coqtop) -R misc/deps/lib lib -R misc/deps/client client -load-vernac-source misc/deps/client/bar.v 2>&1; \ - S=$$?; times; \ - if [ $$R = 0 -a $$S = 0 ]; then \ - echo $(log_success); \ - echo " misc/deps-order...Ok"; \ - else \ - echo $(log_failure); \ - echo " misc/deps-order...Error! (unexpected order)"; \ - fi; \ - rm $$tmpoutput; \ - } > "$@" - -deps-checksum: misc/deps-checksum.log -misc/deps-checksum.log: - @echo "TEST misc/deps-checksum" - $(HIDE){ \ - echo $(call log_intro,deps-checksum); \ - rm -f misc/deps/*/*.vo; \ - $(coqc) -R misc/deps/A A misc/deps/A/A.v; \ - $(coqc) -R misc/deps/B A misc/deps/B/A.v; \ - $(coqc) -R misc/deps/B A misc/deps/B/B.v; \ - $(coqtop) -R misc/deps/B A -R misc/deps/A A -load-vernac-source misc/deps/checksum.v 2>&1; \ - if [ $$? = 0 ]; then \ - echo $(log_success); \ - echo " misc/deps-checksum...Ok"; \ - else \ - echo $(log_failure); \ - echo " misc/deps-checksum...Error!"; \ - fi; \ - } > "$@" - +misc: $(patsubst %.sh,%.log,$(wildcard misc/*.sh)) -# Sort universes for the whole standard library -EXPECTED_UNIVERSES := 4 # Prop is not counted -universes: misc/universes.log -misc/universes.log: misc/universes/all_stdlib.v - @echo "TEST misc/universes" - $(HIDE){ \ - $(coqc) -R misc/universes Universes misc/universes/all_stdlib 2>&1; \ - $(coqc) -R misc/universes Universes misc/universes/universes 2>&1; \ - mv universes.txt misc/universes; \ - N=`awk '{print $$3}' misc/universes/universes.txt | sort -u | wc -l`; \ - times; \ - if [ "$$N" -eq $(EXPECTED_UNIVERSES) ]; then \ - echo $(log_success); \ - echo " misc/universes...Ok ($(EXPECTED_UNIVERSES) universes)"; \ - else \ - echo $(log_failure); \ - echo " misc/universes...Error! ($$N/$(EXPECTED_UNIVERSES) universes)"; \ - fi; \ - } > "$@" - -misc/universes/all_stdlib.v: - cd .. && $(MAKE) test-suite/$@ - -misc/printers.log: - @echo "TEST printers" - $(HIDE){ \ - echo $(call log_intro,$<); \ - printf "Drop. #use\"include\";; #quit;;" | $(coqtopbyte) 2>&1 | grep Unbound; R=$$?; times; \ - if [ $$R != 0 ]; then \ - echo $(log_success); \ - echo " misc/printers...Ok"; \ - else \ - echo $(log_failure); \ - echo " misc/printers...Error!"; \ - fi; \ - } > "$@" - -misc/exitstatus.log: - @echo "TEST exitstatus" +$(patsubst %.sh,%.log,$(wildcard misc/*.sh)): %.log: %.sh $(PREREQUISITELOG) + @echo "TEST $<" $(HIDE){ \ echo $(call log_intro,$<); \ - $(coqc) misc/exitstatus/illtyped.v 2>&1; R=$$?; times; \ - if [ $$R != 0 ]; then \ + export coqc="$(coqc)"; \ + export coqtop="$(coqtop)"; \ + export coqdep="$(coqdep)"; \ + export coqtopbyte="$(coqtopbyte)"; \ + "$<" 2>&1; R=$$?; times; \ + if [ $$R = 0 ]; then \ echo $(log_success); \ - echo " misc/exitstatus...Ok"; \ + echo " $<...Ok"; \ else \ echo $(log_failure); \ - echo " misc/exitstatus...Error!"; \ + echo " $<...Error!"; \ fi; \ } > "$@" diff --git a/test-suite/misc/deps-checksum.sh b/test-suite/misc/deps-checksum.sh new file mode 100755 index 000000000..1e2afb754 --- /dev/null +++ b/test-suite/misc/deps-checksum.sh @@ -0,0 +1,5 @@ +rm -f misc/deps/*/*.vo +$coqc -R misc/deps/A A misc/deps/A/A.v +$coqc -R misc/deps/B A misc/deps/B/A.v +$coqc -R misc/deps/B A misc/deps/B/B.v +$coqtop -R misc/deps/B A -R misc/deps/A A -load-vernac-source misc/deps/checksum.v diff --git a/test-suite/misc/deps-order.sh b/test-suite/misc/deps-order.sh new file mode 100755 index 000000000..375b706f0 --- /dev/null +++ b/test-suite/misc/deps-order.sh @@ -0,0 +1,20 @@ +# Check that both coqdep and coqtop/coqc supports -R +# Check that both coqdep and coqtop/coqc takes the later -R +# See bugs 2242, 2337, 2339 +rm -f misc/deps/*/*.vo +tmpoutput=`mktemp /tmp/coqcheck.XXXXXX` +$coqdep -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/bar.v 2>&1 | head -n 1 > $tmpoutput +diff -u misc/deps/deps.out $tmpoutput 2>&1 +R=$? +times +$coqc -R misc/deps/lib lib misc/deps/lib/foo.v 2>&1 +$coqc -R misc/deps/lib lib -R misc/deps/client client misc/deps/client/foo.v 2>&1 +$coqtop -R misc/deps/lib lib -R misc/deps/client client -load-vernac-source misc/deps/client/bar.v 2>&1 +S=$? +if [ $R = 0 -a $S = 0 ]; then + printf "coqdep and coqtop agree\n" + exit 0 +else + printf "coqdep and coqtop disagree\n" + exit 1 +fi diff --git a/test-suite/misc/exitstatus.sh b/test-suite/misc/exitstatus.sh new file mode 100755 index 000000000..cea1de862 --- /dev/null +++ b/test-suite/misc/exitstatus.sh @@ -0,0 +1,7 @@ +$coqtop -load-vernac-source misc/exitstatus/illtyped.v +N=$? +$coqc misc/exitstatus/illtyped.v +P=$? +printf "On ill-typed input, coqtop returned $N.\n" +printf "On ill-typed input, coqc returned $P.\n" +if [ $N = 1 -a $P = 1 ]; then exit 0; else exit 1; fi diff --git a/test-suite/misc/printers.sh b/test-suite/misc/printers.sh new file mode 100755 index 000000000..c822d0eb3 --- /dev/null +++ b/test-suite/misc/printers.sh @@ -0,0 +1,3 @@ +printf "Drop. #use\"include\";; #quit;;\n" | $coqtopbyte 2>&1 | grep Unbound +if [ $? = 0 ]; then exit 1; else exit 0; fi + diff --git a/test-suite/misc/universes.sh b/test-suite/misc/universes.sh new file mode 100755 index 000000000..d87a86035 --- /dev/null +++ b/test-suite/misc/universes.sh @@ -0,0 +1,8 @@ +# Sort universes for the whole standard library +EXPECTED_UNIVERSES=4 # Prop is not counted +$coqc -R misc/universes Universes misc/universes/all_stdlib 2>&1 +$coqc -R misc/universes Universes misc/universes/universes 2>&1 +mv universes.txt misc/universes +N=`awk '{print $3}' misc/universes/universes.txt | sort -u | wc -l` +printf "Found %s/%s universes\n" $N $EXPECTED_UNIVERSES +if [ "$N" -eq $EXPECTED_UNIVERSES ]; then exit 0; else exit 1; fi -- cgit v1.2.3 From f1707b32c96931eb1c1d8153b8570234321153c2 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 10 May 2017 17:23:24 -0400 Subject: Switch bedrock to mit-plv base --- dev/ci/ci-basic-overlay.sh | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh index 5da13c100..1dfade261 100644 --- a/dev/ci/ci-basic-overlay.sh +++ b/dev/ci/ci-basic-overlay.sh @@ -98,13 +98,13 @@ # bedrock_src ######################################################################## : ${bedrock_src_CI_BRANCH:=master} -: ${bedrock_src_CI_GITURL:=https://github.com/JasonGross/bedrock.git} +: ${bedrock_src_CI_GITURL:=https://github.com/mit-plv/bedrock.git} ######################################################################## # bedrock_facade ######################################################################## : ${bedrock_facade_CI_BRANCH:=master} -: ${bedrock_facade_CI_GITURL:=https://github.com/JasonGross/bedrock.git} +: ${bedrock_facade_CI_GITURL:=https://github.com/mit-plv/bedrock.git} ######################################################################## # formal-topology -- cgit v1.2.3 From e0577588056110ea13a904aa1f01c86dbc931f02 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Thu, 11 May 2017 13:14:27 +0200 Subject: Remove an unused open introduced by the previous commit. --- tactics/tactics.ml | 1 - 1 file changed, 1 deletion(-) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 3842b432d..556df6e55 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5004,7 +5004,6 @@ let name_op_to_name name_op object_kind suffix = add_suffix name suffix, gk let tclABSTRACT ?(opaque=true) name_op tac = - let open Proof_global in let s, gk = if opaque then name_op_to_name name_op (Proof Theorem) "_subproof" else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in -- cgit v1.2.3 From fa530442ba05e9b60efff9c726616ae00d6d09e7 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 4 May 2017 18:30:38 +0200 Subject: Obligations shrinking: shrink abstraction too --- vernac/obligations.ml | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/vernac/obligations.ml b/vernac/obligations.ml index e0520216b..62c594389 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -631,6 +631,16 @@ let unfold_entry cst = Hints.HintsUnfoldEntry [EvalConstRef cst] let add_hint local prg cst = Hints.add_hints local [Id.to_string prg.prg_name] (unfold_entry cst) +let it_mkLambda_or_LetIn_or_clean t ctx = + let open Context.Rel.Declaration in + let fold t decl = + if is_local_assum decl then mkLambda_or_LetIn decl t + else + if noccurn 1 t then subst1 mkProp t + else mkLambda_or_LetIn decl t + in + Context.Rel.fold_inside fold ctx ~init:t + let declare_obligation prg obl body ty uctx = let body = prg.prg_reduce body in let ty = Option.map prg.prg_reduce ty in @@ -664,7 +674,7 @@ let declare_obligation prg obl body ty uctx = if poly then Some (DefinedObl constant) else - Some (TermObl (it_mkLambda_or_LetIn (mkApp (mkConst constant, args)) ctx)) } + Some (TermObl (it_mkLambda_or_LetIn_or_clean (mkApp (mkConst constant, args)) ctx)) } let init_prog_info ?(opaque = false) sign n pl b t ctx deps fixkind notations obls impls kind reduce hook = -- cgit v1.2.3 From 06f3ce00971283d2718e272ec9f123430d75ffa6 Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Thu, 11 May 2017 15:16:21 +0200 Subject: Documenting Printing Compact Contexts + CHANGES --- CHANGES | 6 ++++++ doc/refman/RefMan-oth.tex | 16 ++++++++++++++++ 2 files changed, 22 insertions(+) diff --git a/CHANGES b/CHANGES index 60b88ea8d..9318a0df1 100644 --- a/CHANGES +++ b/CHANGES @@ -19,6 +19,12 @@ Tactics be fairly uncommon. - "auto with real" can now discharge comparisons of literals +Vernacular Commands + +- Goals context can be printed in a more compact way when "Set + Printing Compact Contexts" is activated. + + Standard Library - New file PropExtensionality.v to explicitly work in the axiomatic diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex index 56ce753cd..0f6978b25 100644 --- a/doc/refman/RefMan-oth.tex +++ b/doc/refman/RefMan-oth.tex @@ -960,6 +960,22 @@ time of writing this documentation, the default value is 50). \subsection[\tt Test Printing Depth.]{\tt Test Printing Depth.\optindex{Printing Depth}} This command displays the current nesting depth used for display. +\subsection[\tt Unset Printing Compact Contexts.]{\tt Unset Printing Compact Contexts.\optindex{Printing Compact Contexts}} +This command resets the displaying of goals contexts to non compact +mode (default at the time of writing this documentation). Non compact +means that consecutive variables of different types are printed on +different lines. + +\subsection[\tt Set Printing Compact Contexts.]{\tt Set Printing Compact Contexts.\optindex{Printing Compact Contexts}} +This command sets the displaying of goals contexts to compact mode. +The printer tries to reduce the vertical size of goals contexts by +putting several variables (even if of different types) on the same +line provided it does not exceed the printing width (See {\tt Set + Printing Width} above). + +\subsection[\tt Test Printing Compact Contexts.]{\tt Test Printing Compact Contexts.\optindex{Printing Compact Contexts}} +This command displays the current state of compaction of goal d'isolat. + \subsection[\tt Set Printing Dependent Evars Line.]{\tt Set Printing Dependent Evars Line.\optindex{Printing Dependent Evars Line}} This command enables the printing of the ``{\tt (dependent evars: \ldots)}'' line when {\tt -emacs} is passed. -- cgit v1.2.3 From 6c9fb0b16fa5674a3135a49adff201d6e4415cd1 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Wed, 10 May 2017 20:48:37 -0400 Subject: Add documentation for Set Ltac Batch Debug --- doc/refman/RefMan-ltac.tex | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex index c2f52e23b..0346c4a55 100644 --- a/doc/refman/RefMan-ltac.tex +++ b/doc/refman/RefMan-ltac.tex @@ -1242,7 +1242,7 @@ This will automatically print the same trace as {\tt Info \num} at each tactic c The current value for the {\tt Info Level} option can be checked using the {\tt Test Info Level} command. -\subsection[Interactive debugger]{Interactive debugger\optindex{Ltac Debug}} +\subsection[Interactive debugger]{Interactive debugger\optindex{Ltac Debug}\optindex{Ltac Batch Debug}} The {\ltac} interpreter comes with a step-by-step debugger. The debugger can be activated using the command @@ -1273,6 +1273,17 @@ r $n$: & advance $n$ steps further\\ r {\qstring}: & advance up to the next call to ``{\tt idtac} {\qstring}''\\ \end{tabular} +A non-interactive mode for the debugger is available via the command + +\begin{quote} +{\tt Set Ltac Batch Debug.} +\end{quote} + +This option has the effect of presenting a newline at every prompt, +when the debugger is on. The debug log thus created, which does not +require user input to generate when this option is set, can then be +run through external tools such as \texttt{diff}. + \subsection[Profiling {\ltac} tactics]{Profiling {\ltac} tactics\optindex{Ltac Profiling}\comindex{Show Ltac Profile}\comindex{Reset Ltac Profile}} It is possible to measure the time spent in invocations of primitive tactics as well as tactics defined in {\ltac} and their inner invocations. The primary use is the development of complex tactics, which can sometimes be so slow as to impede interactive usage. The reasons for the performence degradation can be intricate, like a slowly performing {\ltac} match or a sub-tactic whose performance only degrades in certain situations. The profiler generates a call tree and indicates the time spent in a tactic depending its calling context. Thus it allows to locate the part of a tactic definition that contains the performance bug. -- cgit v1.2.3 From ba88e99d37817cd3a591b14bbceb61e28f3e3382 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Thu, 27 Apr 2017 17:26:17 +0200 Subject: Aligning on standard layout of CHANGES. --- CHANGES | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/CHANGES b/CHANGES index 8cb5573b2..69271f932 100644 --- a/CHANGES +++ b/CHANGES @@ -17,13 +17,12 @@ Tactics Most notably, the new implementation recognizes Miller patterns that were missed before because of a missing normalization step. Hopefully this should be fairly uncommon. -- "auto with real" can now discharge comparisons of literals +- Tactic "auto with real" can now discharge comparisons of literals. - The types of variables in patterns of "match" are now beta-iota-reduced after type-checking. This has an impact on the type of the variables that the tactic "refine" introduces in the context, producing types a priori closer to the expectations. - Standard Library - New file PropExtensionality.v to explicitly work in the axiomatic -- cgit v1.2.3 From e3a3b4bb17c40b6af2ef8501c405f0600cc6d65b Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 22 Apr 2017 21:09:05 +0200 Subject: Typo in a comment of plugin Quote. --- plugins/quote/quote.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index fc9d70ae7..c649cfb2c 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -8,7 +8,7 @@ (* The `Quote' tactic *) -(* The basic idea is to automatize the inversion of interpetation functions +(* The basic idea is to automatize the inversion of interpretation functions in 2-level approach Examples are given in \texttt{theories/DEMOS/DemoQuote.v} -- cgit v1.2.3 From e3550a0acc39e235e01a727267b12a7c06f23b2c Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 26 Aug 2016 14:46:33 +0200 Subject: Uniformity of style for selecting plural or not; spacing for comma. --- tactics/class_tactics.ml | 2 +- vernac/obligations.ml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 2d6dffdd2..05eb0a976 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1202,7 +1202,7 @@ module Search = struct Feedback.msg_debug (pr_depth info.search_depth ++ str": no match for " ++ Printer.pr_econstr_env (Goal.env gl) s concl ++ - spc () ++ str ", " ++ int (List.length poss) ++ + str ", " ++ int (List.length poss) ++ str" possibilities"); match e with | (ReachedLimitEx,ie) -> Proofview.tclZERO ~info:ie ReachedLimitEx diff --git a/vernac/obligations.ml b/vernac/obligations.ml index e0520216b..5233fab15 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -1088,7 +1088,7 @@ let add_definition n ?term t ctx ?pl ?(implicits=[]) ?(kind=Global,false,Definit Defined cst) else ( let len = Array.length obls in - let _ = Flags.if_verbose Feedback.msg_info (info ++ str ", generating " ++ int len ++ str " obligation(s)") in + let _ = Flags.if_verbose Feedback.msg_info (info ++ str ", generating " ++ int len ++ str (String.plural len " obligation")) in progmap_add n (CEphemeron.create prg); let res = auto_solve_obligations (Some n) tactic in match res with -- cgit v1.2.3 From 4242ccd602e1295a2bc3b95e9569c2355b3cb02f Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 13 May 2017 22:06:19 +0200 Subject: [travis] Move VST to required suite. --- .travis.yml | 1 - 1 file changed, 1 deletion(-) diff --git a/.travis.yml b/.travis.yml index adaae5487..7b6d41b38 100644 --- a/.travis.yml +++ b/.travis.yml @@ -53,7 +53,6 @@ matrix: allow_failures: - env: TEST_TARGET="ci-geocoq" - - env: TEST_TARGET="ci-vst" # Full Coq test-suite with two compilers # [TODO: use yaml refs and avoid duplication for packages list] -- cgit v1.2.3 From 9bcd13252bfb52dfcf8f2c31049bafba3d482f36 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 13 May 2017 22:11:10 +0200 Subject: [travis] Update OCaml to 4.04.1 --- .travis.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.travis.yml b/.travis.yml index 7b6d41b38..560455187 100644 --- a/.travis.yml +++ b/.travis.yml @@ -82,7 +82,7 @@ matrix: - imagemagick - env: - TEST_TARGET="test-suite" - - COMPILER="4.04.0" + - COMPILER="4.04.1" - CAMLP5_VER="6.17" - EXTRA_CONF="-coqide opt -with-doc yes" - EXTRA_OPAM="lablgtk-extras hevea" -- cgit v1.2.3 From 684dd06c538ea6024e5dd01bfc6eb416b08ddc14 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 14 May 2017 00:26:16 +0200 Subject: Removing a variable warned unused. --- toplevel/metasyntax.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml index 349c05a38..0787b058a 100644 --- a/toplevel/metasyntax.ml +++ b/toplevel/metasyntax.ml @@ -520,7 +520,7 @@ let warn_skip_spaces_curly = (fun () ->strbrk "Skipping spaces inside curly brackets") let rec drop_spacing = function - | UnpCut _ as u :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt + | UnpCut _ :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt | UnpTerminal s' :: fmt when String.equal s' (String.make (String.length s') ' ') -> warn_skip_spaces_curly (); drop_spacing fmt | fmt -> fmt -- cgit v1.2.3 From 3908fb1c6d68678daa65b4a2fa944424575acf87 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 14 May 2017 12:29:33 +0200 Subject: Removing a line warned unused. --- pretyping/constr_matching.ml | 1 - 1 file changed, 1 deletion(-) diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index afdf601c2..daac33f50 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -84,7 +84,6 @@ let rec build_lambda vars ctx m = match vars with | n :: vars -> (* change [ x1 ... xn y z1 ... zm |- t ] into [ x1 ... xn z1 ... zm |- lam y. t ] *) - let len = List.length ctx in let pre, suf = List.chop (pred n) ctx in let (na, t, suf) = match suf with | [] -> assert false -- cgit v1.2.3 From a05cdcb00edbf0e35190f2d724c4a8c46d6cf9a3 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 15 May 2017 03:00:04 +0200 Subject: Typo in comments of constrintern. --- interp/constrintern.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 3f99a3c7c..b57a04679 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -900,7 +900,7 @@ let interp_reference vars r = (**********************************************************************) (** {5 Cases } *) -(** {6 Elemtary bricks } *) +(** {6 Elementary bricks } *) let apply_scope_env env = function | [] -> {env with tmp_scope = None}, [] | sc::scl -> {env with tmp_scope = sc}, scl -- cgit v1.2.3 From b643916aed4093eb7f21116aaef726cab561bc27 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 15 May 2017 01:10:54 +0200 Subject: [interp] [ast] Make raw_cases_pattern_expr private. The type `raw_cases_pattern_expr` is used only in the interpretation of notation patterns. Indeed, this should be a private type thus we make it local to `Constrintern`; it makes no sense to expose it in the public AST. The patch is routine, except for the case used to interpret primitives in patterns. We now return a `glob_constr` representing the raw pattern, instead of the private raw pattern type. This could be further refactored but have opted to be conservative here. This patch is a refinement of b2953849b999d1c3b42c0f494b234f2a93ac7754 , see the commentaries there for more information about `raw_cases_pattern_expr`. --- interp/constrexpr_ops.ml | 6 ------ interp/constrexpr_ops.mli | 1 - interp/constrintern.ml | 36 +++++++++++++++++++++++++++++++----- interp/notation.ml | 25 +++++++++++++------------ interp/notation.mli | 4 +++- intf/constrexpr.mli | 8 -------- 6 files changed, 47 insertions(+), 33 deletions(-) diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index a592b4cff..542f9feaf 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -263,12 +263,6 @@ let cases_pattern_expr_loc = function | CPatDelimiters (loc,_,_) -> loc | CPatCast(loc,_,_) -> loc -let raw_cases_pattern_expr_loc = function - | RCPatAlias (loc,_,_) -> loc - | RCPatCstr (loc,_,_,_) -> loc - | RCPatAtom (loc,_) -> loc - | RCPatOr (loc,_) -> loc - let local_binder_loc = function | CLocalAssum ((loc,_)::_,_,t) | CLocalDef ((loc,_),t,None) -> Loc.merge loc (constr_loc t) diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli index f6d97e107..b547288e3 100644 --- a/interp/constrexpr_ops.mli +++ b/interp/constrexpr_ops.mli @@ -36,7 +36,6 @@ val binder_kind_eq : binder_kind -> binder_kind -> bool val constr_loc : constr_expr -> Loc.t val cases_pattern_expr_loc : cases_pattern_expr -> Loc.t -val raw_cases_pattern_expr_loc : raw_cases_pattern_expr -> Loc.t val local_binders_loc : local_binder_expr list -> Loc.t (** {6 Constructors}*) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 3f99a3c7c..b18341800 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -900,6 +900,21 @@ let interp_reference vars r = (**********************************************************************) (** {5 Cases } *) +(** Private internalization patterns *) +type raw_cases_pattern_expr = + | RCPatAlias of Loc.t * raw_cases_pattern_expr * Id.t + | RCPatCstr of Loc.t * Globnames.global_reference + * raw_cases_pattern_expr list * raw_cases_pattern_expr list + (** [RCPatCstr (loc, c, l1, l2)] represents ((@c l1) l2) *) + | RCPatAtom of Loc.t * Id.t option + | RCPatOr of Loc.t * raw_cases_pattern_expr list + +let raw_cases_pattern_expr_loc = function + | RCPatAlias (loc,_,_) -> loc + | RCPatCstr (loc,_,_,_) -> loc + | RCPatAtom (loc,_) -> loc + | RCPatOr (loc,_) -> loc + (** {6 Elemtary bricks } *) let apply_scope_env env = function | [] -> {env with tmp_scope = None}, [] @@ -1198,8 +1213,8 @@ let rec subst_pat_iterator y t p = match p with | RCPatAtom (_,id) -> begin match id with Some x when Id.equal x y -> t | _ -> p end | RCPatCstr (loc,id,l1,l2) -> - RCPatCstr (loc,id,List.map (subst_pat_iterator y t) l1, - List.map (subst_pat_iterator y t) l2) + RCPatCstr (loc,id,List.map (subst_pat_iterator y t) l1, + List.map (subst_pat_iterator y t) l2) | RCPatAlias (l,p,a) -> RCPatAlias (l,subst_pat_iterator y t p,a) | RCPatOr (l,pl) -> RCPatOr (l,List.map (subst_pat_iterator y t) pl) @@ -1216,6 +1231,14 @@ let drop_notations_pattern looked_for = let test_kind top = if top then looked_for else function ConstructRef _ -> () | _ -> raise Not_found in + (** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *) + let rec rcp_of_glob = function + | GVar (loc,id) -> RCPatAtom (loc,Some id) + | GHole (loc,_,_,_) -> RCPatAtom (loc,None) + | GRef (loc,g,_) -> RCPatCstr (loc, g,[],[]) + | GApp (loc,GRef (_,g,_),l) -> RCPatCstr (loc, g, List.map rcp_of_glob l,[]) + | _ -> CErrors.anomaly Pp.(str "Invalid return pattern from Notation.interp_prim_token_cases_pattern_expr ") + in let rec drop_syndef top scopes re pats = let (loc,qid) = qualid_of_reference re in try @@ -1285,8 +1308,9 @@ let drop_notations_pattern looked_for = let (argscs1,_) = find_remaining_scopes expl_pl pl g in RCPatCstr (loc, g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, []) | CPatNotation (loc,"- _",([CPatPrim(_,Numeral p)],[]),[]) - when Bigint.is_strictly_pos p -> - fst (Notation.interp_prim_token_cases_pattern_expr loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes) + when Bigint.is_strictly_pos p -> + let (pat, _df) = Notation.interp_prim_token_cases_pattern_expr loc (ensure_kind false loc) (Numeral (Bigint.neg p)) scopes in + rcp_of_glob pat | CPatNotation (_,"( _ )",([a],[]),[]) -> in_pat top scopes a | CPatNotation (loc, ntn, fullargs,extrargs) -> @@ -1299,7 +1323,9 @@ let drop_notations_pattern looked_for = in_not top loc scopes (subst,substlist) extrargs c | CPatDelimiters (loc, key, e) -> in_pat top (None,find_delimiters_scope loc key::snd scopes) e - | CPatPrim (loc,p) -> fst (Notation.interp_prim_token_cases_pattern_expr loc (test_kind false) p scopes) + | CPatPrim (loc,p) -> + let (pat, _df) = Notation.interp_prim_token_cases_pattern_expr loc (test_kind false) p scopes in + rcp_of_glob pat | CPatAtom (loc, Some id) -> begin match drop_syndef top scopes id [] with diff --git a/interp/notation.ml b/interp/notation.ml index 6aa6f54c0..7be2fe0f0 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -444,16 +444,20 @@ let notation_of_prim_token = function | Numeral n -> "- "^(to_string (neg n)) | String _ -> raise Not_found -let find_prim_token g loc p sc = +let find_prim_token check_allowed loc p sc = (* Try for a user-defined numerical notation *) try let (_,c),df = find_notation (notation_of_prim_token p) sc in - g (Notation_ops.glob_constr_of_notation_constr loc c),df + let pat = Notation_ops.glob_constr_of_notation_constr loc c in + check_allowed pat; + pat, df with Not_found -> (* Try for a primitive numerical notation *) let (spdir,interp) = Hashtbl.find prim_token_interpreter_tab sc loc p in check_required_module loc sc spdir; - g (interp ()), ((dirpath (fst spdir),DirPath.empty),"") + let pat = interp () in + check_allowed pat; + pat, ((dirpath (fst spdir),DirPath.empty),"") let interp_prim_token_gen g loc p local_scopes = let scopes = make_current_scopes local_scopes in @@ -466,20 +470,17 @@ let interp_prim_token_gen g loc p local_scopes = | String s -> str "No interpretation for string " ++ qs s) ++ str ".") let interp_prim_token = - interp_prim_token_gen (fun x -> x) + interp_prim_token_gen (fun _ -> ()) -(** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *) - -let rec rcp_of_glob looked_for = function - | GVar (loc,id) -> RCPatAtom (loc,Some id) - | GHole (loc,_,_,_) -> RCPatAtom (loc,None) - | GRef (loc,g,_) -> looked_for g; RCPatCstr (loc, g,[],[]) +let rec check_allowed_ref_in_pat looked_for = function + | GVar _ | GHole _ -> () + | GRef (_,g,_) -> looked_for g | GApp (loc,GRef (_,g,_),l) -> - looked_for g; RCPatCstr (loc, g, List.map (rcp_of_glob looked_for) l,[]) + looked_for g; List.iter (check_allowed_ref_in_pat looked_for) l | _ -> raise Not_found let interp_prim_token_cases_pattern_expr loc looked_for p = - interp_prim_token_gen (rcp_of_glob looked_for) loc p + interp_prim_token_gen (check_allowed_ref_in_pat looked_for) loc p let interp_notation loc ntn local_scopes = let scopes = make_current_scopes local_scopes in diff --git a/interp/notation.mli b/interp/notation.mli index 2e92a00a8..300480ff1 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -85,8 +85,10 @@ val declare_string_interpreter : scope_name -> required_module -> val interp_prim_token : Loc.t -> prim_token -> local_scopes -> glob_constr * (notation_location * scope_name option) + +(* This function returns a glob_const representing a pattern *) val interp_prim_token_cases_pattern_expr : Loc.t -> (global_reference -> unit) -> prim_token -> - local_scopes -> raw_cases_pattern_expr * (notation_location * scope_name option) + local_scopes -> glob_constr * (notation_location * scope_name option) (** Return the primitive token associated to a [term]/[cases_pattern]; raise [No_match] if no such token *) diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli index a4a6eb909..77f052ddb 100644 --- a/intf/constrexpr.mli +++ b/intf/constrexpr.mli @@ -36,14 +36,6 @@ type prim_token = | Numeral of Bigint.bigint (** representation of integer literals that appear in Coq scripts. *) | String of string -type raw_cases_pattern_expr = - | RCPatAlias of Loc.t * raw_cases_pattern_expr * Id.t - | RCPatCstr of Loc.t * Globnames.global_reference - * raw_cases_pattern_expr list * raw_cases_pattern_expr list - (** [CPatCstr (_, c, l1, l2)] represents ((@c l1) l2) *) - | RCPatAtom of Loc.t * Id.t option - | RCPatOr of Loc.t * raw_cases_pattern_expr list - type instance_expr = Misctypes.glob_level list type cases_pattern_expr = -- cgit v1.2.3 From e4ca8679e6700cfd6563890eb7d9e4ee59bede57 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Mon, 15 May 2017 02:25:52 +0200 Subject: [interp] Rework check for casts inside patterns. 1969e10f25df0c913600099b7b98ea273a064017 introduced a check so a cast in a pattern is not a fatal error. We move this check to the internalization function, which is the logical place to raise it, removing a bit boilerplate code. --- interp/constrintern.ml | 51 +++++++++++++++----------------------------------- 1 file changed, 15 insertions(+), 36 deletions(-) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index b18341800..405d63dfe 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1335,8 +1335,21 @@ let drop_notations_pattern looked_for = | CPatAtom (loc,None) -> RCPatAtom (loc,None) | CPatOr (loc, pl) -> RCPatOr (loc,List.map (in_pat top scopes) pl) - | CPatCast _ -> - assert false + | CPatCast (loc,_,_) -> + (* We raise an error if the pattern contains a cast, due to + current restrictions on casts in patterns. Cast in patterns + are supportted only in local binders and only at top + level. In fact, they are currently eliminated by the + parser. The only reason why they are in the + [cases_pattern_expr] type is that the parser needs to factor + the "(c : t)" notation with user defined notations (such as + the pair). In the long term, we will try to support such + casts everywhere, and use them to print the domains of + lambdas in the encoding of match in constr. This check is + here and not in the parser because it would require + duplicating the levels of the [pattern] rule. *) + CErrors.user_err ~loc ~hdr:"drop_notations_pattern" + (Pp.strbrk "Casts are not supported in this pattern.") and in_pat_sc scopes x = in_pat false (x,snd scopes) and in_not top loc scopes (subst,substlist as fullsubst) args = function | NVar id -> @@ -1418,40 +1431,7 @@ let rec intern_pat genv aliases pat = check_or_pat_variables loc ids (List.tl idsl); (ids,List.flatten pl') -(* [check_no_patcast p] raises an error if [p] contains a cast. This code is a - bit ad-hoc, and is due to current restrictions on casts in patterns. We - support them only in local binders and only at top level. In fact, they are - currently eliminated by the parser. The only reason why they are in the - [cases_pattern_expr] type is that the parser needs to factor the "(c : t)" - notation with user defined notations (such as the pair). In the long term, we - will try to support such casts everywhere, and use them to print the domains - of lambdas in the encoding of match in constr. We put this check here and not - in the parser because it would require to duplicate the levels of the - [pattern] rule. *) -let rec check_no_patcast = function - | CPatCast (loc,_,_) -> - CErrors.user_err ~loc ~hdr:"check_no_patcast" - (Pp.strbrk "Casts are not supported here.") - | CPatDelimiters(_,_,p) - | CPatAlias(_,p,_) -> check_no_patcast p - | CPatCstr(_,_,opl,pl) -> - Option.iter (List.iter check_no_patcast) opl; - List.iter check_no_patcast pl - | CPatOr(_,pl) -> - List.iter check_no_patcast pl - | CPatNotation(_,_,subst,pl) -> - check_no_patcast_subst subst; - List.iter check_no_patcast pl - | CPatRecord(_,prl) -> - List.iter (fun (_,p) -> check_no_patcast p) prl - | CPatAtom _ | CPatPrim _ -> () - -and check_no_patcast_subst (pl,pll) = - List.iter check_no_patcast pl; - List.iter (List.iter check_no_patcast) pll - let intern_cases_pattern genv scopes aliases pat = - check_no_patcast pat; intern_pat genv aliases (drop_notations_pattern (function ConstructRef _ -> () | _ -> raise Not_found) scopes pat) @@ -1460,7 +1440,6 @@ let _ = fun scopes p -> intern_cases_pattern (Global.env ()) scopes empty_alias p let intern_ind_pattern genv scopes pat = - check_no_patcast pat; let no_not = try drop_notations_pattern (function (IndRef _ | ConstructRef _) -> () | _ -> raise Not_found) scopes pat -- cgit v1.2.3 From 9ddfdab6a4715a08a78296bf8824d086f358bdc0 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 14 May 2017 11:14:43 +0200 Subject: Dead code in coqdep. Was introduced in 5268efdef, reverted then readded in 1be9c4d. --- tools/coqdep_common.ml | 7 ------- tools/coqdep_common.mli | 3 --- 2 files changed, 10 deletions(-) diff --git a/tools/coqdep_common.ml b/tools/coqdep_common.ml index 0064aebda..93b25e2ed 100644 --- a/tools/coqdep_common.ml +++ b/tools/coqdep_common.ml @@ -544,13 +544,6 @@ let add_rec_dir_no_import add_file phys_dir log_dir = let add_rec_dir_import add_file phys_dir log_dir = add_directory true (add_file true) phys_dir log_dir -(** -R semantic but only on immediate capitalized subdirs *) - -let add_rec_uppercase_subdirs add_file phys_dir log_dir = - process_subdirectories (fun phys_dir f -> - add_directory true (add_file true) phys_dir (log_dir@[String.capitalize f])) - phys_dir - (** -I semantic: do not go in subdirs. *) let add_caml_dir phys_dir = add_directory false add_caml_known phys_dir [] diff --git a/tools/coqdep_common.mli b/tools/coqdep_common.mli index 633c474ad..10da0240d 100644 --- a/tools/coqdep_common.mli +++ b/tools/coqdep_common.mli @@ -64,8 +64,5 @@ val add_rec_dir_no_import : val add_rec_dir_import : (bool -> string -> string list -> string -> unit) -> string -> string list -> unit -val add_rec_uppercase_subdirs : - (bool -> string -> string list -> string -> unit) -> string -> string list -> unit - val treat_file : dir -> string -> unit val error_cannot_parse : string -> int * int -> 'a -- cgit v1.2.3 From e2de94b90e8802fa5c5dc33c7daf6b8ce5646bfa Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 14 May 2017 00:10:57 +0200 Subject: Fixing a bug with nested "as" clauses in "match". --- interp/constrintern.ml | 28 ++++++++++++++-------------- test-suite/success/Cases.v | 5 +++++ 2 files changed, 19 insertions(+), 14 deletions(-) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index c916fcd88..80de11e3e 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -947,17 +947,6 @@ let find_remaining_scopes pl1 pl2 ref = in ((try List.firstn len_pl1 allscs with Failure _ -> simple_adjust_scopes len_pl1 allscs), simple_adjust_scopes len_pl2 (aux (impl_list,scope_list))) -let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2 - -let product_of_cases_patterns ids idspl = - List.fold_right (fun (ids,pl) (ids',ptaill) -> - (ids @ ids', - (* Cartesian prod of the or-pats for the nth arg and the tail args *) - List.flatten ( - List.map (fun (subst,p) -> - List.map (fun (subst',ptail) -> (merge_subst subst subst',p::ptail)) ptaill) pl))) - idspl (ids,[Id.Map.empty,[]]) - (* @return the first variable that occurs twice in a pattern naive n^2 algo *) @@ -1212,6 +1201,17 @@ let alias_of als = match als.alias_ids with *) +let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2 + +let product_of_cases_patterns aliases idspl = + List.fold_right (fun (ids,pl) (ids',ptaill) -> + (ids @ ids', + (* Cartesian prod of the or-pats for the nth arg and the tail args *) + List.flatten ( + List.map (fun (subst,p) -> + List.map (fun (subst',ptail) -> (merge_subst subst subst',p::ptail)) ptaill) pl))) + idspl (aliases.alias_ids,[aliases.alias_map,[]]) + let rec subst_pat_iterator y t p = match p with | RCPatAtom (_,id) -> begin match id with Some x when Id.equal x y -> t | _ -> p end @@ -1376,7 +1376,7 @@ let drop_notations_pattern looked_for = let rec intern_pat genv aliases pat = let intern_cstr_with_all_args loc c with_letin idslpl1 pl2 = let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in - let (ids',pll) = product_of_cases_patterns aliases.alias_ids (idslpl1@idslpl2) in + let (ids',pll) = product_of_cases_patterns aliases (idslpl1@idslpl2) in let pl' = List.map (fun (asubst,pl) -> (asubst, PatCstr (loc,c,chop_params_pattern loc (fst c) pl with_letin,alias_of aliases))) pll in ids',pl' in @@ -1466,7 +1466,7 @@ let intern_ind_pattern genv scopes pat = let idslpl1 = List.rev_map (intern_pat genv empty_alias) expl_pl in let idslpl2 = List.map (intern_pat genv empty_alias) pl2 in (with_letin, - match product_of_cases_patterns [] (List.rev_append idslpl1 idslpl2) with + match product_of_cases_patterns empty_alias (List.rev_append idslpl1 idslpl2) with | _,[_,pl] -> (c,chop_params_pattern loc c pl with_letin) | _ -> error_bad_inductive_type loc) | x -> error_bad_inductive_type (raw_cases_pattern_expr_loc x) @@ -1796,7 +1796,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = and intern_multiple_pattern env n (loc,pl) = let idsl_pll = List.map (intern_cases_pattern globalenv (None,env.scopes) empty_alias) pl in check_number_of_pattern loc n pl; - product_of_cases_patterns [] idsl_pll + product_of_cases_patterns empty_alias idsl_pll (* Expands a disjunction of multiple pattern *) and intern_disjunctive_multiple_pattern env loc n mpl = diff --git a/test-suite/success/Cases.v b/test-suite/success/Cases.v index 49c465b6c..52fe98ac0 100644 --- a/test-suite/success/Cases.v +++ b/test-suite/success/Cases.v @@ -1868,3 +1868,8 @@ Definition transport {A} (P : A->Type) {x y : A} (p : x=y) (u : P x) : P y := Check match eq_refl 0 in _=O return O=O with eq_refl => eq_refl end. Check match niln in listn O return O=O with niln => eq_refl end. + +(* A test about nested "as" clauses *) +(* (was failing up to May 2017) *) + +Check fun x => match x with (y,z) as t as w => (y+z,t) = (0,w) end. -- cgit v1.2.3 From 00964706efe8f6b13e38b57ddb45fac516feb958 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 13 May 2017 23:31:08 +0200 Subject: Fixing bug #5222 (anomaly with "`pat" in the presence of scope delimiters). We seized this opportunity to factorize the code for interning `pat with more general pre-existing code. More precisely: There was already a function to compute the free variables of a pattern. Commit c6d9d4fb rewrote an approximation of it and #5222 hits cases non-treated by this function. We remove the partially-defined redundant code and use instead the existing code since intern_cases_pattern, already called, was already doing this computation. (In doing so, we discover a bug in merging names in the presence of nested "as" clauses, which we fix in previous commit. Additionally, intern_local_pattern was misleadingly overkill to simply mean a folding on Id.Set.add and we avoid the detour. --- interp/constrintern.ml | 35 ++++++----------------------------- test-suite/bugs/closed/5522.v | 7 +++++++ 2 files changed, 13 insertions(+), 29 deletions(-) create mode 100644 test-suite/bugs/closed/5522.v diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 80de11e3e..4e76fe9aa 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -432,31 +432,6 @@ let intern_assumption intern lvar env nal bk ty = let env, b = intern_generalized_binder intern_type lvar env (List.hd nal) b b' t ty in env, b -let rec free_vars_of_pat il = - function - | CPatCstr (loc, c, l1, l2) -> - let il = List.fold_left free_vars_of_pat il (Option.default [] l1) in - List.fold_left free_vars_of_pat il l2 - | CPatAtom (loc, ro) -> - begin match ro with - | Some (Ident (loc, i)) -> (loc, i) :: il - | Some _ | None -> il - end - | CPatNotation (loc, n, l1, l2) -> - let il = List.fold_left free_vars_of_pat il (fst l1) in - List.fold_left (List.fold_left free_vars_of_pat) il (snd l1) - | _ -> anomaly (str "free_vars_of_pat") - -let intern_local_pattern intern lvar env p = - List.fold_left - (fun env (loc, i) -> - let bk = Default Implicit in - let ty = CHole (loc, None, Misctypes.IntroAnonymous, None) in - let n = Name i in - let env, _ = intern_assumption intern lvar env [(loc, n)] bk ty in - env) - env (free_vars_of_pat [] p) - type binder_data = | BDRawDef of (Loc.t * glob_binder) | BDPattern of @@ -490,13 +465,15 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio | Some ty -> ty | None -> CHole(loc,None,Misctypes.IntroAnonymous,None) in - let env = intern_local_pattern intern lvar env p in - let cp = + let il,cp = match !intern_cases_pattern_fwd (None,env.scopes) p with - | (_, [(_, cp)]) -> cp + | (il, [(subst,cp)]) -> + if not (Id.Map.equal Id.equal subst Id.Map.empty) then + user_err_loc (loc,"",str "Unsupported nested \"as\" clause."); + il,cp | _ -> assert false in - let il = List.map snd (free_vars_of_pat [] p) in + let env = {env with ids = List.fold_right Id.Set.add il env.ids} in (env, BDPattern(loc,(cp,il),lvar,env,tyc) :: bl) let intern_generalization intern env lvar loc bk ak c = diff --git a/test-suite/bugs/closed/5522.v b/test-suite/bugs/closed/5522.v new file mode 100644 index 000000000..0fae9ede4 --- /dev/null +++ b/test-suite/bugs/closed/5522.v @@ -0,0 +1,7 @@ +(* Checking support for scope delimiters and as clauses in 'pat + applied to notations with binders *) + +Notation "'multifun' x .. y 'in' f" := (fun x => .. (fun y => f) .. ) + (at level 200, x binder, y binder, f at level 200). + +Check multifun '((x, y)%core as z) in (x+y,0)=z. -- cgit v1.2.3 From 697cd5a8e7927873ed6700c7e906ae3675bd98b1 Mon Sep 17 00:00:00 2001 From: Pierre Courtieu Date: Fri, 12 May 2017 10:30:50 +0200 Subject: Simplified compaction criterion + tests. --- printing/printer.ml | 25 +++++-------------------- test-suite/output/CompactContexts.out | 7 +++++++ test-suite/output/CompactContexts.v | 5 +++++ 3 files changed, 17 insertions(+), 20 deletions(-) create mode 100644 test-suite/output/CompactContexts.out create mode 100644 test-suite/output/CompactContexts.v diff --git a/printing/printer.ml b/printing/printer.ml index 997d866f9..3feea6596 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -349,27 +349,12 @@ let pr_ne_context_of header env sigma = List.is_empty (Environ.named_context env) then (mt ()) else let penv = pr_context_unlimited env sigma in (header ++ penv ++ fnl ()) -(* Heuristic for horizontalizing hypothesis: - Detecting variable which type is a simple id or of the form (t x y ...) - where t is a product or only sorts (typically [Type -> Type -> ...] - and not [nat -> nat -> ...] ). - + Special case for non-Prop dependent terms. *) -let rec should_compact env sigma typ = +(* Heuristic for horizontalizing hypothesis that the user probably + considers as "variables": An hypothesis H:T where T:S and S<>Prop. *) +let should_compact env sigma typ = get_compact_context() && - match kind_of_term typ with - | Rel _ | Var _ | Sort _ | Const _ | Ind _ -> true - | App (c,args) -> - let _,type_of_c = Typing.type_of env sigma (EConstr.of_constr c) in - let _,type_of_typ = Typing.type_of env sigma (EConstr.of_constr typ) in - not (is_Prop (EConstr.to_constr sigma type_of_typ)) - && (* These two more tests detect rare cases of non-Prop-sorted - dependent hypothesis: *) - let lnamedtyp , _ = EConstr.decompose_prod sigma type_of_c in - (* c has a non dependent type *) - List.for_all (fun (_,typarg) -> EConstr.isSort sigma typarg) lnamedtyp - && (* and real arguments are recursively elligible to compaction. *) - Array.for_all (should_compact env sigma) args - | _ -> false + let type_of_typ = Retyping.get_type_of env sigma (EConstr.of_constr typ) in + not (is_Prop (EConstr.to_constr sigma type_of_typ)) (* If option Compact Contexts is set, we pack "simple" hypothesis in a diff --git a/test-suite/output/CompactContexts.out b/test-suite/output/CompactContexts.out new file mode 100644 index 000000000..9d1d19877 --- /dev/null +++ b/test-suite/output/CompactContexts.out @@ -0,0 +1,7 @@ +1 subgoal + + hP1 : True + a : nat b : list nat h : forall x : nat, {y : nat | y > x} + h2 : True + ============================ + False diff --git a/test-suite/output/CompactContexts.v b/test-suite/output/CompactContexts.v new file mode 100644 index 000000000..07588d34f --- /dev/null +++ b/test-suite/output/CompactContexts.v @@ -0,0 +1,5 @@ +Set Printing Compact Contexts. + +Lemma f (hP1:True) (a:nat) (b:list nat) (h:forall (x:nat) , { y:nat | y > x}) (h2:True): False. +Show. +Abort. \ No newline at end of file -- cgit v1.2.3 From cd3971e53b76cb62e14822eb3e275d3968a4f215 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 15 May 2017 22:30:09 +0200 Subject: Adding support for using grammar entries returning no value in EXTEND. --- dev/top_printers.ml | 4 ++-- grammar/argextend.mlp | 12 ++++++++---- grammar/q_util.mli | 2 +- grammar/q_util.mlp | 2 +- grammar/tacextend.mlp | 17 ++++++++++------- grammar/vernacextend.mlp | 16 ++++++++++------ parsing/egramml.ml | 4 ++-- parsing/egramml.mli | 2 +- plugins/ltac/g_ltac.ml4 | 8 ++++++-- plugins/ltac/pptactic.ml | 7 ++++--- plugins/ltac/pptactic.mli | 2 +- plugins/ltac/tacentries.ml | 14 +++++++------- plugins/ltac/tacentries.mli | 2 +- 13 files changed, 54 insertions(+), 38 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index f8498c402..ce6d5dff0 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -510,7 +510,7 @@ let _ = extend_vernac_command_grammar ("PrintConstr", 0) None [GramTerminal "PrintConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.ghost,Some (rawwit wit_constr),Extend.Aentry Pcoq.Constr.constr)] let _ = try @@ -526,7 +526,7 @@ let _ = extend_vernac_command_grammar ("PrintPureConstr", 0) None [GramTerminal "PrintPureConstr"; GramNonTerminal - (Loc.ghost,rawwit wit_constr,Extend.Aentry Pcoq.Constr.constr)] + (Loc.ghost,Some (rawwit wit_constr),Extend.Aentry Pcoq.Constr.constr)] (* Setting printer of unbound global reference *) open Names diff --git a/grammar/argextend.mlp b/grammar/argextend.mlp index d00ee4e5d..5cfcc6fd2 100644 --- a/grammar/argextend.mlp +++ b/grammar/argextend.mlp @@ -40,7 +40,8 @@ let make_topwit loc arg = <:expr< Genarg.topwit $make_wit loc arg$ >> let make_act loc act pil = let rec make = function | [] -> <:expr< (fun loc -> $act$) >> - | ExtNonTerminal (_, p) :: tl -> <:expr< (fun $lid:p$ -> $make tl$) >> + | ExtNonTerminal (_, None) :: tl -> <:expr< (fun $lid:"_"$ -> $make tl$) >> + | ExtNonTerminal (_, Some p) :: tl -> <:expr< (fun $lid:p$ -> $make tl$) >> | ExtTerminal _ :: tl -> <:expr< (fun _ -> $make tl$) >> in make (List.rev pil) @@ -63,7 +64,7 @@ let is_ident x = function | _ -> false let make_extend loc s cl wit = match cl with -| [[ExtNonTerminal (Uentry e, id)], act] when is_ident id act -> +| [[ExtNonTerminal (Uentry e, Some id)], act] when is_ident id act -> (** Special handling of identity arguments by not redeclaring an entry *) <:str_item< value $lid:s$ = @@ -246,10 +247,13 @@ EXTEND genarg: [ [ e = LIDENT; "("; s = LIDENT; ")" -> let e = parse_user_entry e "" in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) | e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" -> let e = parse_user_entry e sep in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) + | e = LIDENT -> + let e = parse_user_entry e "" in + ExtNonTerminal (e, None) | s = STRING -> ExtTerminal s ] ] ; diff --git a/grammar/q_util.mli b/grammar/q_util.mli index 5b3eb3202..0b3def058 100644 --- a/grammar/q_util.mli +++ b/grammar/q_util.mli @@ -23,7 +23,7 @@ type user_symbol = type extend_token = | ExtTerminal of string -| ExtNonTerminal of user_symbol * string +| ExtNonTerminal of user_symbol * string option val mlexpr_of_list : ('a -> MLast.expr) -> 'a list -> MLast.expr diff --git a/grammar/q_util.mlp b/grammar/q_util.mlp index 8b930cf36..87262e1c8 100644 --- a/grammar/q_util.mlp +++ b/grammar/q_util.mlp @@ -25,7 +25,7 @@ type user_symbol = type extend_token = | ExtTerminal of string -| ExtNonTerminal of user_symbol * string +| ExtNonTerminal of user_symbol * string option let mlexpr_of_list f l = List.fold_right diff --git a/grammar/tacextend.mlp b/grammar/tacextend.mlp index b14fba975..3057ee58c 100644 --- a/grammar/tacextend.mlp +++ b/grammar/tacextend.mlp @@ -31,19 +31,19 @@ let mlexpr_of_ident id = let rec make_patt = function | [] -> <:patt< [] >> - | ExtNonTerminal (_, p) :: l -> + | ExtNonTerminal (_, Some p) :: l -> <:patt< [ $lid:p$ :: $make_patt l$ ] >> | _::l -> make_patt l let rec make_let raw e = function | [] -> <:expr< fun $lid:"ist"$ -> $e$ >> - | ExtNonTerminal (g, p) :: l -> + | ExtNonTerminal (g, Some p) :: l -> let t = type_of_user_symbol g in let loc = MLast.loc_of_expr e in let e = make_let raw e l in let v = if raw then <:expr< Genarg.out_gen $make_rawwit loc t$ $lid:p$ >> - else <:expr< Tacinterp.Value.cast $make_topwit loc t$ $lid:p$ >> in + else <:expr< Tacinterp.Value.cast $make_topwit loc t$ $lid:p$ >> in <:expr< let $lid:p$ = $v$ in $e$ >> | _::l -> make_let raw e l @@ -75,7 +75,7 @@ let rec mlexpr_of_symbol = function let make_prod_item = function | ExtTerminal s -> <:expr< Tacentries.TacTerm $str:s$ >> | ExtNonTerminal (g, id) -> - <:expr< Tacentries.TacNonTerm $default_loc$ $mlexpr_of_symbol g$ $mlexpr_of_ident id$ >> + <:expr< Tacentries.TacNonTerm $default_loc$ $mlexpr_of_symbol g$ $mlexpr_of_option mlexpr_of_ident id$ >> let mlexpr_of_clause cl = mlexpr_of_list (fun (a,_,_) -> mlexpr_of_list make_prod_item a) cl @@ -87,7 +87,7 @@ let is_constr_gram = function | _ -> false let make_var = function - | ExtNonTerminal (_, p) -> Some p + | ExtNonTerminal (_, p) -> p | _ -> assert false let declare_tactic loc tacname ~level classification clause = match clause with @@ -158,10 +158,13 @@ EXTEND tacargs: [ [ e = LIDENT; "("; s = LIDENT; ")" -> let e = parse_user_entry e "" in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) | e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" -> let e = parse_user_entry e sep in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) + | e = LIDENT -> + let e = parse_user_entry e "" in + ExtNonTerminal (e, None) | s = STRING -> let () = if s = "" then failwith "Empty terminal." in ExtTerminal s diff --git a/grammar/vernacextend.mlp b/grammar/vernacextend.mlp index 436a7f0d9..7c99b52e8 100644 --- a/grammar/vernacextend.mlp +++ b/grammar/vernacextend.mlp @@ -27,7 +27,7 @@ type rule = { let rec make_let e = function | [] -> e - | ExtNonTerminal (g, p) :: l -> + | ExtNonTerminal (g, Some p) :: l -> let t = type_of_user_symbol g in let loc = MLast.loc_of_expr e in let e = make_let e l in @@ -42,7 +42,7 @@ let make_clause { r_patt = pt; r_branch = e; } = (* To avoid warnings *) let mk_ignore c pt = let fold accu = function - | ExtNonTerminal (_, p) -> p :: accu + | ExtNonTerminal (_, Some p) -> p :: accu | _ -> accu in let names = List.fold_left fold [] pt in @@ -101,10 +101,11 @@ let make_fun_classifiers loc s c l = let make_prod_item = function | ExtTerminal s -> <:expr< Egramml.GramTerminal $str:s$ >> - | ExtNonTerminal (g, id) -> + | ExtNonTerminal (g, ido) -> let nt = type_of_user_symbol g in let base s = <:expr< Pcoq.genarg_grammar ($mk_extraarg loc s$) >> in - <:expr< Egramml.GramNonTerminal $default_loc$ $make_rawwit loc nt$ + let typ = match ido with None -> None | Some _ -> Some nt in + <:expr< Egramml.GramNonTerminal $default_loc$ $mlexpr_of_option (make_rawwit loc) typ$ $mlexpr_of_prod_entry_key base g$ >> let mlexpr_of_clause cl = @@ -178,10 +179,13 @@ EXTEND args: [ [ e = LIDENT; "("; s = LIDENT; ")" -> let e = parse_user_entry e "" in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) | e = LIDENT; "("; s = LIDENT; ","; sep = STRING; ")" -> let e = parse_user_entry e sep in - ExtNonTerminal (e, s) + ExtNonTerminal (e, Some s) + | e = LIDENT -> + let e = parse_user_entry e "" in + ExtNonTerminal (e, None) | s = STRING -> ExtTerminal s ] ] diff --git a/parsing/egramml.ml b/parsing/egramml.ml index 97a3e89a5..984957589 100644 --- a/parsing/egramml.ml +++ b/parsing/egramml.ml @@ -17,7 +17,7 @@ open Vernacexpr type 's grammar_prod_item = | GramTerminal of string | GramNonTerminal : - Loc.t * 'a raw_abstract_argument_type * ('s, 'a) symbol -> 's grammar_prod_item + Loc.t * 'a raw_abstract_argument_type option * ('s, 'a) symbol -> 's grammar_prod_item type 'a ty_arg = ('a -> raw_generic_argument) @@ -38,7 +38,7 @@ let rec ty_rule_of_gram = function AnyTyRule r | GramNonTerminal (_, t, tok) :: rem -> let AnyTyRule rem = ty_rule_of_gram rem in - let inj = Some (fun obj -> Genarg.in_gen t obj) in + let inj = Option.map (fun t obj -> Genarg.in_gen t obj) t in let r = TyNext (rem, tok, inj) in AnyTyRule r diff --git a/parsing/egramml.mli b/parsing/egramml.mli index 1ad947200..29baaf052 100644 --- a/parsing/egramml.mli +++ b/parsing/egramml.mli @@ -15,7 +15,7 @@ open Vernacexpr type 's grammar_prod_item = | GramTerminal of string - | GramNonTerminal : Loc.t * 'a Genarg.raw_abstract_argument_type * + | GramNonTerminal : Loc.t * 'a Genarg.raw_abstract_argument_type option * ('s, 'a) Extend.symbol -> 's grammar_prod_item val extend_vernac_command_grammar : diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4 index ca5d198c2..d717ed0a5 100644 --- a/plugins/ltac/g_ltac.ml4 +++ b/plugins/ltac/g_ltac.ml4 @@ -460,7 +460,9 @@ END let pr_ltac_production_item = function | Tacentries.TacTerm s -> quote (str s) -| Tacentries.TacNonTerm (_, (arg, sep), id) -> +| Tacentries.TacNonTerm (_, (arg, None), None) -> str arg +| Tacentries.TacNonTerm (_, (arg, Some _), None) -> assert false +| Tacentries.TacNonTerm (_, (arg, sep), Some id) -> let sep = match sep with | None -> mt () | Some sep -> str "," ++ spc () ++ quote (str sep) @@ -470,7 +472,9 @@ let pr_ltac_production_item = function VERNAC ARGUMENT EXTEND ltac_production_item PRINTED BY pr_ltac_production_item | [ string(s) ] -> [ Tacentries.TacTerm s ] | [ ident(nt) "(" ident(p) ltac_production_sep_opt(sep) ")" ] -> - [ Tacentries.TacNonTerm (loc, (Names.Id.to_string nt, sep), p) ] + [ Tacentries.TacNonTerm (loc, (Names.Id.to_string nt, sep), Some p) ] +| [ ident(nt) ] -> + [ Tacentries.TacNonTerm (loc, (Names.Id.to_string nt, None), None) ] END VERNAC COMMAND EXTEND VernacTacticNotation diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml index b73b66e56..a61957559 100644 --- a/plugins/ltac/pptactic.ml +++ b/plugins/ltac/pptactic.ml @@ -51,7 +51,7 @@ let pr_global x = Nametab.pr_global_env Id.Set.empty x type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of Loc.t * 'a * Names.Id.t option type grammar_terminals = Genarg.ArgT.any Extend.user_symbol grammar_tactic_prod_item_expr list @@ -264,8 +264,9 @@ type 'a extra_genarg_printer = let rec pack prods args = match prods, args with | [], [] -> [] | TacTerm s :: prods, args -> TacTerm s :: pack prods args - | TacNonTerm (loc, symb, id) :: prods, arg :: args -> - TacNonTerm (loc, (symb, arg), id) :: pack prods args + | TacNonTerm (_, _, None) :: prods, args -> pack prods args + | TacNonTerm (loc, symb, (Some _ as ido)) :: prods, arg :: args -> + TacNonTerm (loc, (symb, arg), ido) :: pack prods args | _ -> raise Not_found in let prods = pack pp.pptac_prods l in diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli index 729338fb9..433f342c4 100644 --- a/plugins/ltac/pptactic.mli +++ b/plugins/ltac/pptactic.mli @@ -21,7 +21,7 @@ open Ppextend type 'a grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of Loc.t * 'a * Names.Id.t option type 'a raw_extra_genarg_printer = (constr_expr -> std_ppcmds) -> diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index 32750383b..91262f6fd 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -21,7 +21,7 @@ open Nameops type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of Loc.t * 'a * Names.Id.t option type raw_argument = string * string option type argument = Genarg.ArgT.any Extend.user_symbol @@ -174,9 +174,9 @@ let add_tactic_entry (kn, ml, tg) state = in let map = function | TacTerm s -> GramTerminal s - | TacNonTerm (loc, s, _) -> + | TacNonTerm (loc, s, ido) -> let EntryName (typ, e) = prod_item_of_symbol tg.tacgram_level s in - GramNonTerminal (loc, typ, e) + GramNonTerminal (loc, Option.map (fun _ -> typ) ido, e) in let prods = List.map map tg.tacgram_prods in let rules = make_rule mkact prods in @@ -202,7 +202,7 @@ let register_tactic_notation_entry name entry = let interp_prod_item = function | TacTerm s -> TacTerm s - | TacNonTerm (loc, (nt, sep), id) -> + | TacNonTerm (loc, (nt, sep), ido) -> let symbol = parse_user_entry nt sep in let interp s = function | None -> @@ -220,7 +220,7 @@ let interp_prod_item = function end in let symbol = interp_entry_name interp symbol in - TacNonTerm (loc, symbol, id) + TacNonTerm (loc, symbol, ido) let make_fresh_key = let id = Summary.ref ~name:"TACTIC-NOTATION-COUNTER" 0 in @@ -296,7 +296,7 @@ let inTacticGrammar : tactic_grammar_obj -> obj = let cons_production_parameter = function | TacTerm _ -> None -| TacNonTerm (_, _, id) -> Some id +| TacNonTerm (_, _, ido) -> ido let add_glob_tactic_notation local ~level prods forml ids tac = let parule = { @@ -362,7 +362,7 @@ let add_ml_tactic_notation name ~level prods = let open Tacexpr in let get_id = function | TacTerm s -> None - | TacNonTerm (_, _, id) -> Some id + | TacNonTerm (_, _, ido) -> ido in let ids = List.map_filter get_id prods in let entry = { mltac_name = name; mltac_index = len - i - 1 } in diff --git a/plugins/ltac/tacentries.mli b/plugins/ltac/tacentries.mli index 069504473..dac62dad3 100644 --- a/plugins/ltac/tacentries.mli +++ b/plugins/ltac/tacentries.mli @@ -20,7 +20,7 @@ val register_ltac : locality_flag -> Tacexpr.tacdef_body list -> unit type 'a grammar_tactic_prod_item_expr = 'a Pptactic.grammar_tactic_prod_item_expr = | TacTerm of string -| TacNonTerm of Loc.t * 'a * Names.Id.t +| TacNonTerm of Loc.t * 'a * Names.Id.t option type raw_argument = string * string option (** An argument type as provided in Tactic notations, i.e. a string like -- cgit v1.2.3 From b82f27726f5ae891689e3b958323c2a61d4c154b Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Mon, 15 May 2017 22:31:08 +0200 Subject: Fixing grammar for "evar" by exporting the test_lpar_id_colon trick to EXTEND. --- plugins/ltac/extraargs.ml4 | 20 ++++++++++++++++++++ plugins/ltac/extraargs.mli | 4 ++++ plugins/ltac/extratactics.ml4 | 2 +- plugins/ltac/g_tactic.ml4 | 13 +------------ test-suite/success/ltac.v | 6 ++++++ 5 files changed, 32 insertions(+), 13 deletions(-) diff --git a/plugins/ltac/extraargs.ml4 b/plugins/ltac/extraargs.ml4 index 53b726432..ec3a49df4 100644 --- a/plugins/ltac/extraargs.ml4 +++ b/plugins/ltac/extraargs.ml4 @@ -274,6 +274,26 @@ ARGUMENT EXTEND in_clause | [ in_clause'(cl) ] -> [ cl ] END +let local_test_lpar_id_colon = + let err () = raise Stream.Failure in + Pcoq.Gram.Entry.of_parser "lpar_id_colon" + (fun strm -> + match Util.stream_nth 0 strm with + | Tok.KEYWORD "(" -> + (match Util.stream_nth 1 strm with + | Tok.IDENT _ -> + (match Util.stream_nth 2 strm with + | Tok.KEYWORD ":" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) + +let pr_lpar_id_colon _ _ _ _ = mt () + +ARGUMENT EXTEND test_lpar_id_colon TYPED AS unit PRINTED BY pr_lpar_id_colon +| [ local_test_lpar_id_colon(x) ] -> [ () ] +END + (* spiwack: the print functions are incomplete, but I don't know what they are used for *) let pr_r_nat_field natf = diff --git a/plugins/ltac/extraargs.mli b/plugins/ltac/extraargs.mli index 7d4bccfad..9b4167512 100644 --- a/plugins/ltac/extraargs.mli +++ b/plugins/ltac/extraargs.mli @@ -67,6 +67,10 @@ val pr_by_arg_tac : (int * Ppextend.parenRelation -> raw_tactic_expr -> Pp.std_ppcmds) -> raw_tactic_expr option -> Pp.std_ppcmds +val test_lpar_id_colon : unit Pcoq.Gram.entry + +val wit_test_lpar_id_colon : (unit, unit, unit) Genarg.genarg_type + (** Spiwack: Primitive for retroknowledge registration *) val retroknowledge_field : Retroknowledge.field Pcoq.Gram.entry diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 3e6ccaf84..bd48614db 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -463,7 +463,7 @@ open Evar_tactics (* TODO: add support for some test similar to g_constr.name_colon so that expressions like "evar (list A)" do not raise a syntax error *) TACTIC EXTEND evar - [ "evar" "(" ident(id) ":" lconstr(typ) ")" ] -> [ let_evar (Name id) typ ] + [ "evar" test_lpar_id_colon "(" ident(id) ":" lconstr(typ) ")" ] -> [ let_evar (Name id) typ ] | [ "evar" constr(typ) ] -> [ let_evar Anonymous typ ] END diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4 index 4b3ca80af..e33c25cf8 100644 --- a/plugins/ltac/g_tactic.ml4 +++ b/plugins/ltac/g_tactic.ml4 @@ -72,18 +72,7 @@ let test_lpar_idnum_coloneq = | _ -> err ()) (* idem for (x:t) *) -let test_lpar_id_colon = - Gram.Entry.of_parser "lpar_id_colon" - (fun strm -> - match stream_nth 0 strm with - | KEYWORD "(" -> - (match stream_nth 1 strm with - | IDENT _ -> - (match stream_nth 2 strm with - | KEYWORD ":" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) +open Extraargs (* idem for (x1..xn:t) [n^2 complexity but exceptional use] *) let check_for_coloneq = diff --git a/test-suite/success/ltac.v b/test-suite/success/ltac.v index ce9099059..9ab47fede 100644 --- a/test-suite/success/ltac.v +++ b/test-suite/success/ltac.v @@ -317,3 +317,9 @@ let T := constr:(fun a b : nat => a) in end. exact (eq_refl n). Qed. + +(* Test evar syntax *) + +Goal True. +evar (0=0). +Abort. -- cgit v1.2.3 From 530ce019b7bdcc2603027082f6b3f6841d5990e1 Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Tue, 16 May 2017 18:28:41 +0200 Subject: Stop using auto with * in two proofs. auto with * is an overkill for people who do not care to understand what they really need. In these two cases, one lemma which was available in the typeclass_instances hint db. --- theories/Structures/DecidableType.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Structures/DecidableType.v b/theories/Structures/DecidableType.v index f85222dfb..d811f04ef 100644 --- a/theories/Structures/DecidableType.v +++ b/theories/Structures/DecidableType.v @@ -86,7 +86,7 @@ Module KeyDecidableType(D:DecidableType). Lemma InA_eqk : forall p q m, eqk p q -> InA eqk p m -> InA eqk q m. Proof. - intros; apply InA_eqA with p; auto with *. + intros; apply InA_eqA with p; auto using eqk_equiv. Qed. Definition MapsTo (k:key)(e:elt):= InA eqke (k,e). @@ -109,7 +109,7 @@ Module KeyDecidableType(D:DecidableType). Lemma MapsTo_eq : forall l x y e, eq x y -> MapsTo x e l -> MapsTo y e l. Proof. - intros; unfold MapsTo in *; apply InA_eqA with (x,e); eauto with *. + intros; unfold MapsTo in *; apply InA_eqA with (x,e); auto using eqke_equiv. Qed. Lemma In_eq : forall l x y, eq x y -> In x l -> In y l. -- cgit v1.2.3 From f6b0d8b78ae25c8b1c6b901e57a5f4e38f20cdbd Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 5 May 2017 13:19:21 +0200 Subject: Travis: do not run the tests if building Coq fails --- .travis.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.travis.yml b/.travis.yml index 560455187..8dcc34a8d 100644 --- a/.travis.yml +++ b/.travis.yml @@ -115,6 +115,7 @@ install: script: +- set -e - echo 'Configuring Coq...' && echo -en 'travis_fold:start:coq.config\\r' - ./configure -local -usecamlp5 -native-compiler yes ${EXTRA_CONF} - echo -en 'travis_fold:end:coq.config\\r' @@ -126,6 +127,7 @@ script: - echo 'Running tests...' && echo -en 'travis_fold:start:coq.test\\r' - ${TW} make -j ${NJOBS} ${TEST_TARGET} - echo -en 'travis_fold:end:coq.test\\r' +- set +e # Testing Gitter webhook notifications: -- cgit v1.2.3 From 8252ff7ef7159a2493dd5ac76a647a8b96a5a692 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 5 May 2017 14:11:12 +0200 Subject: Travis: deduplicate package list for coqide+documentation targets --- .travis.yml | 18 ++---------------- 1 file changed, 2 insertions(+), 16 deletions(-) diff --git a/.travis.yml b/.travis.yml index 8dcc34a8d..158f61cfd 100644 --- a/.travis.yml +++ b/.travis.yml @@ -65,7 +65,7 @@ matrix: apt: sources: - avsm - packages: + packages: &extra-packages - opam - aspcud - libgtk2.0-dev @@ -90,21 +90,7 @@ matrix: apt: sources: - avsm - packages: - - opam - - aspcud - - libgtk2.0-dev - - libgtksourceview2.0-dev - - texlive-latex-base - - texlive-latex-recommended - - texlive-latex-extra - - texlive-math-extra - - texlive-fonts-recommended - - texlive-fonts-extra - - latex-xcolor - - ghostscript - - transfig - - imagemagick + packages: *extra-packages install: - opam init -j ${NJOBS} --compiler=${COMPILER} -n -y -- cgit v1.2.3 From d8d56dfadc571fdf23ff9f8cb97d4c8cd96691ee Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 5 May 2017 14:24:23 +0200 Subject: Travis: add -warn-error targets (standard and 4.04.1 ocaml) --- .travis.yml | 28 ++++++++++++++++++++++++++++ Makefile.dev | 5 ++++- 2 files changed, 32 insertions(+), 1 deletion(-) diff --git a/.travis.yml b/.travis.yml index 158f61cfd..06ce3cae2 100644 --- a/.travis.yml +++ b/.travis.yml @@ -91,6 +91,34 @@ matrix: sources: - avsm packages: *extra-packages + - env: + - TEST_TARGET="coqocaml" + - EXTRA_CONF="-coqide opt -warn-error" + - EXTRA_OPAM="lablgtk-extras hevea" + # dummy target + - BUILD_TARGET="clean" + addons: + apt: + sources: + - avsm + packages: &coqide-packages + - opam + - aspcud + - libgtk2.0-dev + - libgtksourceview2.0-dev + - env: + - TEST_TARGET="coqocaml" + - COMPILER="4.04.1" + - CAMLP5_VER="6.17" + - EXTRA_CONF="-coqide opt -warn-error" + - EXTRA_OPAM="lablgtk-extras hevea" + # dummy target + - BUILD_TARGET="clean" + addons: + apt: + sources: + - avsm + packages: *coqide-packages install: - opam init -j ${NJOBS} --compiler=${COMPILER} -n -y diff --git a/Makefile.dev b/Makefile.dev index 5931e46dd..1803cc8ff 100644 --- a/Makefile.dev +++ b/Makefile.dev @@ -97,7 +97,10 @@ minibyte: $(COQTOPBYTE) pluginsbyte pluginsopt: $(PLUGINSOPT) pluginsbyte: $(PLUGINS) -.PHONY: coqlight states miniopt minibyte pluginsopt pluginsbyte +# This should build all the ocaml code but not (most of) the .v files +coqocaml: tools coqbinaries pluginsopt coqide printers + +.PHONY: coqlight states miniopt minibyte pluginsopt pluginsbyte coqocaml ########################## ### 2) core ML components -- cgit v1.2.3 From e6883abb10bd394f3066e6883e563e3bab79fea2 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Tue, 16 May 2017 18:54:14 +0200 Subject: [toplevel] Restore 8.6 goal printing behavior. When porting the toplevel to the STM, the logic for goal printing was simplified too much under optimistic assumptions. The idea was not to rely on the `vernac_classifier` which is an internal and complicated beast. However, it seems there are many cases to consider other than `is_query`, so at the risk of reimplementing the classifier we revert to the old approach of using the full classification. This gives maximum 8.6 compatibility, with the pitfall of having to call the classifier. Indeed, due to the dynamic nature of the "undo classifier", we cannot call it after `Stm.add`, as the document tree will be not the right one, making the classification of undo commands incorrect (actually raising an error "Cannot undo"). --- toplevel/vernac.ml | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 4fca4ea18..e79ee1a42 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -124,6 +124,16 @@ let rec interp_vernac sid (loc,com) = let f = Loadpath.locate_file fname in load_vernac verbosely sid f | v -> + + (* XXX: We need to run this before add as the classification is + highly dynamic and depends on the structure of the + document. Hopefully this is fixed when VtBack can be removed + and Undo etc... are just interpreted regularly. *) + let is_proof_step = match fst (Vernac_classifier.classify_vernac v) with + | VtProofStep _ | VtStm (VtBack _, _) | VtStartProof _ -> true + | _ -> false + in + let nsid, ntip = Stm.add ~ontop:sid (not !Flags.quiet) (loc,v) in (* Main STM interaction *) @@ -137,11 +147,12 @@ let rec interp_vernac sid (loc,com) = if check_proof then Stm.finish (); (* We could use a more refined criteria that depends on the - vernac. For now we imitate the old approach. *) - let hide_goals = !Flags.batch_mode || is_query v || !Flags.quiet || - not (Proof_global.there_are_pending_proofs ()) in + vernac. For now we imitate the old approach and rely on the + classification. *) + let print_goals = not !Flags.batch_mode && not !Flags.quiet && + is_proof_step && Proof_global.there_are_pending_proofs () in - if not hide_goals then Feedback.msg_notice (pr_open_cur_subgoals ()); + if print_goals then Feedback.msg_notice (pr_open_cur_subgoals ()); nsid in try -- cgit v1.2.3 From 94e8664074cfb987f8a63ca29c5436c861184b3a Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 11 Sep 2016 11:55:08 +0200 Subject: Stopping injection not to work on discriminable atoms (see #4890). --- tactics/equality.ml | 19 ++++++++++--------- tactics/equality.mli | 3 +++ 2 files changed, 13 insertions(+), 9 deletions(-) diff --git a/tactics/equality.ml b/tactics/equality.ml index cc7701ad5..18a1f0201 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -736,7 +736,7 @@ let _ = optread = (fun () -> !keep_proof_equalities_for_injection) ; optwrite = (fun b -> keep_proof_equalities_for_injection := b) } -let find_positions env sigma t1 t2 = +let find_positions env sigma ~no_discr t1 t2 = let project env sorts posn t1 t2 = let ty1 = get_type_of env sigma t1 in let s = get_sort_family_of env sigma ty1 in @@ -767,7 +767,7 @@ let find_positions env sigma t1 t2 = List.flatten (List.map2_i (fun i -> findrec sorts' ((sp1,adjust i)::posn)) 0 rargs1 rargs2) - else if Sorts.List.mem InType sorts' + else if Sorts.List.mem InType sorts' && not no_discr then (* see build_discriminator *) raise (DiscrFound (List.rev posn,sp1,sp2)) else @@ -791,13 +791,14 @@ let find_positions env sigma t1 t2 = Inl (path,c1,c2) let discriminable env sigma t1 t2 = - match find_positions env sigma t1 t2 with + match find_positions env sigma ~no_discr:false t1 t2 with | Inl _ -> true | _ -> false let injectable env sigma t1 t2 = - match find_positions env sigma t1 t2 with - | Inl _ | Inr [] | Inr [([],_,_)] -> false + match find_positions env sigma ~no_discr:true t1 t2 with + | Inl _ -> assert false + | Inr [] | Inr [([],_,_)] -> false | Inr _ -> true @@ -1032,7 +1033,7 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause = let sigma = eq_clause.evd in Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in - match find_positions env sigma t1 t2 with + match find_positions env sigma ~no_discr:false t1 t2 with | Inr _ -> tclZEROMSG (str"Not a discriminable equality.") | Inl (cpath, (_,dirn), _) -> @@ -1414,9 +1415,9 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = let sigma = eq_clause.evd in let env = eq_clause.env in - match find_positions env sigma t1 t2 with + match find_positions env sigma ~no_discr:true t1 t2 with | Inl _ -> - tclZEROMSG (strbrk"This equality is discriminable. You should use the discriminate tactic to solve the goal.") + assert false | Inr [] -> let suggestion = if !keep_proof_equalities_for_injection then @@ -1483,7 +1484,7 @@ let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = Proofview.Goal.enter { enter = begin fun gl -> let sigma = clause.evd in let env = Proofview.Goal.env gl in - match find_positions env sigma t1 t2 with + match find_positions env sigma ~no_discr:false t1 t2 with | Inl (cpath, (_,dirn), _) -> discr_positions env sigma u clause cpath dirn | Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *) diff --git a/tactics/equality.mli b/tactics/equality.mli index d979c580a..b47be3bbc 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -96,7 +96,10 @@ val cutRewriteInConcl : bool -> constr -> unit Proofview.tactic val rewriteInHyp : bool -> constr -> Id.t -> unit Proofview.tactic val rewriteInConcl : bool -> constr -> unit Proofview.tactic +(* Tells if tactic "discriminate" is applicable *) val discriminable : env -> evar_map -> constr -> constr -> bool + +(* Tells if tactic "injection" is applicable *) val injectable : env -> evar_map -> constr -> constr -> bool (* Subst *) -- cgit v1.2.3 From 80fc13e3115058230586246e72d69e3eac467f73 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 24 Mar 2017 21:13:38 +0100 Subject: Documenting relaxing of constraints on injection. We seized this opportunity to do a little refreshing of the section describing injection. --- doc/refman/RefMan-tac.tex | 39 ++++++++++++++++----------------------- 1 file changed, 16 insertions(+), 23 deletions(-) diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex index 87b9e4914..fc3fdd002 100644 --- a/doc/refman/RefMan-tac.tex +++ b/doc/refman/RefMan-tac.tex @@ -2187,32 +2187,24 @@ the given bindings to instantiate parameters or hypotheses of {\term}. \label{injection} \tacindex{injection} -The {\tt injection} tactic is based on the fact that constructors of -inductive sets are injections. That means that if $c$ is a constructor -of an inductive set, and if $(c~\vec{t_1})$ and $(c~\vec{t_2})$ are two -terms that are equal then $~\vec{t_1}$ and $~\vec{t_2}$ are equal -too. +The {\tt injection} tactic exploits the property that constructors of +inductive types are injective, i.e. that if $c$ is a constructor +of an inductive type and $c~\vec{t_1}$ and $c~\vec{t_2}$ are equal +then $\vec{t_1}$ and $\vec{t_2}$ are equal too. If {\term} is a proof of a statement of conclusion {\tt {\term$_1$} = {\term$_2$}}, -then {\tt injection} applies injectivity as deep as possible to -derive the equality of all the subterms of {\term$_1$} and {\term$_2$} -placed in the same positions. For example, from {\tt (S - (S n))=(S (S (S m)))} we may derive {\tt n=(S m)}. To use this -tactic {\term$_1$} and {\term$_2$} should be elements of an inductive -set and they should be neither explicitly equal, nor structurally -different. We mean by this that, if {\tt n$_1$} and {\tt n$_2$} are -their respective normal forms, then: -\begin{itemize} -\item {\tt n$_1$} and {\tt n$_2$} should not be syntactically equal, -\item there must not exist any pair of subterms {\tt u} and {\tt w}, - {\tt u} subterm of {\tt n$_1$} and {\tt w} subterm of {\tt n$_2$} , - placed in the same positions and having different constructors as - head symbols. -\end{itemize} -If these conditions are satisfied, then, the tactic derives the -equality of all the subterms of {\term$_1$} and {\term$_2$} placed in -the same positions and puts them as antecedents of the current goal. +then {\tt injection} applies the injectivity of constructors as deep as possible to +derive the equality of all the subterms of {\term$_1$} and {\term$_2$} at positions +where {\term$_1$} and {\term$_2$} start to differ. +For example, from {\tt (S p, S n) = (q, S (S m)} we may derive {\tt S + p = q} and {\tt n = S m}. For this tactic to work, {\term$_1$} and +{\term$_2$} should be typed with an inductive +type and they should be neither convertible, nor having a different +head constructor. If these conditions are satisfied, the tactic +derives the equality of all the subterms of {\term$_1$} and +{\term$_2$} at positions where they differ and adds them as +antecedents to the conclusion of the current goal. \Example Consider the following goal: @@ -2251,6 +2243,7 @@ context using \texttt{intros until \ident}. \item \errindex{Not a projectable equality but a discriminable one} \item \errindex{Nothing to do, it is an equality between convertible terms} \item \errindex{Not a primitive equality} +\item \errindex{Nothing to inject} \end{ErrMsgs} \begin{Variants} -- cgit v1.2.3 From 741f3fab052b91eaec57f32b639ca722c3d8dc34 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 12 Mar 2017 13:18:42 +0100 Subject: A fix for #5390 (a useful error on used introduction names was masked). With the "apply in" introduction pattern, and the backtrack possibly done in "apply" over the components of conjunctions (descend_in_conjunctions), the reasons for failing might have different sources. We give priority to the detection of used names over other (unification) errors so that an error there is not masked in the backtracking made by descend_in_conjunctions. Of course, the question of what best unification error to give in the presence of backtracking is still open. --- COMPATIBILITY | 7 +++++++ tactics/tactics.ml | 40 ++++++++++++++++++++++++++++++++++++---- test-suite/output/Tactics.out | 4 ++++ test-suite/output/Tactics.v | 10 ++++++++++ 4 files changed, 57 insertions(+), 4 deletions(-) diff --git a/COMPATIBILITY b/COMPATIBILITY index d423e71df..78dfabaa3 100644 --- a/COMPATIBILITY +++ b/COMPATIBILITY @@ -1,3 +1,10 @@ +Potential sources of incompatibilities between Coq V8.6 and V8.7 +---------------------------------------------------------------- + +- Extra superfluous names in introduction patterns may now raise an + error rather than a warning when the superfluous name is already in + use. The easy fix is to remove the superfluous name. + Potential sources of incompatibilities between Coq V8.5 and V8.6 ---------------------------------------------------------------- diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 556df6e55..15cef676e 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2395,6 +2395,29 @@ let rec explicit_intro_names = function explicit_intro_names l | [] -> [] +let rec check_name_unicity env ok seen = function +| (_, IntroForthcoming _) :: l -> check_name_unicity env ok seen l +| (loc, IntroNaming (IntroIdentifier id)) :: l -> + (try + ignore (if List.mem_f Id.equal id ok then raise Not_found else lookup_named id env); + user_err ~loc (pr_id id ++ str" is already used.") + with Not_found -> + if List.mem_f Id.equal id seen then + user_err ~loc (pr_id id ++ str" is used twice.") + else + check_name_unicity env ok (id::seen) l) +| (_, IntroAction (IntroOrAndPattern l)) :: l' -> + let ll = match l with IntroAndPattern l -> [l] | IntroOrPattern ll -> ll in + List.iter (fun l -> check_name_unicity env ok seen (l@l')) ll +| (_, IntroAction (IntroInjection l)) :: l' -> + check_name_unicity env ok seen (l@l') +| (_, IntroAction (IntroApplyOn (c,pat))) :: l' -> + check_name_unicity env ok seen (pat::l') +| (_, (IntroNaming (IntroAnonymous | IntroFresh _) + | IntroAction (IntroWildcard | IntroRewrite _))) :: l -> + check_name_unicity env ok seen l +| [] -> () + let wild_id = Id.of_string "_tmp" let rec list_mem_assoc_right id = function @@ -2530,13 +2553,21 @@ and prepare_intros_loc loc with_evars dft destopt = function | IntroForthcoming _ -> user_err ~loc (str "Introduction pattern for one hypothesis expected.") +let intro_patterns_head_core with_evars b destopt bound pat = + Proofview.Goal.enter { enter = begin fun gl -> + let env = Proofview.Goal.env gl in + check_name_unicity env [] [] pat; + intro_patterns_core with_evars b [] [] [] destopt + bound 0 (fun _ l -> clear_wildcards l) pat + end } + let intro_patterns_bound_to with_evars n destopt = - intro_patterns_core with_evars true [] [] [] destopt - (Some (true,n)) 0 (fun _ l -> clear_wildcards l) + intro_patterns_head_core with_evars true destopt + (Some (true,n)) let intro_patterns_to with_evars destopt = - intro_patterns_core with_evars (use_bracketing_last_or_and_intro_pattern ()) - [] [] [] destopt None 0 (fun _ l -> clear_wildcards l) + intro_patterns_head_core with_evars (use_bracketing_last_or_and_intro_pattern ()) + destopt None let intro_pattern_to with_evars destopt pat = intro_patterns_to with_evars destopt [dloc,pat] @@ -4191,6 +4222,7 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let f (_,is_not_let,_,_) = is_not_let in Array.map (fun (_,l) -> List.map f l) indsign in let names = compute_induction_names branchletsigns names in + Array.iter (check_name_unicity env toclear []) names; let tac = (if isrec then Tacticals.New.tclTHENFIRSTn else Tacticals.New.tclTHENLASTn) (Tacticals.New.tclTHENLIST [ diff --git a/test-suite/output/Tactics.out b/test-suite/output/Tactics.out index 239edd1da..19c9fc442 100644 --- a/test-suite/output/Tactics.out +++ b/test-suite/output/Tactics.out @@ -2,3 +2,7 @@ Ltac f H := split; [ a H | e H ] Ltac g := match goal with | |- context [ if ?X then _ else _ ] => case X end +The command has indeed failed with message: +H is already used. +The command has indeed failed with message: +H is already used. diff --git a/test-suite/output/Tactics.v b/test-suite/output/Tactics.v index a7c497cfa..9a5edb813 100644 --- a/test-suite/output/Tactics.v +++ b/test-suite/output/Tactics.v @@ -11,3 +11,13 @@ Print Ltac f. Ltac g := match goal with |- context [if ?X then _ else _ ] => case X end. Print Ltac g. + +(* Test an error message (#5390) *) +Lemma myid (P : Prop) : P <-> P. +Proof. split; auto. Qed. + +Goal True -> (True /\ True) -> True. +Proof. +intros H. +Fail intros [H%myid ?]. +Fail destruct 1 as [H%myid ?]. -- cgit v1.2.3 From 6a412da0f2681ede8f2753666bb9098e7ac8bfd2 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Thu, 18 May 2017 02:55:28 +0200 Subject: [ide] Disable `print_ast` call. So far this part of the system has shown little utility other than having developers put time to fix it every time they change something in the system. I have never seen this functionality used in the wild, and a large part of the vernac was marked TODO. Given that we have automatic methods to provide this functionality these days (PPX), we remove Texmacspp. --- ide/coqidetop.mllib | 1 - ide/ide_slave.ml | 10 +- ide/texmacspp.ml | 766 ---------------------------------------------------- ide/texmacspp.mli | 12 - 4 files changed, 2 insertions(+), 787 deletions(-) delete mode 100644 ide/texmacspp.ml delete mode 100644 ide/texmacspp.mli diff --git a/ide/coqidetop.mllib b/ide/coqidetop.mllib index 043ad6008..df988d8f1 100644 --- a/ide/coqidetop.mllib +++ b/ide/coqidetop.mllib @@ -4,6 +4,5 @@ Xml_printer Serialize Richpp Xmlprotocol -Texmacspp Document Ide_slave diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index 56ada9d13..dbca959ea 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -388,14 +388,8 @@ let interp ((_raw, verbose), s) = let quit = ref false -(** Serializes the output of Stm.get_ast *) -let print_ast id = - match Stm.get_ast id with - | Some (expr, loc) -> begin - try Texmacspp.tmpp expr loc - with e -> Xml_datatype.PCData ("ERROR " ^ Printexc.to_string e) - end - | None -> Xml_datatype.PCData "ERROR" +(** Disabled *) +let print_ast id = Xml_datatype.PCData "ERROR" (** Grouping all call handlers together + error handling *) diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml deleted file mode 100644 index 05f1820cf..000000000 --- a/ide/texmacspp.ml +++ /dev/null @@ -1,766 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* (List.filter (fun (a,_) -> a = attr) attrs) - | _ -> []) - xml_list in - match List.flatten attrs_list with - | [] -> (attr, "") - | l -> (List.hd l) - -let backstep_loc xmllist = - let start_att = get_fst_attr_in_xml_list "begin" xmllist in - let stop_att = get_fst_attr_in_xml_list "end" (List.rev xmllist) in - [start_att ; stop_att] - -let compare_begin_att xml1 xml2 = - let att1 = get_fst_attr_in_xml_list "begin" [xml1] in - let att2 = get_fst_attr_in_xml_list "begin" [xml2] in - match att1, att2 with - | (_, s1), (_, s2) when s1 == "" || s2 == "" -> 0 - | (_, s1), (_, s2) when int_of_string s1 > int_of_string s2 -> 1 - | (_, s1), (_, s2) when int_of_string s1 < int_of_string s2 -> -1 - | _ -> 0 - -let xmlBeginSection loc name = xmlWithLoc loc "beginsection" ["name", name] [] - -let xmlEndSegment loc name = xmlWithLoc loc "endsegment" ["name", name] [] - -let xmlThm typ name loc xml = - xmlWithLoc loc "theorem" ["type", typ; "name", name] xml - -let xmlDef typ name loc xml = - xmlWithLoc loc "definition" ["type", typ; "name", name] xml - -let xmlNotation attr name loc xml = - xmlWithLoc loc "notation" (("name", name) :: attr) xml - -let xmlReservedNotation attr name loc = - xmlWithLoc loc "reservednotation" (("name", name) :: attr) [] - -let xmlCst name ?(attr=[]) loc = - xmlWithLoc loc "constant" (("name", name) :: attr) [] - -let xmlOperator name ?(attr=[]) ?(pprules=[]) loc = - xmlWithLoc loc "operator" - (("name", name) :: List.map (fun (a,b) -> "format"^a,b) pprules @ attr) [] - -let xmlApply loc ?(attr=[]) xml = xmlWithLoc loc "apply" attr xml - -let xmlToken loc ?(attr=[]) xml = xmlWithLoc loc "token" attr xml - -let xmlTyped xml = Element("typed", (backstep_loc xml), xml) - -let xmlReturn ?(attr=[]) xml = Element("return", attr, xml) - -let xmlCase xml = Element("case", [], xml) - -let xmlScrutinee ?(attr=[]) xml = Element("scrutinee", attr, xml) - -let xmlWith xml = Element("with", [], xml) - -let xmlAssign id xml = Element("assign", ["target",string_of_id id], [xml]) - -let xmlInductive kind loc xml = xmlWithLoc loc "inductive" ["kind",kind] xml - -let xmlCoFixpoint xml = Element("cofixpoint", [], xml) - -let xmlFixpoint xml = Element("fixpoint", [], xml) - -let xmlCheck loc xml = xmlWithLoc loc "check" [] xml - -let xmlAssumption kind loc xml = xmlWithLoc loc "assumption" ["kind",kind] xml - -let xmlComment loc xml = xmlWithLoc loc "comment" [] xml - -let xmlCanonicalStructure attr loc = xmlWithLoc loc "canonicalstructure" attr [] - -let xmlQed ?(attr=[]) loc = xmlWithLoc loc "qed" attr [] - -let xmlPatvar id loc = xmlWithLoc loc "patvar" ["id", id] [] - -let xmlReference ref = - let name = Libnames.string_of_reference ref in - let i, j = Loc.unloc (Libnames.loc_of_reference ref) in - let b, e = string_of_int i, string_of_int j in - Element("reference",["name", name; "begin", b; "end", e] ,[]) - -let xmlRequire loc ?(attr=[]) xml = xmlWithLoc loc "require" attr xml -let xmlImport loc ?(attr=[]) xml = xmlWithLoc loc "import" attr xml - -let xmlAddLoadPath loc ?(attr=[]) xml = xmlWithLoc loc "addloadpath" attr xml -let xmlRemoveLoadPath loc ?(attr=[]) = xmlWithLoc loc "removeloadpath" attr -let xmlAddMLPath loc ?(attr=[]) = xmlWithLoc loc "addmlpath" attr - -let xmlExtend loc xml = xmlWithLoc loc "extend" [] xml - -let xmlScope loc action ?(attr=[]) name xml = - xmlWithLoc loc "scope" (["name",name;"action",action] @ attr) xml - -let xmlProofMode loc name = xmlWithLoc loc "proofmode" ["name",name] [] - -let xmlProof loc xml = xmlWithLoc loc "proof" [] xml - -let xmlSectionSubsetDescr name ssd = - Element("sectionsubsetdescr",["name",name], - [PCData (Proof_using.to_string ssd)]) - -let xmlDeclareMLModule loc s = - xmlWithLoc loc "declarexmlmodule" [] - (List.map (fun x -> Element("path",["value",x],[])) s) - -(* tactics *) -let xmlLtac loc xml = xmlWithLoc loc "ltac" [] xml - -(* toplevel commands *) -let xmlGallina loc xml = xmlWithLoc loc "gallina" [] xml - -let xmlTODO loc x = - xmlWithLoc loc "todo" [] [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] - -let string_of_name n = - match n with - | Anonymous -> "_" - | Name id -> Id.to_string id - -let string_of_glob_sort s = - match s with - | GProp -> "Prop" - | GSet -> "Set" - | GType _ -> "Type" - -let string_of_cast_sort c = - match c with - | CastConv _ -> "CastConv" - | CastVM _ -> "CastVM" - | CastNative _ -> "CastNative" - | CastCoerce -> "CastCoerce" - -let string_of_case_style s = - match s with - | LetStyle -> "Let" - | IfStyle -> "If" - | LetPatternStyle -> "LetPattern" - | MatchStyle -> "Match" - | RegularStyle -> "Regular" - -let attribute_of_syntax_modifier sm = -match sm with - | SetItemLevel (sl, NumLevel n) -> - List.map (fun s -> ("itemlevel", s)) sl @ ["level", string_of_int n] - | SetItemLevel (sl, NextLevel) -> - List.map (fun s -> ("itemlevel", s)) sl @ ["level", "next"] - | SetLevel i -> ["level", string_of_int i] - | SetAssoc a -> - begin match a with - | NonA -> ["",""] - | RightA -> ["associativity", "right"] - | LeftA -> ["associativity", "left"] - end - | SetEntryType (s, _) -> ["entrytype", s] - | SetOnlyPrinting -> ["onlyprinting", ""] - | SetOnlyParsing -> ["onlyparsing", ""] - | SetCompatVersion v -> ["compat", Flags.pr_version v] - | SetFormat (system, (loc, s)) -> - let start, stop = unlock loc in - ["format-"^system, s; "begin", start; "end", stop] - -let string_of_assumption_kind l a many = - match l, a, many with - | (Discharge, Logical, true) -> "Hypotheses" - | (Discharge, Logical, false) -> "Hypothesis" - | (Discharge, Definitional, true) -> "Variables" - | (Discharge, Definitional, false) -> "Variable" - | (Global, Logical, true) -> "Axioms" - | (Global, Logical, false) -> "Axiom" - | (Global, Definitional, true) -> "Parameters" - | (Global, Definitional, false) -> "Parameter" - | (Local, Logical, true) -> "Local Axioms" - | (Local, Logical, false) -> "Local Axiom" - | (Local, Definitional, true) -> "Local Parameters" - | (Local, Definitional, false) -> "Local Parameter" - | (Global, Conjectural, _) -> "Conjecture" - | ((Discharge | Local), Conjectural, _) -> assert false - -let rec pp_bindlist bl = - let tlist = - List.flatten - (List.map - (fun (loc_names, _, e) -> - let names = - (List.map - (fun (loc, name) -> - xmlCst (string_of_name name) loc) loc_names) in - match e with - | CHole _ -> names - | _ -> names @ [pp_expr e]) - bl) in - match tlist with - | [e] -> e - | l -> xmlTyped l -and pp_decl_notation ((_, s), ce, sc) = (* don't know what it is for now *) - Element ("decl_notation", ["name", s], [pp_expr ce]) -and pp_local_binder lb = (* don't know what it is for now *) - match lb with - | CLocalDef ((loc, nam), ce, ty) -> - let attrs = ["name", string_of_name nam] in - let value = match ty with Some t -> CCast (Loc.merge (constr_loc ce) (constr_loc t),ce, CastConv t) | None -> ce in - pp_expr ~attr:attrs value - | CLocalAssum (namll, _, ce) -> - let ppl = - List.map (fun (loc, nam) -> (xmlCst (string_of_name nam) loc)) namll in - xmlTyped (ppl @ [pp_expr ce]) - | CLocalPattern _ -> - assert false -and pp_local_decl_expr lde = (* don't know what it is for now *) - match lde with - | AssumExpr (_, ce) -> pp_expr ce - | DefExpr (_, ce, _) -> pp_expr ce -and pp_inductive_expr ((_, ((l, id),_)), lbl, ceo, _, cl_or_rdexpr) = - (* inductive_expr *) - let b,e = Loc.unloc l in - let location = ["begin", string_of_int b; "end", string_of_int e] in - [Element ("lident", ["name", Id.to_string id] @ location, [])] @ (* inductive name *) - begin match cl_or_rdexpr with - | Constructors coel -> List.map (fun (_, (_, ce)) -> pp_expr ce) coel - | RecordDecl (_, ldewwwl) -> - List.map (fun (((_, x), _), _) -> pp_local_decl_expr x) ldewwwl - end @ - begin match ceo with (* don't know what it is for now *) - | Some ce -> [pp_expr ce] - | None -> [] - end @ - (List.map pp_local_binder lbl) -and pp_recursion_order_expr optid roe = (* don't know what it is for now *) - let attrs = - match optid with - | None -> [] - | Some (loc, id) -> - let start, stop = unlock loc in - ["begin", start; "end", stop ; "name", Id.to_string id] in - let kind, expr = - match roe with - | CStructRec -> "struct", [] - | CWfRec e -> "rec", [pp_expr e] - | CMeasureRec (e, None) -> "mesrec", [pp_expr e] - | CMeasureRec (e, Some rel) -> "mesrec", [pp_expr e] @ [pp_expr rel] in - Element ("recursion_order", ["kind", kind] @ attrs, expr) -and pp_fixpoint_expr (((loc, id), pl), (optid, roe), lbl, ce, ceo) = - (* fixpoint_expr *) - let start, stop = unlock loc in - let id = Id.to_string id in - [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ - (* fixpoint name *) - [pp_recursion_order_expr optid roe] @ - (List.map pp_local_binder lbl) @ - [pp_expr ce] @ - begin match ceo with (* don't know what it is for now *) - | Some ce -> [pp_expr ce] - | None -> [] - end -and pp_cofixpoint_expr (((loc, id), pl), lbl, ce, ceo) = (* cofixpoint_expr *) - (* Nota: it is like fixpoint_expr without (optid, roe) - * so could be merged if there is no more differences *) - let start, stop = unlock loc in - let id = Id.to_string id in - [Element ("lident", ["begin", start; "end", stop ; "name", id], [])] @ - (* cofixpoint name *) - (List.map pp_local_binder lbl) @ - [pp_expr ce] @ - begin match ceo with (* don't know what it is for now *) - | Some ce -> [pp_expr ce] - | None -> [] - end -and pp_lident (loc, id) = xmlCst (Id.to_string id) loc -and pp_simple_binder (idl, ce) = List.map pp_lident idl @ [pp_expr ce] -and pp_cases_pattern_expr cpe = - match cpe with - | CPatAlias (loc, cpe, id) -> - xmlApply loc - (xmlOperator "alias" ~attr:["name", string_of_id id] loc :: - [pp_cases_pattern_expr cpe]) - | CPatCstr (loc, ref, None, cpel2) -> - xmlApply loc - (xmlOperator "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: - [Element ("impargs", [], []); - Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) - | CPatCstr (loc, ref, Some cpel1, cpel2) -> - xmlApply loc - (xmlOperator "reference" - ~attr:["name", Libnames.string_of_reference ref] loc :: - [Element ("impargs", [], (List.map pp_cases_pattern_expr cpel1)); - Element ("args", [], (List.map pp_cases_pattern_expr cpel2))]) - | CPatAtom (loc, optr) -> - let attrs = match optr with - | None -> [] - | Some r -> ["name", Libnames.string_of_reference r] in - xmlApply loc (xmlOperator "atom" ~attr:attrs loc :: []) - | CPatOr (loc, cpel) -> - xmlApply loc (xmlOperator "or" loc :: List.map pp_cases_pattern_expr cpel) - | CPatNotation (loc, n, (subst_constr, subst_rec), cpel) -> - xmlApply loc - (xmlOperator "notation" loc :: - [xmlOperator n loc; - Element ("subst", [], - [Element ("subterms", [], - List.map pp_cases_pattern_expr subst_constr); - Element ("recsubterms", [], - List.map - (fun (cpel) -> - Element ("recsubterm", [], - List.map pp_cases_pattern_expr cpel)) - subst_rec)]); - Element ("args", [], (List.map pp_cases_pattern_expr cpel))]) - | CPatPrim (loc, tok) -> pp_token loc tok - | CPatRecord (loc, rcl) -> - xmlApply loc - (xmlOperator "record" loc :: - List.map (fun (r, cpe) -> - Element ("field", - ["reference", Libnames.string_of_reference r], - [pp_cases_pattern_expr cpe])) - rcl) - | CPatDelimiters (loc, delim, cpe) -> - xmlApply loc - (xmlOperator "delimiter" ~attr:["name", delim] loc :: - [pp_cases_pattern_expr cpe]) - | CPatCast _ -> assert false -and pp_case_expr (e, name, pat) = - match name, pat with - | None, None -> xmlScrutinee [pp_expr e] - | Some (loc, name), None -> - let start, stop= unlock loc in - xmlScrutinee ~attr:["name", string_of_name name; - "begin", start; "end", stop] [pp_expr e] - | Some (loc, name), Some p -> - let start, stop= unlock loc in - xmlScrutinee ~attr:["name", string_of_name name; - "begin", start; "end", stop] - [Element ("in", [], [pp_cases_pattern_expr p]) ; pp_expr e] - | None, Some p -> - xmlScrutinee [Element ("in", [], [pp_cases_pattern_expr p]) ; pp_expr e] -and pp_branch_expr_list bel = - xmlWith - (List.map - (fun (_, cpel, e) -> - let ppcepl = - List.map pp_cases_pattern_expr (List.flatten (List.map snd cpel)) in - let ppe = [pp_expr e] in - xmlCase (ppcepl @ ppe)) - bel) -and pp_token loc tok = - let tokstr = - match tok with - | String s -> PCData s - | Numeral n -> PCData (to_string n) in - xmlToken loc [tokstr] -and pp_local_binder_list lbl = - let l = (List.map pp_local_binder lbl) in - Element ("recurse", (backstep_loc l), l) -and pp_const_expr_list cel = - let l = List.map pp_expr cel in - Element ("recurse", (backstep_loc l), l) -and pp_expr ?(attr=[]) e = - match e with - | CRef (r, _) -> - xmlCst ~attr - (Libnames.string_of_reference r) (Libnames.loc_of_reference r) - | CProdN (loc, bl, e) -> - xmlApply loc - (xmlOperator "forall" loc :: [pp_bindlist bl] @ [pp_expr e]) - | CApp (loc, (_, hd), args) -> - xmlApply ~attr loc (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args) - | CAppExpl (loc, (_, r, _), args) -> - xmlApply ~attr loc - (xmlCst (Libnames.string_of_reference r) - (Libnames.loc_of_reference r) :: List.map pp_expr args) - | CNotation (loc, notation, ([],[],[])) -> - xmlOperator notation loc - | CNotation (loc, notation, (args, cell, lbll)) -> - let fmts = Notation.find_notation_extra_printing_rules notation in - let oper = xmlOperator notation loc ~pprules:fmts in - let cels = List.map pp_const_expr_list cell in - let lbls = List.map pp_local_binder_list lbll in - let args = List.map pp_expr args in - xmlApply loc (oper :: (List.sort compare_begin_att (args @ cels @ lbls))) - | CSort(loc, s) -> - xmlOperator (string_of_glob_sort s) loc - | CDelimiters (loc, scope, ce) -> - xmlApply loc (xmlOperator "delimiter" ~attr:["name", scope] loc :: - [pp_expr ce]) - | CPrim (loc, tok) -> pp_token loc tok - | CGeneralization (loc, kind, _, e) -> - let kind= match kind with - | Explicit -> "explicit" - | Implicit -> "implicit" in - xmlApply loc - (xmlOperator "generalization" ~attr:["kind", kind] loc :: [pp_expr e]) - | CCast (loc, e, tc) -> - begin match tc with - | CastConv t | CastVM t |CastNative t -> - xmlApply loc - (xmlOperator ":" loc ~attr:["kind", (string_of_cast_sort tc)] :: - [pp_expr e; pp_expr t]) - | CastCoerce -> - xmlApply loc - (xmlOperator ":" loc ~attr:["kind", "CastCoerce"] :: - [pp_expr e]) - end - | CEvar (loc, ek, cel) -> - let ppcel = List.map (fun (id,e) -> xmlAssign id (pp_expr e)) cel in - xmlApply loc - (xmlOperator "evar" loc ~attr:["id", string_of_id ek] :: - ppcel) - | CPatVar (loc, id) -> xmlPatvar (string_of_id id) loc - | CHole (loc, _, _, _) -> xmlCst ~attr "_" loc - | CIf (loc, test, (_, ret), th, el) -> - let return = match ret with - | None -> [] - | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "if" loc :: - return @ [pp_expr th] @ [pp_expr el]) - | CLetTuple (loc, names, (_, ret), value, body) -> - let return = match ret with - | None -> [] - | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "lettuple" loc :: - return @ - (List.map (fun (loc, var) -> xmlCst (string_of_name var) loc) names) @ - [pp_expr value; pp_expr body]) - | CCases (loc, sty, ret, cel, bel) -> - let return = match ret with - | None -> [] - | Some r -> [xmlReturn [pp_expr r]] in - xmlApply loc - (xmlOperator "match" loc ~attr:["style", (string_of_case_style sty)] :: - (return @ - [Element ("scrutinees", [], List.map pp_case_expr cel)] @ - [pp_branch_expr_list bel])) - | CRecord (_, _) -> assert false - | CLetIn (loc, (varloc, var), value, typ, body) -> - let value = match typ with Some t -> CCast (Loc.merge (constr_loc value) (constr_loc t),value, CastConv t) | None -> value in - xmlApply loc - (xmlOperator "let" loc :: - [xmlCst (string_of_name var) varloc; pp_expr value; pp_expr body]) - | CLambdaN (loc, bl, e) -> - xmlApply loc - (xmlOperator "lambda" loc :: [pp_bindlist bl] @ [pp_expr e]) - | CCoFix (_, _, _) -> assert false - | CFix (loc, lid, fel) -> - xmlApply loc - (xmlOperator "fix" loc :: - List.flatten (List.map - (fun (a,b,cl,c,d) -> pp_fixpoint_expr ((a,None),b,cl,c,Some d)) - fel)) - -let pp_comment (c) = - match c with - | CommentConstr e -> [pp_expr e] - | CommentString s -> [Element ("string", [], [PCData s])] - | CommentInt i -> [PCData (string_of_int i)] - -let rec tmpp v loc = - match v with - (* Control *) - | VernacLoad (verbose,f) -> - xmlWithLoc loc "load" ["verbose",string_of_bool verbose;"file",f] [] - | VernacTime (loc,e) -> - xmlApply loc (Element("time",[],[]) :: - [tmpp e loc]) - | VernacRedirect (s, (loc,e)) -> - xmlApply loc (Element("redirect",["path", s],[]) :: - [tmpp e loc]) - | VernacTimeout (s,e) -> - xmlApply loc (Element("timeout",["val",string_of_int s],[]) :: - [tmpp e loc]) - | VernacFail e -> xmlApply loc (Element("fail",[],[]) :: [tmpp e loc]) - - (* Syntax *) - | VernacSyntaxExtension (_, ((_, name), sml)) -> - let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in - xmlReservedNotation attrs name loc - - | VernacOpenCloseScope (_,(true,name)) -> xmlScope loc "open" name [] - | VernacOpenCloseScope (_,(false,name)) -> xmlScope loc "close" name [] - | VernacDelimiters (name,Some tag) -> - xmlScope loc "delimit" name ~attr:["delimiter",tag] [] - | VernacDelimiters (name,None) -> - xmlScope loc "undelimit" name ~attr:[] [] - | VernacInfix (_,((_,name),sml),ce,sn) -> - let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in - let sc_attr = - match sn with - | Some scope -> ["scope", scope] - | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] - | VernacNotation (_, ce, (lstr, sml), sn) -> - let name = snd lstr in - let attrs = List.flatten (List.map attribute_of_syntax_modifier sml) in - let sc_attr = - match sn with - | Some scope -> ["scope", scope] - | None -> [] in - xmlNotation (sc_attr @ attrs) name loc [pp_expr ce] - | VernacBindScope _ as x -> xmlTODO loc x - | VernacNotationAddFormat _ as x -> xmlTODO loc x - | VernacUniverse _ - | VernacConstraint _ - | VernacPolymorphic (_, _) as x -> xmlTODO loc x - (* Gallina *) - | VernacDefinition (ldk, ((_,id),_), de) -> - let l, dk = - match ldk with - | Some l, dk -> (l, dk) - | None, dk -> (Global, dk) in (* Like in ppvernac.ml, l 585 *) - let e = - match de with - | ProveBody (_, ce) -> ce - | DefineBody (_, Some _, ce, None) -> ce - | DefineBody (_, None , ce, None) -> ce - | DefineBody (_, Some _, ce, Some _) -> ce - | DefineBody (_, None , ce, Some _) -> ce in - let str_dk = Kindops.string_of_definition_kind (l, false, dk) in - let str_id = Id.to_string id in - (xmlDef str_dk str_id loc [pp_expr e]) - | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement) ], b) -> - let str_tk = Kindops.string_of_theorem_kind tk in - let str_id = Id.to_string id in - (xmlThm str_tk str_id loc [pp_expr statement]) - | VernacStartTheoremProof _ as x -> xmlTODO loc x - | VernacEndProof pe -> - begin - match pe with - | Admitted -> xmlQed loc - | Proved (_, Some ((_, id), Some tk)) -> - let nam = Id.to_string id in - let typ = Kindops.string_of_theorem_kind tk in - xmlQed ~attr:["name", nam; "type", typ] loc - | Proved (_, Some ((_, id), None)) -> - let nam = Id.to_string id in - xmlQed ~attr:["name", nam] loc - | Proved _ -> xmlQed loc - end - | VernacExactProof _ as x -> xmlTODO loc x - | VernacAssumption ((l, a), _, sbwcl) -> - let binders = List.map (fun (_, (id, c)) -> (List.map fst id, c)) sbwcl in - let many = - List.length (List.flatten (List.map fst binders)) > 1 in - let exprs = - List.flatten (List.map pp_simple_binder binders) in - let l = match l with Some x -> x | None -> Decl_kinds.Global in - let kind = string_of_assumption_kind l a many in - xmlAssumption kind loc exprs - | VernacInductive (_, _, iednll) -> - let kind = - let (_, _, _, k, _),_ = List.hd iednll in - begin - match k with - | Record -> "Record" - | Structure -> "Structure" - | Inductive_kw -> "Inductive" - | CoInductive -> "CoInductive" - | Class _ -> "Class" - | Variant -> "Variant" - end in - let exprs = - List.flatten (* should probably not be flattened *) - (List.map - (fun (ie, dnl) -> (pp_inductive_expr ie) @ - (List.map pp_decl_notation dnl)) iednll) in - xmlInductive kind loc exprs - | VernacFixpoint (_, fednll) -> - let exprs = - List.flatten (* should probably not be flattened *) - (List.map - (fun (fe, dnl) -> (pp_fixpoint_expr fe) @ - (List.map pp_decl_notation dnl)) fednll) in - xmlFixpoint exprs - | VernacCoFixpoint (_, cfednll) -> - (* Nota: it is like VernacFixpoint without so could be merged *) - let exprs = - List.flatten (* should probably not be flattened *) - (List.map - (fun (cfe, dnl) -> (pp_cofixpoint_expr cfe) @ - (List.map pp_decl_notation dnl)) cfednll) in - xmlCoFixpoint exprs - | VernacScheme _ as x -> xmlTODO loc x - | VernacCombinedScheme _ as x -> xmlTODO loc x - - (* Gallina extensions *) - | VernacBeginSection (_, id) -> xmlBeginSection loc (Id.to_string id) - | VernacEndSegment (_, id) -> xmlEndSegment loc (Id.to_string id) - | VernacNameSectionHypSet _ as x -> xmlTODO loc x - | VernacRequire (from, import, l) -> - let import = match import with - | None -> [] - | Some true -> ["export","true"] - | Some false -> ["import","true"] - in - let from = match from with - | None -> [] - | Some r -> ["from", Libnames.string_of_reference r] - in - xmlRequire loc ~attr:(from @ import) (List.map (fun ref -> - xmlReference ref) l) - | VernacImport (true,l) -> - xmlImport loc ~attr:["export","true"] (List.map (fun ref -> - xmlReference ref) l) - | VernacImport (false,l) -> - xmlImport loc (List.map (fun ref -> - xmlReference ref) l) - | VernacCanonical r -> - let attr = - match r with - | AN (Qualid (_, q)) -> ["qualid", string_of_qualid q] - | AN (Ident (_, id)) -> ["id", Id.to_string id] - | ByNotation (_, s, _) -> ["notation", s] in - xmlCanonicalStructure attr loc - | VernacCoercion _ as x -> xmlTODO loc x - | VernacIdentityCoercion _ as x -> xmlTODO loc x - - (* Type classes *) - | VernacInstance _ as x -> xmlTODO loc x - - | VernacContext _ as x -> xmlTODO loc x - - | VernacDeclareInstances _ as x -> xmlTODO loc x - - | VernacDeclareClass _ as x -> xmlTODO loc x - - (* Modules and Module Types *) - | VernacDeclareModule _ as x -> xmlTODO loc x - | VernacDefineModule _ as x -> xmlTODO loc x - | VernacDeclareModuleType _ as x -> xmlTODO loc x - | VernacInclude _ as x -> xmlTODO loc x - - (* Solving *) - - | (VernacSolveExistential _) as x -> - xmlLtac loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] - - (* Auxiliary file and library management *) - | VernacAddLoadPath (recf,name,None) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] [] - | VernacAddLoadPath (recf,name,Some dp) -> - xmlAddLoadPath loc ~attr:["rec",string_of_bool recf;"path",name] - [PCData (Names.DirPath.to_string dp)] - | VernacRemoveLoadPath name -> xmlRemoveLoadPath loc ~attr:["path",name] [] - | VernacAddMLPath (recf,name) -> - xmlAddMLPath loc ~attr:["rec",string_of_bool recf;"path",name] [] - | VernacDeclareMLModule sl -> xmlDeclareMLModule loc sl - | VernacChdir _ as x -> xmlTODO loc x - - (* State management *) - | VernacWriteState _ as x -> xmlTODO loc x - | VernacRestoreState _ as x -> xmlTODO loc x - - (* Resetting *) - | VernacResetName _ as x -> xmlTODO loc x - | VernacResetInitial as x -> xmlTODO loc x - | VernacBack _ as x -> xmlTODO loc x - | VernacBackTo _ -> PCData "VernacBackTo" - - (* Commands *) - | VernacCreateHintDb _ as x -> xmlTODO loc x - | VernacRemoveHints _ as x -> xmlTODO loc x - | VernacHints _ as x -> xmlTODO loc x - | VernacSyntacticDefinition ((_, name), (idl, ce), _, _) -> - let name = Id.to_string name in - let attrs = List.map (fun id -> ("id", Id.to_string id)) idl in - xmlNotation attrs name loc [pp_expr ce] - | VernacDeclareImplicits _ as x -> xmlTODO loc x - | VernacArguments _ as x -> xmlTODO loc x - | VernacArgumentsScope _ as x -> xmlTODO loc x - | VernacReserve _ as x -> xmlTODO loc x - | VernacGeneralizable _ as x -> xmlTODO loc x - | VernacSetOpacity _ as x -> xmlTODO loc x - | VernacSetStrategy _ as x -> xmlTODO loc x - | VernacUnsetOption _ as x -> xmlTODO loc x - | VernacSetOption _ as x -> xmlTODO loc x - | VernacSetAppendOption _ as x -> xmlTODO loc x - | VernacAddOption _ as x -> xmlTODO loc x - | VernacRemoveOption _ as x -> xmlTODO loc x - | VernacMemOption _ as x -> xmlTODO loc x - | VernacPrintOption _ as x -> xmlTODO loc x - | VernacCheckMayEval (_,_,e) -> xmlCheck loc [pp_expr e] - | VernacGlobalCheck _ as x -> xmlTODO loc x - | VernacDeclareReduction _ as x -> xmlTODO loc x - | VernacPrint _ as x -> xmlTODO loc x - | VernacSearch _ as x -> xmlTODO loc x - | VernacLocate _ as x -> xmlTODO loc x - | VernacRegister _ as x -> xmlTODO loc x - | VernacComments (cl) -> - xmlComment loc (List.flatten (List.map pp_comment cl)) - - (* Stm backdoor *) - | VernacStm _ as x -> xmlTODO loc x - - (* Proof management *) - | VernacGoal _ as x -> xmlTODO loc x - | VernacAbort _ as x -> xmlTODO loc x - | VernacAbortAll -> PCData "VernacAbortAll" - | VernacRestart as x -> xmlTODO loc x - | VernacUndo _ as x -> xmlTODO loc x - | VernacUndoTo _ as x -> xmlTODO loc x - | VernacBacktrack _ as x -> xmlTODO loc x - | VernacFocus _ as x -> xmlTODO loc x - | VernacUnfocus as x -> xmlTODO loc x - | VernacUnfocused as x -> xmlTODO loc x - | VernacBullet _ as x -> xmlTODO loc x - | VernacSubproof _ as x -> xmlTODO loc x - | VernacEndSubproof as x -> xmlTODO loc x - | VernacShow _ as x -> xmlTODO loc x - | VernacCheckGuard as x -> xmlTODO loc x - | VernacProof (tac,using) -> - let tac = None (** FIXME *) in - let using = Option.map (xmlSectionSubsetDescr "using") using in - xmlProof loc (Option.List.(cons tac (cons using []))) - | VernacProofMode name -> xmlProofMode loc name - - (* Toplevel control *) - | VernacToplevelControl _ as x -> xmlTODO loc x - - (* For extension *) - | VernacExtend _ as x -> - xmlExtend loc [PCData (Pp.string_of_ppcmds (Ppvernac.pr_vernac x))] - - (* Flags *) - | VernacProgram e -> xmlApply loc (Element("program",[],[]) :: [tmpp e loc]) - | VernacLocal (b,e) -> - xmlApply loc (Element("local",["flag",string_of_bool b],[]) :: - [tmpp e loc]) - -let tmpp v loc = - match tmpp v loc with - | Element("ltac",_,_) as x -> x - | xml -> xmlGallina loc [xml] diff --git a/ide/texmacspp.mli b/ide/texmacspp.mli deleted file mode 100644 index 858847fb6..000000000 --- a/ide/texmacspp.mli +++ /dev/null @@ -1,12 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* Loc.t -> xml -- cgit v1.2.3 From 6ef8a007e77c61d9f8c2b19b19dc22548f1896ca Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Thu, 18 May 2017 04:47:38 +0200 Subject: [toplevel] [stm] Avoid edit_at in batch mode (bug #5520) In non-interactive mode, `edit_at` seems to do very weird things, for instance will try to recompute all the previous states which seems weird. We better avoid that to approximate 8.6 behavior while we investigate more. --- toplevel/coqtop.ml | 3 +++ toplevel/vernac.ml | 4 +++- 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index 8f50bfb3d..41d370ea5 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -641,6 +641,9 @@ let init_toplevel arglist = init_library_roots (); load_vernac_obj (); require (); + (* XXX: This is incorrect in batch mode, as we will initialize + the STM before having done Declaremods.start_library, thus + state 1 is invalid. This bug was present in 8.5/8.6. *) Stm.init (); let sid = load_rcfile (Stm.get_current_state ()) in (* XXX: We ignore this for now, but should be threaded to the toplevels *) diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index eaf685b18..f8378f876 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -153,7 +153,9 @@ let rec interp_vernac sid (loc,com) = let com = if !Flags.time then VernacTime (loc,com) else com in interp com with reraise -> - ignore(Stm.edit_at sid); + (* XXX: In non-interactive mode edit_at seems to do very weird + things, so we better avoid it while we investigate *) + if not !Flags.batch_mode then ignore(Stm.edit_at sid); let (reraise, info) = CErrors.push reraise in let loc' = Option.default Loc.ghost (Loc.get_loc info) in if Loc.is_ghost loc' then iraise (reraise, Loc.add_loc info loc) -- cgit v1.2.3 From 79595b45e68c6234aff54f076687a8024c4a2aaa Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 18 May 2017 12:50:36 -0400 Subject: Add .dir-locals.el to plugins As requested in https://github.com/coq/coq/pull/386#issuecomment-302358542 --- plugins/.dir-locals.el | 4 ++++ 1 file changed, 4 insertions(+) create mode 100644 plugins/.dir-locals.el diff --git a/plugins/.dir-locals.el b/plugins/.dir-locals.el new file mode 100644 index 000000000..4e8830f6c --- /dev/null +++ b/plugins/.dir-locals.el @@ -0,0 +1,4 @@ +((coq-mode . ((eval . (let ((default-directory (locate-dominating-file + buffer-file-name ".dir-locals.el"))) + (setq-local coq-prog-args `("-coqlib" ,(expand-file-name "..") "-R" ,(expand-file-name ".") "Coq")) + (setq-local coq-prog-name (expand-file-name "../bin/coqtop"))))))) -- cgit v1.2.3 From 2b00427c545be77805c6e954c2fa9c45682fec85 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Thu, 18 May 2017 14:07:04 -0400 Subject: Fix a typo --- dev/build/windows/makecoq_mingw.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/dev/build/windows/makecoq_mingw.sh b/dev/build/windows/makecoq_mingw.sh index 52b158871..98e80c765 100644 --- a/dev/build/windows/makecoq_mingw.sh +++ b/dev/build/windows/makecoq_mingw.sh @@ -1082,7 +1082,7 @@ function make_coq { copq_coq_gtk copy_coq_license - # make clean seems to br broken for 8.5pl2 + # make clean seems to be broken for 8.5pl2 # 1.) find | xargs fails on cygwin, can be fixed by sed -i 's|\| xargs rm -f|-exec rm -fv \{\} \+|' Makefile # 2.) clean of test suites fails with "cannot run complexity tests (no bogomips found)" # make clean -- cgit v1.2.3 From 5c68b57c37c23a3b7b2afe4f6ff073568c7b8db9 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 19 May 2017 12:20:19 +0200 Subject: Fixing an extra bug with pattern_of_constr. Ensure in type constr_pattern that those preexisting existential variables of the goal which do not contribute as pattern variables are expanded: constr_pattern is not observed up to evar expansion (like EConstr does), so we need to pre-normalize defined evars in patterns to that matching against an EConstr works. --- pretyping/patternops.ml | 5 +++-- test-suite/success/change.v | 9 +++++++++ 2 files changed, 12 insertions(+), 2 deletions(-) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index a22db1407..638e12d7c 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -155,14 +155,15 @@ let pattern_of_constr env sigma t = | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) - | Evar (evk,ctxt) -> + | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> assert (not b); PMeta (Some id) | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ -> (* These are the two evar kinds used for existing goals *) (* see Proofview.mark_in_evm *) - PEvar (evk,Array.map (pattern_of_constr env) ctxt) + if Evd.is_defined sigma evk then pattern_of_constr env (Evd.existential_value sigma ev) + else PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> PMeta None) | Case (ci,p,a,br) -> diff --git a/test-suite/success/change.v b/test-suite/success/change.v index 1f0b7d38a..a9821b027 100644 --- a/test-suite/success/change.v +++ b/test-suite/success/change.v @@ -59,3 +59,12 @@ unfold x. (* check that n in 0+n is not interpreted as the n from "fun n" *) change n with (0+n). Abort. + +(* Check non-collision of non-normalized defined evars with pattern variables *) + +Goal exists x, 1=1 -> x=1/\x=1. +eexists ?[n]; intros; split. +eassumption. +match goal with |- ?x=1 => change (x=1) with (0+x=1) end. +match goal with |- 0+1=1 => trivial end. +Qed. -- cgit v1.2.3 From fa1e921257735d246d04be9e456f11ec7330b37a Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 19 May 2017 13:17:08 +0200 Subject: Re-adding explicit dependency of misc universe test into all_stdlib.v. Was working when calling test-suite from main Makefile but not when calling directly from the test-suite Makefile. The dependency was mistakenly dropped in 98a927aef. --- test-suite/Makefile | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/test-suite/Makefile b/test-suite/Makefile index 39ef05269..afaa48463 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -408,6 +408,11 @@ modules/%.vo: modules/%.v misc: $(patsubst %.sh,%.log,$(wildcard misc/*.sh)) +misc/universes.log: misc/universes/all_stdlib.v + +misc/universes/all_stdlib.v: + cd .. && $(MAKE) test-suite/$@ + $(patsubst %.sh,%.log,$(wildcard misc/*.sh)): %.log: %.sh $(PREREQUISITELOG) @echo "TEST $<" $(HIDE){ \ -- cgit v1.2.3 From 931cb1d5b337b0fda54133954c2e3025e1637beb Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 19 May 2017 19:36:22 +0200 Subject: Travis: do not cache opam logs (+prettier spacing) --- .travis.yml | 15 +++++++++++++-- 1 file changed, 13 insertions(+), 2 deletions(-) diff --git a/.travis.yml b/.travis.yml index aa2e0e63c..a959fbf96 100644 --- a/.travis.yml +++ b/.travis.yml @@ -1,14 +1,21 @@ dist: trusty + # Travis builds are slower using sudo: false (the container-based # infrastructure) as of March 2017; see # https://github.com/coq/coq/pull/467 for some discussion. sudo: required + # Until Ocaml becomes a language, we set a known one. language: c + cache: apt: true directories: - $HOME/.opam + +before_cache: + - rm -rf ~/.opam/log/ + addons: apt: sources: @@ -17,6 +24,7 @@ addons: - opam - aspcud - gcc-multilib + env: global: - NJOBS=2 @@ -55,9 +63,8 @@ matrix: allow_failures: - env: TEST_TARGET="ci-geocoq" - # Full Coq test-suite with two compilers - # [TODO: use yaml refs and avoid duplication for packages list] include: + # Full Coq test-suite with two compilers - env: - TEST_TARGET="test-suite" - EXTRA_CONF="-coqide opt -with-doc yes" @@ -81,6 +88,7 @@ matrix: - ghostscript - transfig - imagemagick + - env: - TEST_TARGET="test-suite" - COMPILER="4.04.1" @@ -92,6 +100,8 @@ matrix: sources: - avsm packages: *extra-packages + + # Ocaml warnings with two compilers - env: - TEST_TARGET="coqocaml" - EXTRA_CONF="-coqide opt -warn-error" @@ -107,6 +117,7 @@ matrix: - aspcud - libgtk2.0-dev - libgtksourceview2.0-dev + - env: - TEST_TARGET="coqocaml" - COMPILER="4.04.1" -- cgit v1.2.3 From ac333e6297bd7c32282fa5f47bebeb35826072df Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Sat, 20 May 2017 01:24:46 +0200 Subject: [test-suite] Add tests for goal printing. - https://coq.inria.fr/bugs/show_bug.cgi?id=5529 - https://coq.inria.fr/bugs/show_bug.cgi?id=5537 See also PR #640 --- test-suite/output/goal_output.out | 8 ++++++++ test-suite/output/goal_output.v | 13 +++++++++++++ 2 files changed, 21 insertions(+) create mode 100644 test-suite/output/goal_output.out create mode 100644 test-suite/output/goal_output.v diff --git a/test-suite/output/goal_output.out b/test-suite/output/goal_output.out new file mode 100644 index 000000000..773533a8d --- /dev/null +++ b/test-suite/output/goal_output.out @@ -0,0 +1,8 @@ +Nat.t = nat + : Set +Nat.t = nat + : Set +1 subgoal + + ============================ + False diff --git a/test-suite/output/goal_output.v b/test-suite/output/goal_output.v new file mode 100644 index 000000000..327b80b0a --- /dev/null +++ b/test-suite/output/goal_output.v @@ -0,0 +1,13 @@ +(* From + - https://coq.inria.fr/bugs/show_bug.cgi?id=5529 + - https://coq.inria.fr/bugs/show_bug.cgi?id=5537 + *) + +Print Nat.t. +Timeout 1 Print Nat.t. + +Lemma toto: False. +Set Printing All. +Show. +Abort. + -- cgit v1.2.3 From aae960fc0ed0a3a8d244c1c3d7c5363c17044cbc Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Sat, 20 May 2017 10:28:17 +0200 Subject: Mention ./configure in INSTALL.doc As prompted in https://coq.inria.fr/bugs/show_bug.cgi?id=2831 --- INSTALL.doc | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/INSTALL.doc b/INSTALL.doc index 2472d2b2a..21b21163c 100644 --- a/INSTALL.doc +++ b/INSTALL.doc @@ -41,6 +41,7 @@ Compilation To produce all documentation about Coq, just run: + ./configure (if you hadn't already) make doc @@ -71,6 +72,10 @@ Alternatively, you can use some specific targets: to produce all formats of the Coq standard library +Also note the "-with-doc yes" option of ./configure to enable the +build of the documentation as part of the default make target. + + Installation ------------ -- cgit v1.2.3 From 5b218f87bd59cfe9d510410c9acf78b5485391e1 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Fri, 5 May 2017 15:01:15 +0200 Subject: Revised behavior on ill-formed identifiers. Namely: Replacing (currently deactivated) warning on illegal ident by an error in strict mode and nothing in soft mode. --- kernel/names.ml | 18 +++++------------- kernel/names.mli | 5 +++-- 2 files changed, 8 insertions(+), 15 deletions(-) diff --git a/kernel/names.ml b/kernel/names.ml index 811b4a62a..f5b3f4e00 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -34,16 +34,8 @@ struct let hash = String.hash - let warn_invalid_identifier = - CWarnings.create ~name:"invalid-identifier" ~category:"parsing" - ~default:CWarnings.Disabled - (fun s -> str s) - - let check_soft ?(warn = true) x = - let iter (fatal, x) = - if fatal then CErrors.error x else - if warn then warn_invalid_identifier x - in + let check_valid ?(strict=true) x = + let iter (fatal, x) = if fatal || strict then CErrors.error x in Option.iter iter (Unicode.ident_refutation x) let is_valid s = match Unicode.ident_refutation s with @@ -52,15 +44,15 @@ struct let of_bytes s = let s = Bytes.to_string s in - check_soft s; + check_valid s; String.hcons s let of_string s = - let () = check_soft s in + let () = check_valid s in String.hcons s let of_string_soft s = - let () = check_soft ~warn:false s in + let () = check_valid ~strict:false s in String.hcons s let to_string id = id diff --git a/kernel/names.mli b/kernel/names.mli index be9b9422b..5b0163aa5 100644 --- a/kernel/names.mli +++ b/kernel/names.mli @@ -46,11 +46,12 @@ sig val of_bytes : bytes -> t val of_string : string -> t (** Converts a string into an identifier. - @raise UserError if the string is not valid, or echo a warning if it contains invalid identifier characters. + @raise UserError if the string is invalid as an identifier. @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *) val of_string_soft : string -> t - (** Same as {!of_string} except that no warning is ever issued. + (** Same as {!of_string} except that any string made of supported UTF-8 characters is accepted. + @raise UserError if the string is invalid as an UTF-8 string. @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *) val to_string : t -> string -- cgit v1.2.3 From c1bf3bc4956d0b8af879fcf4854cbd650e1821c0 Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Sat, 20 May 2017 14:06:06 +0200 Subject: Change wrong bullet message. Remove a space before colon. Remove the use of term mandatory (this closes https://coq.inria.fr/bugs/show_bug.cgi?id=3994). --- proofs/proof_global.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 5f4a7766f..99fab0848 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -516,7 +516,7 @@ module Bullet = struct | NeedClosingBrace -> str"Try unfocusing with \"}\"." | NoBulletInUse -> assert false (* This should never raise an error. *) | ProofFinished -> str"No more subgoals." - | Suggest b -> str"Bullet " ++ pr_bullet b ++ str" is mandatory here." + | Suggest b -> str"Expecting " ++ pr_bullet b ++ str"." | Unfinished b -> str"Current bullet " ++ pr_bullet b ++ str" is not finished." exception FailedBullet of t * suggestion @@ -525,7 +525,7 @@ module Bullet = struct CErrors.register_handler (function | FailedBullet (b,sugg) -> - let prefix = str"Wrong bullet " ++ pr_bullet b ++ str" : " in + let prefix = str"Wrong bullet " ++ pr_bullet b ++ str": " in CErrors.user_err ~hdr:"Focus" (prefix ++ suggest_on_error sugg) | _ -> raise CErrors.Unhandled) -- cgit v1.2.3 From 87078054ad8af4dfb42c70f3460ef6cc01d48c5c Mon Sep 17 00:00:00 2001 From: Théo Zimmermann Date: Sat, 20 May 2017 15:43:20 +0200 Subject: Deprecate -nodoc. --- configure.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure.ml b/configure.ml index 5330da7d3..679f52417 100644 --- a/configure.ml +++ b/configure.ml @@ -328,7 +328,7 @@ let args_options = Arg.align [ "-browser", arg_string_option Prefs.browser, " Use to open URL %s"; "-nodoc", Arg.Clear Prefs.withdoc, - " Do not compile the documentation"; + " Deprecated: use -with-doc no instead"; "-with-doc", arg_bool Prefs.withdoc, "(yes|no) Compile the documentation or not"; "-with-geoproof", arg_bool Prefs.geoproof, -- cgit v1.2.3 From 04c532d81f69f0b6371b30b524b7aa258a6309c3 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 20 May 2017 17:07:29 +0200 Subject: Added a test for #4765 (an example of printing abbreviation with binders). --- test-suite/output/Notations3.out | 2 ++ test-suite/output/Notations3.v | 5 +++++ 2 files changed, 7 insertions(+) diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out index 360f37967..4d59a92cb 100644 --- a/test-suite/output/Notations3.out +++ b/test-suite/output/Notations3.out @@ -98,3 +98,5 @@ fun n : nat => foo4 n (fun _ y : nat => ETA z : nat, (fun _ : nat => y = 0)) : nat -> Prop tele (t : Type) '(y, z) (x : t0) := tt : forall t : Type, nat * nat -> t -> fpack +foo5 x nat x + : nat -> nat diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v index 4b8bfe312..96d831944 100644 --- a/test-suite/output/Notations3.v +++ b/test-suite/output/Notations3.v @@ -139,3 +139,8 @@ Notation "'tele' x .. z := b" := (at level 85, x binder, z binder). Check tele (t:Type) '((y,z):nat*nat) (x:t) := tt. + +(* Cyprien's part of bug #4765 *) + +Notation foo5 x T y := (fun x : T => y). +Check foo5 x nat x. -- cgit v1.2.3 From 3c0d8d08bda81b9fbd7210e4e352a08bbe8219e8 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 17 May 2017 20:55:32 +0200 Subject: [vernac] Remove `Save.` command. It has been deprecated for a while in favor of `Qed`. --- doc/refman/RefMan-pro.tex | 5 ----- doc/refman/RefMan-tus.tex | 2 +- doc/tutorial/Tutorial.tex | 8 +++----- ide/coq_commands.ml | 3 +-- ide/coqide.ml | 4 ++-- man/gallina.1 | 4 ++-- parsing/g_proofs.ml4 | 1 - test-suite/bugs/closed/348.v | 2 +- test-suite/bugs/closed/38.v | 2 +- test-suite/success/dependentind.v | 4 ++-- tools/gallina-syntax.el | 4 +--- tools/gallina_lexer.mll | 1 - 12 files changed, 14 insertions(+), 26 deletions(-) diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex index 8ba28b32f..7e2b9db24 100644 --- a/doc/refman/RefMan-pro.tex +++ b/doc/refman/RefMan-pro.tex @@ -90,11 +90,6 @@ memory overflow. Defines the proved term as a transparent constant. -\item {\tt Save.} -\comindex{Save} - - This is a deprecated equivalent to {\tt Qed}. - \item {\tt Save {\ident}.} Forces the name of the original goal to be {\ident}. This command diff --git a/doc/refman/RefMan-tus.tex b/doc/refman/RefMan-tus.tex index 017de6d48..7e5bb81a9 100644 --- a/doc/refman/RefMan-tus.tex +++ b/doc/refman/RefMan-tus.tex @@ -707,7 +707,7 @@ Once all the existential variables have been defined the derivation is completed, and a construction can be generated from the proof tree, replacing each of the existential variables by its definition. This is exactly what happens when one of the commands -\texttt{Qed}, \texttt{Save} or \texttt{Defined} is invoked +\texttt{Qed} or \texttt{Defined} is invoked (see Section~\ref{Qed}). The saved theorem becomes a defined constant, whose body is the proof object generated. diff --git a/doc/tutorial/Tutorial.tex b/doc/tutorial/Tutorial.tex index 0d537256b..30b6304c1 100644 --- a/doc/tutorial/Tutorial.tex +++ b/doc/tutorial/Tutorial.tex @@ -385,13 +385,11 @@ apply H; [ assumption | apply H0; assumption ]. Let us now save lemma \verb:distr_impl:: \begin{coq_example} -Save. +Qed. \end{coq_example} -Here \verb:Save: needs no argument, since we gave the name \verb:distr_impl: -in advance; -it is however possible to override the given name by giving a different -argument to command \verb:Save:. +Here \verb:Qed: needs no argument, since we gave the name \verb:distr_impl: +in advance. Actually, such an easy combination of tactics \verb:intro:, \verb:apply: and \verb:assumption: may be found completely automatically by an automatic diff --git a/ide/coq_commands.ml b/ide/coq_commands.ml index d55e7f9dd..5912bec35 100644 --- a/ide/coq_commands.ml +++ b/ide/coq_commands.ml @@ -105,8 +105,7 @@ let commands = [ "Reset Extraction Inline"; "Restore State"; ]; - [ "Save."; - "Scheme"; + [ "Scheme"; "Section"; "Set Extraction AutoInline"; "Set Extraction Optimize"; diff --git a/ide/coqide.ml b/ide/coqide.ml index 0b7567a5a..663e28d36 100644 --- a/ide/coqide.ml +++ b/ide/coqide.ml @@ -1103,8 +1103,8 @@ let build_ui () = menu templates_menu [ item "Templates" ~label:"Te_mplates"; - template_item ("Lemma new_lemma : .\nProof.\n\nSave.\n", 6,9, "J"); - template_item ("Theorem new_theorem : .\nProof.\n\nSave.\n", 8,11, "T"); + template_item ("Lemma new_lemma : .\nProof.\n\nQed.\n", 6,9, "J"); + template_item ("Theorem new_theorem : .\nProof.\n\nQed.\n", 8,11, "T"); template_item ("Definition ident := .\n", 11,5, "E"); template_item ("Inductive ident : :=\n | : .\n", 10,5, "I"); template_item ("Fixpoint ident (_ : _) {struct _} : _ :=\n.\n", 9,5, "F"); diff --git a/man/gallina.1 b/man/gallina.1 index 8c607216e..f8879c457 100644 --- a/man/gallina.1 +++ b/man/gallina.1 @@ -29,7 +29,7 @@ The suffix '.g' stands for Gallina. For that purpose, gallina removes all commands that follow a "Theorem", "Lemma", "Fact", "Remark" or "Goal" statement until it -reaches a command "Abort.", "Save.", "Qed.", "Defined." or "Proof +reaches a command "Abort.", "Qed.", "Defined." or "Proof <...>.". It also removes every "Hint", "Syntax", "Immediate" or "Transparent" command. @@ -52,7 +52,7 @@ Comments are removed in the *.g file. .SH NOTES Nested comments are correctly handled. In particular, every command -"Save." or "Abort." in a comment is not taken into account. +"Qed." or "Abort." in a comment is not taken into account. .SH BUGS diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 68b8be6b8..53637439a 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -48,7 +48,6 @@ GEXTEND Gram | IDENT "Qed" -> VernacEndProof (Proved (Opaque None,None)) | IDENT "Qed"; IDENT "exporting"; l = LIST0 identref SEP "," -> VernacEndProof (Proved (Opaque (Some l),None)) - | IDENT "Save" -> VernacEndProof (Proved (Opaque None,None)) | IDENT "Save"; tok = thm_token; id = identref -> VernacEndProof (Proved (Opaque None,Some (id,Some tok))) | IDENT "Save"; id = identref -> diff --git a/test-suite/bugs/closed/348.v b/test-suite/bugs/closed/348.v index 28cc5cb1e..48f0b5512 100644 --- a/test-suite/bugs/closed/348.v +++ b/test-suite/bugs/closed/348.v @@ -9,5 +9,5 @@ End D. Module D' (M:S). Import M. - Definition empty:Set. exact nat. Save. + Definition empty:Set. exact nat. Qed. End D'. diff --git a/test-suite/bugs/closed/38.v b/test-suite/bugs/closed/38.v index 4fc8d7c97..6b6e83779 100644 --- a/test-suite/bugs/closed/38.v +++ b/test-suite/bugs/closed/38.v @@ -14,7 +14,7 @@ Definition same := fun (l m : liste) => forall (x : A), e x l <-> e x m. Definition same_refl (x:liste) : (same x x). unfold same; split; intros; trivial. -Save. +Qed. Goal forall (x:liste), (same x x). intro. diff --git a/test-suite/success/dependentind.v b/test-suite/success/dependentind.v index 12ddbda84..f5bb884d2 100644 --- a/test-suite/success/dependentind.v +++ b/test-suite/success/dependentind.v @@ -15,7 +15,7 @@ Proof. intros n H. dependent destruction H. assumption. -Save. +Qed. Require Import ProofIrrelevance. @@ -25,7 +25,7 @@ Proof. dependent destruction v. exists v ; exists a. reflexivity. -Save. +Qed. (* Extraction Unnamed_thm. *) diff --git a/tools/gallina-syntax.el b/tools/gallina-syntax.el index c25abece1..662762b08 100644 --- a/tools/gallina-syntax.el +++ b/tools/gallina-syntax.el @@ -390,7 +390,7 @@ ("Corollary" "cor" "Corollary # : #.\nProof.\n#\nQed." t "Corollary") ("Declare Module :" "dmi" "Declare Module # : #.\n#\nEnd #." t) ("Declare Module <:" "dmi2" "Declare Module # <: #.\n#\nEnd #." t) - ("Definition goal" "defg" "Definition #:#.\n#\nSave." t);; careful + ("Definition goal" "defg" "Definition #:#.\n#\nQed." t);; careful ("Fact" "fct" "Fact # : #." t "Fact") ("Goal" nil "Goal #." t "Goal") ("Lemma" "l" "Lemma # : #.\nProof.\n#\nQed." t "Lemma") @@ -492,7 +492,6 @@ ("Require" nil "Require #." t "Require") ("Reserved Notation" nil "Reserved Notation" nil "Reserved\\s-+Notation") ("Reset Extraction Inline" nil "Reset Extraction Inline." t "Reset\\s-+Extraction\\s-+Inline") - ("Save" nil "Save." t "Save") ("Search" nil "Search #" nil "Search") ("SearchAbout" nil "SearchAbout #" nil "SearchAbout") ("SearchPattern" nil "SearchPattern #" nil "SearchPattern") @@ -710,7 +709,6 @@ Used by `coq-goal-command-p'" (defvar coq-keywords-save-strict '("Defined" - "Save" "Qed" "End" "Admitted" diff --git a/tools/gallina_lexer.mll b/tools/gallina_lexer.mll index 449efd57c..432b18e64 100644 --- a/tools/gallina_lexer.mll +++ b/tools/gallina_lexer.mll @@ -105,7 +105,6 @@ and end_of_line = parse | _ { print (Lexing.lexeme lexbuf) } and skip_proof = parse - | "Save." { end_of_line lexbuf } | "Save" space { skip_until_point lexbuf } | "Qed." { end_of_line lexbuf } -- cgit v1.2.3 From c1e9a27d383688e44ba34ada24fe08151cb5846e Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 17 May 2017 21:04:18 +0200 Subject: [vernac] Remove `Save thm id.` command. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We'd like to cleanup the `proof_end` type so we can have a smaller path in proof save. Note that the construction: ``` Goal Type. ⋮ Save id. ``` has to be handled by the STM in the same path as Defined (but with an opaque flag), as `Save id` will alter the environment and cannot be processed in parallel. We thus try to simply such paths a bit, as complexity of `lemmas.ml` seems like an issue these days. The form `Save Theorem id` doesn't really seem used, and moreover we should really add a type of "Goal", and unify syntax. It is often the case that beginners try `Goal addnC n : n + 0 = n." etc... --- CHANGES | 2 ++ dev/v8-syntax/syntax-v8.tex | 2 +- doc/refman/RefMan-pro.tex | 8 -------- intf/vernacexpr.mli | 4 +++- parsing/g_proofs.ml4 | 6 ++---- printing/ppvernac.ml | 4 +--- proofs/proof_global.ml | 2 +- proofs/proof_global.mli | 2 +- vernac/lemmas.ml | 11 +---------- 9 files changed, 12 insertions(+), 29 deletions(-) diff --git a/CHANGES b/CHANGES index b5f6ba927..31d691be0 100644 --- a/CHANGES +++ b/CHANGES @@ -28,6 +28,8 @@ Vernacular Commands - Goals context can be printed in a more compact way when "Set Printing Compact Contexts" is activated. +- The deprecated `Save` vernacular and its form `Save Theorem id` to + close proofs have been removed from the syntax. Please use `Qed`. Standard Library diff --git a/dev/v8-syntax/syntax-v8.tex b/dev/v8-syntax/syntax-v8.tex index 64431ea16..fa2864cec 100644 --- a/dev/v8-syntax/syntax-v8.tex +++ b/dev/v8-syntax/syntax-v8.tex @@ -1158,7 +1158,7 @@ $$ \nlsep \TERM{Abort}~\NT{ident} \nlsep \TERM{Existential}~\NT{num}~\KWD{:=}~\NT{constr-body} \nlsep \TERM{Qed} -\nlsep \TERM{Save}~\OPTGR{\NT{thm-token}~\NT{ident}} +\nlsep \TERM{Save}~\NT{ident}} \nlsep \TERM{Defined}~\OPT{\NT{ident}} \nlsep \TERM{Suspend} \nlsep \TERM{Resume}~\OPT{\NT{ident}} diff --git a/doc/refman/RefMan-pro.tex b/doc/refman/RefMan-pro.tex index 7e2b9db24..0760d716e 100644 --- a/doc/refman/RefMan-pro.tex +++ b/doc/refman/RefMan-pro.tex @@ -96,14 +96,6 @@ memory overflow. (and the following ones) can only be used if the original goal has been opened using the {\tt Goal} command. -\item {\tt Save Theorem {\ident}.} \\ - {\tt Save Lemma {\ident}.} \\ - {\tt Save Remark {\ident}.}\\ - {\tt Save Fact {\ident}.} - {\tt Save Corollary {\ident}.} - {\tt Save Proposition {\ident}.} - - Are equivalent to {\tt Save {\ident}.} \end{Variants} \subsection[\tt Admitted.]{\tt Admitted.\comindex{Admitted}\label{Admitted}} diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.mli index cb093d85d..94fa37eb6 100644 --- a/intf/vernacexpr.mli +++ b/intf/vernacexpr.mli @@ -143,6 +143,7 @@ type search_restriction = type rec_flag = bool (* true = Rec; false = NoRec *) type verbose_flag = bool (* true = Verbose; false = Silent *) + (* list of idents for qed exporting *) type opacity_flag = Opaque of lident list option | Transparent type coercion_flag = bool (* true = AddCoercion false = NoCoercion *) type instance_flag = bool option @@ -223,7 +224,8 @@ type syntax_modifier = type proof_end = | Admitted - | Proved of opacity_flag * (lident * theorem_kind option) option + (* name in `Save ident` when closing goal *) + | Proved of opacity_flag * lident option type scheme = | InductionScheme of bool * reference or_by_notation * sort_expr diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 53637439a..2def22290 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -48,13 +48,11 @@ GEXTEND Gram | IDENT "Qed" -> VernacEndProof (Proved (Opaque None,None)) | IDENT "Qed"; IDENT "exporting"; l = LIST0 identref SEP "," -> VernacEndProof (Proved (Opaque (Some l),None)) - | IDENT "Save"; tok = thm_token; id = identref -> - VernacEndProof (Proved (Opaque None,Some (id,Some tok))) | IDENT "Save"; id = identref -> - VernacEndProof (Proved (Opaque None,Some (id,None))) + VernacEndProof (Proved (Opaque None, Some id)) | IDENT "Defined" -> VernacEndProof (Proved (Transparent,None)) | IDENT "Defined"; id=identref -> - VernacEndProof (Proved (Transparent,Some (id,None))) + VernacEndProof (Proved (Transparent,Some id)) | IDENT "Restart" -> VernacRestart | IDENT "Undo" -> VernacUndo 1 | IDENT "Undo"; n = natural -> VernacUndo n diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 3e41439c8..76f1d8dd7 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -721,9 +721,7 @@ open Decl_kinds | Opaque (Some l) -> keyword "Qed" ++ spc() ++ str"export" ++ prlist_with_sep (fun () -> str", ") pr_lident l) - | Some (id,th) -> (match th with - | None -> (if opac <> Transparent then keyword "Save" else keyword "Defined") ++ spc() ++ pr_lident id - | Some tok -> keyword "Save" ++ spc() ++ pr_thm_token tok ++ spc() ++ pr_lident id) + | Some id -> (if opac <> Transparent then keyword "Save" else keyword "Defined") ++ spc() ++ pr_lident id ) | VernacExactProof c -> return (hov 2 (keyword "Proof" ++ pr_lconstrarg c)) diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 5f4a7766f..0443d6357 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -82,7 +82,7 @@ type proof_object = { type proof_ending = | Admitted of Names.Id.t * Decl_kinds.goal_kind * Entries.parameter_entry * proof_universes | Proved of Vernacexpr.opacity_flag * - (Vernacexpr.lident * Decl_kinds.theorem_kind option) option * + Vernacexpr.lident option * proof_object type proof_terminator = proof_ending -> unit type closed_proof = proof_object * proof_terminator diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli index 6bb6f5e2c..e90fb5bc9 100644 --- a/proofs/proof_global.mli +++ b/proofs/proof_global.mli @@ -70,7 +70,7 @@ type proof_ending = | Admitted of Names.Id.t * Decl_kinds.goal_kind * Entries.parameter_entry * proof_universes | Proved of Vernacexpr.opacity_flag * - (Vernacexpr.lident * Decl_kinds.theorem_kind option) option * + Vernacexpr.lident option * proof_object type proof_terminator type closed_proof = proof_object * proof_terminator diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index b79795aeb..0088b7079 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -280,13 +280,6 @@ let save_anonymous ?export_seff proof save_ident = check_anonymity id save_ident; save ?export_seff save_ident const cstrs pl do_guard persistence hook -let save_anonymous_with_strength ?export_seff proof kind save_ident = - let id,const,(cstrs,pl),do_guard,_,hook = proof in - check_anonymity id save_ident; - (* we consider that non opaque behaves as local for discharge *) - save ?export_seff save_ident const cstrs pl do_guard - (Global, const.const_entry_polymorphic, Proof kind) hook - (* Admitted *) let warn_let_as_axiom = @@ -337,9 +330,7 @@ let universe_proof_terminator compute_guard hook = (hook (Some (fst proof.Proof_global.universes))) is_opaque in begin match idopt with | None -> save_named ~export_seff proof - | Some ((_,id),None) -> save_anonymous ~export_seff proof id - | Some ((_,id),Some kind) -> - save_anonymous_with_strength ~export_seff proof kind id + | Some (_,id) -> save_anonymous ~export_seff proof id end; check_exist exports end -- cgit v1.2.3 From 47d4f5037926378d56b1a5050bd87415ed31d2dd Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 23 May 2017 16:16:46 -0400 Subject: unification.mli: Fix a spelling typo in a comment --- pretyping/unification.mli | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/pretyping/unification.mli b/pretyping/unification.mli index 148178f2f..8d7e3521d 100644 --- a/pretyping/unification.mli +++ b/pretyping/unification.mli @@ -45,7 +45,7 @@ val elim_no_delta_flags : unit -> unify_flags val is_keyed_unification : unit -> bool -(** The "unique" unification fonction *) +(** The "unique" unification function *) val w_unify : env -> evar_map -> conv_pb -> ?flags:unify_flags -> constr -> constr -> evar_map -- cgit v1.2.3