aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--CHANGES5
-rw-r--r--dev/top_printers.ml11
-rw-r--r--doc/refman/RefMan-ltac.tex21
-rw-r--r--interp/implicit_quantifiers.ml62
-rw-r--r--kernel/uGraph.ml16
-rw-r--r--plugins/ltac/extratactics.ml413
-rw-r--r--plugins/ltac/tactic_debug.ml14
-rw-r--r--plugins/setoid_ring/newring.ml20
-rw-r--r--pretyping/cases.ml6
-rw-r--r--pretyping/cbv.ml27
-rw-r--r--pretyping/glob_ops.ml215
-rw-r--r--pretyping/glob_ops.mli1
-rw-r--r--tactics/autorewrite.ml46
-rw-r--r--tactics/tactics.ml32
-rw-r--r--tactics/tactics.mli4
-rw-r--r--test-suite/bugs/closed/5219.v10
-rw-r--r--test-suite/success/boundvars.v14
-rw-r--r--test-suite/success/transparent_abstract.v21
-rw-r--r--theories/Init/Prelude.v2
-rw-r--r--toplevel/coqloop.ml8
-rw-r--r--toplevel/coqtop.ml10
-rw-r--r--toplevel/usage.ml1
-rw-r--r--toplevel/vernac.ml5
23 files changed, 250 insertions, 314 deletions
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
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")"
diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex
index 9378529cb..c2f52e23b 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,24 @@ 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.
+ 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}
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 7f11c0a3b..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) ->
@@ -131,61 +125,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/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 =
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 21419d1f9..3e6ccaf84 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -812,6 +812,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/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
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
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" *)
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
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index ebbfa195f..6509aaac3 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -214,120 +214,62 @@ 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 same_id na id = match na with
-| Anonymous -> false
-| Name id' -> Id.equal id id'
+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 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
-
-
-let add_name_to_ids set na =
- match na with
- | Anonymous -> set
- | Name id -> Id.Set.add id set
+ | 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 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
@@ -353,57 +295,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/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/tactics/autorewrite.ml b/tactics/autorewrite.ml
index e58ec5a31..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
@@ -127,45 +125,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 ->
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 211a7338b..556df6e55 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 ~opaque ?(goal_type=None) id gk tac tacK =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
@@ -4927,7 +4927,10 @@ let abstract_subproof id gk tac =
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 _ ->
@@ -4957,8 +4960,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 +4979,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 ~opaque id gk tac =
+ 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"
-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 +5001,13 @@ 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 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 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 ba4a9706d..07a803542 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 : 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
val abstract_generalize : ?generalize_vars:bool -> ?force_dep:bool -> Id.t -> unit Proofview.tactic
val specialize_eqs : Id.t -> unit Proofview.tactic
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.
diff --git a/test-suite/success/boundvars.v b/test-suite/success/boundvars.v
new file mode 100644
index 000000000..fafe27292
--- /dev/null
+++ b/test-suite/success/boundvars.v
@@ -0,0 +1,14 @@
+(* 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.
+
+(* 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.
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_".
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/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\
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