aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--dev/doc/changes.txt5
-rw-r--r--dev/printers.mllib13
-rw-r--r--engine/engine.mllib2
-rw-r--r--engine/evarutil.ml (renamed from pretyping/evarutil.ml)244
-rw-r--r--engine/evarutil.mli (renamed from pretyping/evarutil.mli)48
-rw-r--r--engine/proofview.ml (renamed from proofs/proofview.ml)158
-rw-r--r--engine/proofview.mli (renamed from proofs/proofview.mli)48
-rw-r--r--engine/proofview_monad.ml4
-rw-r--r--engine/proofview_monad.mli4
-rw-r--r--pretyping/cases.ml1
-rw-r--r--pretyping/coercion.ml6
-rw-r--r--pretyping/evarconv.ml1
-rw-r--r--pretyping/evardefine.ml209
-rw-r--r--pretyping/evardefine.mli46
-rw-r--r--pretyping/pretyping.ml18
-rw-r--r--pretyping/pretyping.mli4
-rw-r--r--pretyping/pretyping.mllib4
-rw-r--r--pretyping/reductionops.ml30
-rw-r--r--pretyping/typeclasses.ml2
-rw-r--r--pretyping/typing.ml4
-rw-r--r--pretyping/unification.ml1
-rw-r--r--printing/printer.ml2
-rw-r--r--proofs/proof.ml28
-rw-r--r--proofs/proofs.mllib2
-rw-r--r--proofs/refine.ml122
-rw-r--r--proofs/refine.mli37
-rw-r--r--tactics/extratactics.ml42
-rw-r--r--tactics/hints.ml2
-rw-r--r--tactics/inv.ml2
-rw-r--r--tactics/rewrite.ml8
-rw-r--r--tactics/tactics.ml38
-rw-r--r--tactics/tactics.mli2
-rw-r--r--toplevel/classes.ml8
-rw-r--r--toplevel/himsg.ml15
-rw-r--r--toplevel/record.ml2
35 files changed, 618 insertions, 504 deletions
diff --git a/dev/doc/changes.txt b/dev/doc/changes.txt
index 1f5ba7862..2f631c633 100644
--- a/dev/doc/changes.txt
+++ b/dev/doc/changes.txt
@@ -76,6 +76,11 @@
the case of (VERNAC) ARGUMENT EXTEND, the name of the argument entry is bound
in the parsing rules, so beware of recursive calls.
+- Evarutil was split in two parts. The new Evardefine file exposes functions
+define_evar_* mostly used internally in the unification engine.
+
+- The Refine module was move out of Proofview.
+
=========================================
= CHANGES BETWEEN COQ V8.4 AND COQ V8.5 =
=========================================
diff --git a/dev/printers.mllib b/dev/printers.mllib
index 7710033db..a3ba42ba7 100644
--- a/dev/printers.mllib
+++ b/dev/printers.mllib
@@ -123,14 +123,16 @@ Evd
Sigma
Glob_ops
Redops
+Pretype_errors
+Evarutil
Reductionops
Inductiveops
Arguments_renaming
Nativenorm
Retyping
Cbv
-Pretype_errors
-Evarutil
+
+Evardefine
Evarsolve
Recordops
Evarconv
@@ -141,6 +143,9 @@ Find_subterm
Tacred
Classops
Typeclasses_errors
+Logic_monad
+Proofview_monad
+Proofview
Typeclasses
Detyping
Indrec
@@ -181,9 +186,7 @@ Refiner
Clenv
Evar_refiner
Proof_errors
-Logic_monad
-Proofview_monad
-Proofview
+Refine
Proof
Proof_global
Pfedit
diff --git a/engine/engine.mllib b/engine/engine.mllib
index 7197a2583..70b74edca 100644
--- a/engine/engine.mllib
+++ b/engine/engine.mllib
@@ -5,3 +5,5 @@ UState
Evd
Sigma
Proofview_monad
+Evarutil
+Proofview
diff --git a/pretyping/evarutil.ml b/engine/evarutil.ml
index ab70de057..2bd67dcdc 100644
--- a/pretyping/evarutil.ml
+++ b/engine/evarutil.ml
@@ -17,10 +17,12 @@ open Namegen
open Pre_env
open Environ
open Evd
-open Reductionops
-open Pretype_errors
open Sigma.Notations
+let safe_evar_value sigma ev =
+ try Some (Evd.existential_value sigma ev)
+ with NotInstantiatedEvar | Not_found -> None
+
(** Combinators *)
let evd_comb0 f evdref =
@@ -62,33 +64,41 @@ let rec flush_and_check_evars sigma c =
(* let nf_evar_key = Profile.declare_profile "nf_evar" *)
(* let nf_evar = Profile.profile2 nf_evar_key Reductionops.nf_evar *)
-let nf_evar = Reductionops.nf_evar
+
+let rec whd_evar sigma c =
+ match kind_of_term c with
+ | Evar ev ->
+ let (evk, args) = ev in
+ let args = Array.map (fun c -> whd_evar sigma c) args in
+ (match safe_evar_value sigma (evk, args) with
+ Some c -> whd_evar sigma c
+ | None -> c)
+ | Sort (Type u) ->
+ let u' = Evd.normalize_universe sigma u in
+ if u' == u then c else mkSort (Sorts.sort_of_univ u')
+ | Const (c', u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstU (c', u')
+ | Ind (i, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkIndU (i, u')
+ | Construct (co, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstructU (co, u')
+ | _ -> c
+
+let rec nf_evar sigma t = Constr.map (fun t -> nf_evar sigma t) (whd_evar sigma t)
+
let j_nf_evar sigma j =
{ uj_val = nf_evar sigma j.uj_val;
uj_type = nf_evar sigma j.uj_type }
-let j_nf_betaiotaevar sigma j =
- { uj_val = nf_evar sigma j.uj_val;
- uj_type = Reductionops.nf_betaiota sigma j.uj_type }
let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl
-let jv_nf_betaiotaevar sigma jl =
- Array.map (j_nf_betaiotaevar sigma) jl
let jv_nf_evar sigma = Array.map (j_nf_evar sigma)
let tj_nf_evar sigma {utj_val=v;utj_type=t} =
{utj_val=nf_evar sigma v;utj_type=t}
-let env_nf_evar sigma env =
- let open Context.Rel.Declaration in
- process_rel_context
- (fun d e -> push_rel (map_constr (nf_evar sigma) d) e) env
-
-let env_nf_betaiotaevar sigma env =
- let open Context.Rel.Declaration in
- process_rel_context
- (fun d e ->
- push_rel (map_constr (Reductionops.nf_betaiota sigma) d) e) env
-
let nf_evars_universes evm =
- Universes.nf_evars_and_universes_opt_subst (Reductionops.safe_evar_value evm)
+ Universes.nf_evars_and_universes_opt_subst (safe_evar_value evm)
(Evd.universe_subst evm)
let nf_evars_and_universes evm =
@@ -659,23 +669,6 @@ let undefined_evars_of_evar_info evd evi =
(undefined_evars_of_named_context evd
(named_context_of_val evi.evar_hyps)))
-(* [check_evars] fails if some unresolved evar remains *)
-
-let check_evars env initial_sigma sigma c =
- let rec proc_rec c =
- match kind_of_term c with
- | Evar (evk,_ as ev) ->
- (match existential_opt_value sigma ev with
- | Some c -> proc_rec c
- | None ->
- if not (Evd.mem initial_sigma evk) then
- let (loc,k) = evar_source evk sigma in
- match k with
- | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
- | _ -> error_unsolvable_implicit loc env sigma evk None)
- | _ -> iter_constr proc_rec c
- in proc_rec c
-
(* spiwack: this is a more complete version of
{!Termops.occur_evar}. The latter does not look recursively into an
[evar_map]. If unification only need to check superficially, tactics
@@ -688,146 +681,6 @@ let occur_evar_upto sigma n c =
in
try occur_rec c; false with Occur -> true
-
-(****************************************)
-(* Operations on value/type constraints *)
-(****************************************)
-
-type type_constraint = types option
-
-type val_constraint = constr option
-
-(* Old comment...
- * Basically, we have the following kind of constraints (in increasing
- * strength order):
- * (false,(None,None)) -> no constraint at all
- * (true,(None,None)) -> we must build a judgement which _TYPE is a kind
- * (_,(None,Some ty)) -> we must build a judgement which _TYPE is ty
- * (_,(Some v,_)) -> we must build a judgement which _VAL is v
- * Maybe a concrete datatype would be easier to understand.
- * We differentiate (true,(None,None)) from (_,(None,Some Type))
- * because otherwise Case(s) would be misled, as in
- * (n:nat) Case n of bool [_]nat end would infer the predicate Type instead
- * of Set.
- *)
-
-(* The empty type constraint *)
-let empty_tycon = None
-
-(* Builds a type constraint *)
-let mk_tycon ty = Some ty
-
-(* Constrains the value of a type *)
-let empty_valcon = None
-
-(* Builds a value constraint *)
-let mk_valcon c = Some c
-
-let idx = Namegen.default_dependent_ident
-
-(* Refining an evar to a product *)
-
-let define_pure_evar_as_product evd evk =
- let open Context.Named.Declaration in
- let evi = Evd.find_undefined evd evk in
- let evenv = evar_env evi in
- let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in
- let concl = whd_betadeltaiota evenv evd evi.evar_concl in
- let s = destSort concl in
- let evd1,(dom,u1) =
- let evd = Sigma.Unsafe.of_evar_map evd in
- let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in
- (Sigma.to_evar_map evd1, e)
- in
- let evd2,rng =
- let newenv = push_named (LocalAssum (id, dom)) evenv in
- let src = evar_source evk evd1 in
- let filter = Filter.extend 1 (evar_filter evi) in
- if is_prop_sort s then
- (* Impredicative product, conclusion must fall in [Prop]. *)
- new_evar_unsafe newenv evd1 concl ~src ~filter
- else
- let status = univ_flexible_alg in
- let evd3, (rng, srng) =
- let evd1 = Sigma.Unsafe.of_evar_map evd1 in
- let Sigma (e, evd3, _) = new_type_evar newenv evd1 status ~src ~filter in
- (Sigma.to_evar_map evd3, e)
- in
- let prods = Univ.sup (univ_of_sort u1) (univ_of_sort srng) in
- let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in
- evd3, rng
- in
- let prod = mkProd (Name id, dom, subst_var id rng) in
- let evd3 = Evd.define evk prod evd2 in
- evd3,prod
-
-(* Refine an applied evar to a product and returns its instantiation *)
-
-let define_evar_as_product evd (evk,args) =
- let evd,prod = define_pure_evar_as_product evd evk in
- (* Quick way to compute the instantiation of evk with args *)
- let na,dom,rng = destProd prod in
- let evdom = mkEvar (fst (destEvar dom), args) in
- let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evrng = mkEvar (fst (destEvar rng), evrngargs) in
- evd,mkProd (na, evdom, evrng)
-
-(* Refine an evar with an abstraction
-
- I.e., solve x1..xq |- ?e:T(x1..xq) with e:=λy:A.?e'[x1..xq,y] where:
- - either T(x1..xq) = πy:A(x1..xq).B(x1..xq,y)
- or T(x1..xq) = ?d[x1..xq] and we define ?d := πy:?A.?B
- with x1..xq |- ?A:Type and x1..xq,y |- ?B:Type
- - x1..xq,y:A |- ?e':B
-*)
-
-let define_pure_evar_as_lambda env evd evk =
- let open Context.Named.Declaration in
- let evi = Evd.find_undefined evd evk in
- let evenv = evar_env evi in
- let typ = whd_betadeltaiota evenv evd (evar_concl evi) in
- let evd1,(na,dom,rng) = match kind_of_term typ with
- | Prod (na,dom,rng) -> (evd,(na,dom,rng))
- | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ
- | _ -> error_not_product_loc Loc.ghost env evd typ in
- let avoid = ids_of_named_context (evar_context evi) in
- let id =
- next_name_away_with_default_using_types "x" na avoid (whd_evar evd dom) in
- let newenv = push_named (LocalAssum (id, dom)) evenv in
- let filter = Filter.extend 1 (evar_filter evi) in
- let src = evar_source evk evd1 in
- let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in
- let lam = mkLambda (Name id, dom, subst_var id body) in
- Evd.define evk lam evd2, lam
-
-let define_evar_as_lambda env evd (evk,args) =
- let evd,lam = define_pure_evar_as_lambda env evd evk in
- (* Quick way to compute the instantiation of evk with args *)
- let na,dom,body = destLambda lam in
- let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evbody = mkEvar (fst (destEvar body), evbodyargs) in
- evd,mkLambda (na, dom, evbody)
-
-let rec evar_absorb_arguments env evd (evk,args as ev) = function
- | [] -> evd,ev
- | a::l ->
- (* TODO: optimize and avoid introducing intermediate evars *)
- let evd,lam = define_pure_evar_as_lambda env evd evk in
- let _,_,body = destLambda lam in
- let evk = fst (destEvar body) in
- evar_absorb_arguments env evd (evk, Array.cons a args) l
-
-(* Refining an evar to a sort *)
-
-let define_evar_as_sort env evd (ev,args) =
- let evd, u = new_univ_variable univ_rigid evd in
- let evi = Evd.find_undefined evd ev in
- let s = Type u in
- let concl = whd_betadeltaiota (evar_env evi) evd evi.evar_concl in
- let sort = destSort concl in
- let evd' = Evd.define ev (mkSort s) evd in
- Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s
-
(* We don't try to guess in which sort the type should be defined, since
any type has type Type. May cause some trouble, but not so far... *)
@@ -835,38 +688,6 @@ let judge_of_new_Type evd =
let Sigma (s, evd', p) = Sigma.new_univ_variable univ_rigid evd in
Sigma ({ uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) }, evd', p)
-(* Propagation of constraints through application and abstraction:
- Given a type constraint on a functional term, returns the type
- constraint on its domain and codomain. If the input constraint is
- an evar instantiate it with the product of 2 new evars. *)
-
-let split_tycon loc env evd tycon =
- let rec real_split evd c =
- let t = whd_betadeltaiota env evd c in
- match kind_of_term t with
- | Prod (na,dom,rng) -> evd, (na, dom, rng)
- | Evar ev (* ev is undefined because of whd_betadeltaiota *) ->
- let (evd',prod) = define_evar_as_product evd ev in
- let (_,dom,rng) = destProd prod in
- evd',(Anonymous, dom, rng)
- | App (c,args) when isEvar c ->
- let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in
- real_split evd' (mkApp (lam,args))
- | _ -> error_not_product_loc loc env evd c
- in
- match tycon with
- | None -> evd,(Anonymous,None,None)
- | Some c ->
- let evd', (n, dom, rng) = real_split evd c in
- evd', (n, mk_tycon dom, mk_tycon rng)
-
-let valcon_of_tycon x = x
-let lift_tycon n = Option.map (lift n)
-
-let pr_tycon env = function
- None -> str "None"
- | Some t -> Termops.print_constr_env env t
-
let subterm_source evk (loc,k) =
let evk = match k with
| Evar_kinds.SubEvar (evk) -> evk
@@ -876,7 +697,7 @@ let subterm_source evk (loc,k) =
(** Term exploration up to instantiation. *)
let kind_of_term_upto sigma t =
- Constr.kind (Reductionops.whd_evar sigma t)
+ Constr.kind (whd_evar sigma t)
(** [eq_constr_univs_test sigma1 sigma2 t u] tests equality of [t] and
[u] up to existential variable instantiation and equalisable
@@ -897,3 +718,6 @@ let eq_constr_univs_test sigma1 sigma2 t u =
(universes sigma2) fold t u sigma2
in
match ans with None -> false | Some _ -> true
+
+type type_constraint = types option
+type val_constraint = constr option
diff --git a/pretyping/evarutil.mli b/engine/evarutil.mli
index bc4c37a91..ffff2c5dd 100644
--- a/pretyping/evarutil.mli
+++ b/engine/evarutil.mli
@@ -78,6 +78,8 @@ val new_evar_instance :
val make_pure_subst : evar_info -> constr array -> (Id.t * constr) list
+val safe_evar_value : evar_map -> existential -> constr option
+
(** {6 Evars/Metas switching...} *)
val non_instantiated : evar_map -> evar_info Evar.Map.t
@@ -96,20 +98,6 @@ val has_undefined_evars : evar_map -> constr -> bool
val is_ground_term : evar_map -> constr -> bool
val is_ground_env : evar_map -> env -> bool
-(** [check_evars env initial_sigma extended_sigma c] fails if some
- new unresolved evar remains in [c] *)
-val check_evars : env -> evar_map -> evar_map -> constr -> unit
-
-val define_evar_as_product : evar_map -> existential -> evar_map * types
-val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types
-val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
-
-(** Instantiate an evar by as many lambda's as needed so that its arguments
- are moved to the evar substitution (i.e. turn [?x[vars1:=args1] args] into
- [?y[vars1:=args1,vars:=args]] with
- [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *)
-val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
- evar_map * existential
(** [gather_dependent_evars evm seeds] classifies the evars in [evm]
as dependent_evars and goals (these may overlap). A goal is an
@@ -140,26 +128,12 @@ val occur_evar_upto : evar_map -> Evar.t -> Constr.t -> bool
val judge_of_new_Type : 'r Sigma.t -> (unsafe_judgment, 'r) Sigma.sigma
-type type_constraint = types option
-type val_constraint = constr option
-
-val empty_tycon : type_constraint
-val mk_tycon : constr -> type_constraint
-val empty_valcon : val_constraint
-val mk_valcon : constr -> val_constraint
-
-val split_tycon :
- Loc.t -> env -> evar_map -> type_constraint ->
- evar_map * (Name.t * type_constraint * type_constraint)
-
-val valcon_of_tycon : type_constraint -> val_constraint
-val lift_tycon : int -> type_constraint -> type_constraint
-
(***********************************************************)
(** [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains
uninstantiated; [nf_evar] leaves uninstantiated evars as is *)
+val whd_evar : evar_map -> constr -> constr
val nf_evar : evar_map -> constr -> constr
val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment
val jl_nf_evar :
@@ -177,12 +151,6 @@ val nf_evar_info : evar_map -> evar_info -> evar_info
val nf_evar_map : evar_map -> evar_map
val nf_evar_map_undefined : evar_map -> evar_map
-val env_nf_evar : evar_map -> env -> env
-val env_nf_betaiotaevar : evar_map -> env -> env
-
-val j_nf_betaiotaevar : evar_map -> unsafe_judgment -> unsafe_judgment
-val jv_nf_betaiotaevar :
- evar_map -> unsafe_judgment array -> unsafe_judgment array
(** Presenting terms without solved evars *)
val nf_evars_universes : evar_map -> constr -> constr
@@ -212,11 +180,6 @@ val kind_of_term_upto : evar_map -> constr -> (constr,types) kind_of_term
assumed to be an extention of those in [sigma1]. *)
val eq_constr_univs_test : evar_map -> evar_map -> constr -> constr -> bool
-(** {6 debug pretty-printer:} *)
-
-val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
-
-
(** {6 Removing hyps in evars'context}
raise OccurHypInSimpleClause if the removal breaks dependencies *)
@@ -251,3 +214,8 @@ val subterm_source : existential_key -> Evar_kinds.t Loc.located ->
Evar_kinds.t Loc.located
val meta_counter_summary_name : string
+
+(** Deprecater *)
+
+type type_constraint = types option
+type val_constraint = constr option
diff --git a/proofs/proofview.ml b/engine/proofview.ml
index b68fa042e..ba664cafa 100644
--- a/proofs/proofview.ml
+++ b/engine/proofview.ml
@@ -56,10 +56,12 @@ type telescope =
| TNil of Evd.evar_map
| TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+let typeclass_resolvable = Evd.Store.field ()
+
let dependent_init =
(* Goals are created with a store which marks them as unresolvable
for type classes. *)
- let store = Typeclasses.set_resolvable Evd.Store.empty false in
+ let store = Evd.Store.set Evd.Store.empty typeclass_resolvable () in
(* Goals don't have a source location. *)
let src = (Loc.ghost,Evar_kinds.GoalEvar) in
(* Main routine *)
@@ -625,18 +627,18 @@ let shelve_unifiable =
InfoL.leaf (Info.Tactic (fun () -> Pp.str"shelve_unifiable")) >>
Shelf.modify (fun gls -> gls @ u)
-(** [guard_no_unifiable] fails with error [UnresolvedBindings] if some
+(** [guard_no_unifiable] returns the list of unifiable goals if some
goals are unifiable (see {!shelve_unifiable}) in the current focus. *)
let guard_no_unifiable =
let open Proof in
Pv.get >>= fun initial ->
let (u,n) = partition_unifiable initial.solution initial.comb in
match u with
- | [] -> tclUNIT ()
+ | [] -> tclUNIT None
| gls ->
let l = CList.map (fun g -> Evd.dependent_evar_ident g initial.solution) gls in
let l = CList.map (fun id -> Names.Name id) l in
- tclZERO (Logic.RefinerError (Logic.UnresolvedBindings l))
+ tclUNIT (Some l)
(** [unshelve l p] adds all the goals in [l] at the end of the focused
goals of p *)
@@ -901,18 +903,22 @@ module Unsafe = struct
let reset_future_goals p =
{ p with solution = Evd.reset_future_goals p.solution }
- let mark_as_goal_evm evd content =
+ let mark_as_goal evd content =
let info = Evd.find evd content in
let info =
{ info with Evd.evar_source = match info.Evd.evar_source with
| _, (Evar_kinds.VarInstance _ | Evar_kinds.GoalEvar) as x -> x
| loc,_ -> loc,Evar_kinds.GoalEvar }
in
- let info = Typeclasses.mark_unresolvable info in
+ let info = match Evd.Store.get info.Evd.evar_extra typeclass_resolvable with
+ | None -> { info with Evd.evar_extra = Evd.Store.set info.Evd.evar_extra typeclass_resolvable () }
+ | Some () -> info
+ in
Evd.add evd content info
- let mark_as_goal p gl =
- { p with solution = mark_as_goal_evm p.solution gl }
+ let advance = advance
+
+ let typeclass_resolvable = typeclass_resolvable
end
@@ -924,8 +930,20 @@ let (<+>) t1 t2 = tclOR t1 (fun _ -> t2)
(** {6 Goal-dependent tactics} *)
-(* To avoid shadowing by the local [Goal] module *)
-module GoalV82 = Goal.V82
+let goal_env evars gl =
+ let evi = Evd.find evars gl in
+ Evd.evar_filtered_env evi
+
+let goal_nf_evar sigma gl =
+ let evi = Evd.find sigma gl in
+ let evi = Evarutil.nf_evar_info sigma evi in
+ let sigma = Evd.add sigma gl evi in
+ (gl, sigma)
+
+let goal_extra evars gl =
+ let evi = Evd.find evars gl in
+ evi.Evd.evar_extra
+
let catchable_exception = function
| Logic_monad.Exception _ -> false
@@ -950,7 +968,7 @@ module Goal = struct
let sigma { sigma=sigma } = Sigma.Unsafe.of_evar_map sigma
let hyps { env=env } = Environ.named_context env
let concl { concl=concl } = concl
- let extra { sigma=sigma; self=self } = Goal.V82.extra sigma self
+ let extra { sigma=sigma; self=self } = goal_extra sigma self
let raw_concl { concl=concl } = concl
@@ -1063,118 +1081,6 @@ end
-(** {6 The refine tactic} *)
-
-module Refine =
-struct
-
- let extract_prefix env info =
- let ctx1 = List.rev (Environ.named_context env) in
- let ctx2 = List.rev (Evd.evar_context info) in
- let rec share l1 l2 accu = match l1, l2 with
- | d1 :: l1, d2 :: l2 ->
- if d1 == d2 then share l1 l2 (d1 :: accu)
- else (accu, d2 :: l2)
- | _ -> (accu, l2)
- in
- share ctx1 ctx2 []
-
- let typecheck_evar ev env sigma =
- let info = Evd.find sigma ev in
- (** Typecheck the hypotheses. *)
- let type_hyp (sigma, env) decl =
- let t = get_type decl in
- let evdref = ref sigma in
- let _ = Typing.e_sort_of env evdref t in
- let () = match decl with
- | LocalAssum _ -> ()
- | LocalDef (_,body,_) -> Typing.e_check env evdref body t
- in
- (!evdref, Environ.push_named decl env)
- in
- let (common, changed) = extract_prefix env info in
- let env = Environ.reset_with_named_context (Environ.val_of_named_context common) env in
- let (sigma, env) = List.fold_left type_hyp (sigma, env) changed in
- (** Typecheck the conclusion *)
- let evdref = ref sigma in
- let _ = Typing.e_sort_of env evdref (Evd.evar_concl info) in
- !evdref
-
- let typecheck_proof c concl env sigma =
- let evdref = ref sigma in
- let () = Typing.e_check env evdref c concl in
- !evdref
-
- let (pr_constrv,pr_constr) =
- Hook.make ~default:(fun _env _sigma _c -> Pp.str"<constr>") ()
-
- let refine ?(unsafe = true) f = Goal.enter { Goal.enter = begin fun gl ->
- let sigma = Goal.sigma gl in
- let sigma = Sigma.to_evar_map sigma in
- let env = Goal.env gl in
- let concl = Goal.concl gl in
- (** Save the [future_goals] state to restore them after the
- refinement. *)
- let prev_future_goals = Evd.future_goals sigma in
- let prev_principal_goal = Evd.principal_future_goal sigma in
- (** Create the refinement term *)
- let (c, sigma) = Sigma.run (Evd.reset_future_goals sigma) f in
- let evs = Evd.future_goals sigma in
- let evkmain = Evd.principal_future_goal sigma in
- (** Check that the introduced evars are well-typed *)
- let fold accu ev = typecheck_evar ev env accu in
- let sigma = if unsafe then sigma else CList.fold_left fold sigma evs in
- (** Check that the refined term is typesafe *)
- let sigma = if unsafe then sigma else typecheck_proof c concl env sigma in
- (** Check that the goal itself does not appear in the refined term *)
- let _ =
- if not (Evarutil.occur_evar_upto sigma gl.Goal.self c) then ()
- else Pretype_errors.error_occur_check env sigma gl.Goal.self c
- in
- (** Proceed to the refinement *)
- let sigma = match evkmain with
- | None -> Evd.define gl.Goal.self c sigma
- | Some evk ->
- let id = Evd.evar_ident gl.Goal.self sigma in
- let sigma = Evd.define gl.Goal.self c sigma in
- match id with
- | None -> sigma
- | Some id -> Evd.rename evk id sigma
- in
- (** Restore the [future goals] state. *)
- let sigma = Evd.restore_future_goals sigma prev_future_goals prev_principal_goal in
- (** Select the goals *)
- let comb = undefined sigma (CList.rev evs) in
- let sigma = CList.fold_left Unsafe.mark_as_goal_evm sigma comb in
- let open Proof in
- InfoL.leaf (Info.Tactic (fun () -> Pp.(hov 2 (str"refine"++spc()++ Hook.get pr_constrv env sigma c)))) >>
- Pv.modify (fun ps -> { ps with solution = sigma; comb; })
- end }
-
- (** Useful definitions *)
-
- let with_type env evd c t =
- let my_type = Retyping.get_type_of env evd c in
- let j = Environ.make_judge c my_type in
- let (evd,j') =
- Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t
- in
- evd , j'.Environ.uj_val
-
- let refine_casted ?unsafe f = Goal.enter { Goal.enter = begin fun gl ->
- let concl = Goal.concl gl in
- let env = Goal.env gl in
- let f = { run = fun h ->
- let Sigma (c, h, p) = f.run h in
- let sigma, c = with_type env (Sigma.to_evar_map h) c concl in
- Sigma (c, Sigma.Unsafe.of_evar_map sigma, p)
- } in
- refine ?unsafe f
- end }
-end
-
-
-
(** {6 Trace} *)
module Trace = struct
@@ -1225,7 +1131,7 @@ module V82 = struct
in
(* Old style tactics expect the goals normalized with respect to evars. *)
let (initgoals,initevd) =
- Evd.Monad.List.map (fun g s -> GoalV82.nf_evar s g) ps.comb ps.solution
+ Evd.Monad.List.map (fun g s -> goal_nf_evar s g) ps.comb ps.solution
in
let (goalss,evd) = Evd.Monad.List.map tac initgoals initevd in
let sgs = CList.flatten goalss in
@@ -1241,7 +1147,7 @@ module V82 = struct
solution. *)
let nf_evar_goals =
Pv.modify begin fun ps ->
- let map g s = GoalV82.nf_evar s g in
+ let map g s = goal_nf_evar s g in
let (goals,evd) = Evd.Monad.List.map map ps.comb ps.solution in
{ ps with solution = evd; comb = goals; }
end
@@ -1275,7 +1181,7 @@ module V82 = struct
let of_tactic t gls =
try
let init = { shelf = []; solution = gls.Evd.sigma ; comb = [gls.Evd.it] } in
- let (_,final,_,_) = apply (GoalV82.env gls.Evd.sigma gls.Evd.it) t init in
+ let (_,final,_,_) = apply (goal_env gls.Evd.sigma gls.Evd.it) t init in
{ Evd.sigma = final.solution ; it = final.comb }
with Logic_monad.TacticFailure e as src ->
let (_, info) = Errors.push src in
diff --git a/proofs/proofview.mli b/engine/proofview.mli
index 61014468b..7996b7969 100644
--- a/proofs/proofview.mli
+++ b/engine/proofview.mli
@@ -295,9 +295,9 @@ val shelve : unit tactic
considered). *)
val shelve_unifiable : unit tactic
-(** [guard_no_unifiable] fails with error [UnresolvedBindings] if some
+(** [guard_no_unifiable] returns the list of unifiable goals if some
goals are unifiable (see {!shelve_unifiable}) in the current focus. *)
-val guard_no_unifiable : unit tactic
+val guard_no_unifiable : Names.Name.t list option tactic
(** [unshelve l p] adds all the goals in [l] at the end of the focused
goals of p *)
@@ -406,7 +406,16 @@ module Unsafe : sig
(** Give an evar the status of a goal (changes its source location
and makes it unresolvable for type classes. *)
- val mark_as_goal : proofview -> Evar.t -> proofview
+ val mark_as_goal : Evd.evar_map -> Evar.t -> Evd.evar_map
+
+ (** [advance sigma g] returns [Some g'] if [g'] is undefined and is
+ the current avatar of [g] (for instance [g] was changed by [clear]
+ into [g']). It returns [None] if [g] has been (partially)
+ solved. *)
+ val advance : Evd.evar_map -> Evar.t -> Evar.t option
+
+ val typeclass_resolvable : unit Evd.Store.field
+
end
(** This module gives access to the innards of the monad. Its use is
@@ -491,39 +500,6 @@ module Goal : sig
end
-(** {6 The refine tactic} *)
-
-module Refine : sig
-
- (** Printer used to print the constr which refine refines. *)
- val pr_constr :
- (Environ.env -> Evd.evar_map -> Term.constr -> Pp.std_ppcmds) Hook.t
-
- (** {7 Refinement primitives} *)
-
- val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
- (** In [refine ?unsafe t], [t] is a term with holes under some
- [evar_map] context. The term [t] is used as a partial solution
- for the current goal (refine is a goal-dependent tactic), the
- new holes created by [t] become the new subgoals. Exceptions
- raised during the interpretation of [t] are caught and result in
- tactic failures. If [unsafe] is [false] (default is [true]) [t] is
- type-checked beforehand. *)
-
- (** {7 Helper functions} *)
-
- val with_type : Environ.env -> Evd.evar_map ->
- Term.constr -> Term.types -> Evd.evar_map * Term.constr
- (** [with_type env sigma c t] ensures that [c] is of type [t]
- inserting a coercion if needed. *)
-
- val refine_casted : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
- (** Like {!refine} except the refined term is coerced to the conclusion of the
- current goal. *)
-
-end
-
-
(** {6 Trace} *)
module Trace : sig
diff --git a/engine/proofview_monad.ml b/engine/proofview_monad.ml
index 2b9db60b4..6f52b3ee9 100644
--- a/engine/proofview_monad.ml
+++ b/engine/proofview_monad.ml
@@ -154,8 +154,8 @@ end
focused goals. *)
type proofview = {
solution : Evd.evar_map;
- comb : Goal.goal list;
- shelf : Goal.goal list;
+ comb : Evar.t list;
+ shelf : Evar.t list;
}
(** {6 Instantiation of the logic monad} *)
diff --git a/engine/proofview_monad.mli b/engine/proofview_monad.mli
index 7a6ea10fe..0aff0a720 100644
--- a/engine/proofview_monad.mli
+++ b/engine/proofview_monad.mli
@@ -70,8 +70,8 @@ end
focused goals. *)
type proofview = {
solution : Evd.evar_map;
- comb : Goal.goal list;
- shelf : Goal.goal list;
+ comb : Evar.t list;
+ shelf : Evar.t list;
}
(** {6 Instantiation of the logic monad} *)
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 8a55a7aaa..c3968f896 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -25,6 +25,7 @@ open Glob_ops
open Retyping
open Pretype_errors
open Evarutil
+open Evardefine
open Evarsolve
open Evarconv
open Evd
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 57b273d0d..9d0f391e4 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -245,7 +245,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
| Lambda (n, t, t') -> c, t'
(*| Prod (n, t, t') -> t'*)
| Evar (k, args) ->
- let (evs, t) = Evarutil.define_evar_as_lambda env !evdref (k,args) in
+ let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in
evdref := evs;
let (n, dom, rng) = destLambda t in
let dom = whd_evar !evdref dom in
@@ -375,7 +375,7 @@ let inh_app_fun_core env evd j =
match kind_of_term t with
| Prod (_,_,_) -> (evd,j)
| Evar ev ->
- let (evd',t) = define_evar_as_product evd ev in
+ let (evd',t) = Evardefine.define_evar_as_product evd ev in
(evd',{ uj_val = j.uj_val; uj_type = t })
| _ ->
try let t,p =
@@ -416,7 +416,7 @@ let inh_coerce_to_sort loc env evd j =
match kind_of_term typ with
| Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s })
| Evar ev when not (is_defined evd (fst ev)) ->
- let (evd',s) = define_evar_as_sort env evd ev in
+ let (evd',s) = Evardefine.define_evar_as_sort env evd ev in
(evd',{ utj_val = j.uj_val; utj_type = s })
| _ ->
inh_tosort_force loc env evd j
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 489a8a729..08973a05c 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -18,6 +18,7 @@ open Termops
open Environ
open Recordops
open Evarutil
+open Evardefine
open Evarsolve
open Globnames
open Evd
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
new file mode 100644
index 000000000..ef3a3f525
--- /dev/null
+++ b/pretyping/evardefine.ml
@@ -0,0 +1,209 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Errors
+open Util
+open Pp
+open Names
+open Term
+open Vars
+open Termops
+open Namegen
+open Pre_env
+open Environ
+open Evd
+open Evarutil
+open Pretype_errors
+open Sigma.Notations
+
+let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal typ =
+ let evd = Sigma.Unsafe.of_evar_map evd in
+ let Sigma (evk, evd, _) = new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ in
+ (Sigma.to_evar_map evd, evk)
+
+let env_nf_evar sigma env =
+ let open Context.Rel.Declaration in
+ process_rel_context
+ (fun d e -> push_rel (map_constr (nf_evar sigma) d) e) env
+
+let env_nf_betaiotaevar sigma env =
+ let open Context.Rel.Declaration in
+ process_rel_context
+ (fun d e ->
+ push_rel (map_constr (Reductionops.nf_betaiota sigma) d) e) env
+
+(****************************************)
+(* Operations on value/type constraints *)
+(****************************************)
+
+type type_constraint = types option
+
+type val_constraint = constr option
+
+(* Old comment...
+ * Basically, we have the following kind of constraints (in increasing
+ * strength order):
+ * (false,(None,None)) -> no constraint at all
+ * (true,(None,None)) -> we must build a judgement which _TYPE is a kind
+ * (_,(None,Some ty)) -> we must build a judgement which _TYPE is ty
+ * (_,(Some v,_)) -> we must build a judgement which _VAL is v
+ * Maybe a concrete datatype would be easier to understand.
+ * We differentiate (true,(None,None)) from (_,(None,Some Type))
+ * because otherwise Case(s) would be misled, as in
+ * (n:nat) Case n of bool [_]nat end would infer the predicate Type instead
+ * of Set.
+ *)
+
+(* The empty type constraint *)
+let empty_tycon = None
+
+(* Builds a type constraint *)
+let mk_tycon ty = Some ty
+
+(* Constrains the value of a type *)
+let empty_valcon = None
+
+(* Builds a value constraint *)
+let mk_valcon c = Some c
+
+let idx = Namegen.default_dependent_ident
+
+(* Refining an evar to a product *)
+
+let define_pure_evar_as_product evd evk =
+ let open Context.Named.Declaration in
+ let evi = Evd.find_undefined evd evk in
+ let evenv = evar_env evi in
+ let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in
+ let concl = Reductionops.whd_betadeltaiota evenv evd evi.evar_concl in
+ let s = destSort concl in
+ let evd1,(dom,u1) =
+ let evd = Sigma.Unsafe.of_evar_map evd in
+ let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in
+ (Sigma.to_evar_map evd1, e)
+ in
+ let evd2,rng =
+ let newenv = push_named (LocalAssum (id, dom)) evenv in
+ let src = evar_source evk evd1 in
+ let filter = Filter.extend 1 (evar_filter evi) in
+ if is_prop_sort s then
+ (* Impredicative product, conclusion must fall in [Prop]. *)
+ new_evar_unsafe newenv evd1 concl ~src ~filter
+ else
+ let status = univ_flexible_alg in
+ let evd3, (rng, srng) =
+ let evd1 = Sigma.Unsafe.of_evar_map evd1 in
+ let Sigma (e, evd3, _) = new_type_evar newenv evd1 status ~src ~filter in
+ (Sigma.to_evar_map evd3, e)
+ in
+ let prods = Univ.sup (univ_of_sort u1) (univ_of_sort srng) in
+ let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in
+ evd3, rng
+ in
+ let prod = mkProd (Name id, dom, subst_var id rng) in
+ let evd3 = Evd.define evk prod evd2 in
+ evd3,prod
+
+(* Refine an applied evar to a product and returns its instantiation *)
+
+let define_evar_as_product evd (evk,args) =
+ let evd,prod = define_pure_evar_as_product evd evk in
+ (* Quick way to compute the instantiation of evk with args *)
+ let na,dom,rng = destProd prod in
+ let evdom = mkEvar (fst (destEvar dom), args) in
+ let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
+ let evrng = mkEvar (fst (destEvar rng), evrngargs) in
+ evd,mkProd (na, evdom, evrng)
+
+(* Refine an evar with an abstraction
+
+ I.e., solve x1..xq |- ?e:T(x1..xq) with e:=λy:A.?e'[x1..xq,y] where:
+ - either T(x1..xq) = πy:A(x1..xq).B(x1..xq,y)
+ or T(x1..xq) = ?d[x1..xq] and we define ?d := πy:?A.?B
+ with x1..xq |- ?A:Type and x1..xq,y |- ?B:Type
+ - x1..xq,y:A |- ?e':B
+*)
+
+let define_pure_evar_as_lambda env evd evk =
+ let open Context.Named.Declaration in
+ let evi = Evd.find_undefined evd evk in
+ let evenv = evar_env evi in
+ let typ = Reductionops.whd_betadeltaiota evenv evd (evar_concl evi) in
+ let evd1,(na,dom,rng) = match kind_of_term typ with
+ | Prod (na,dom,rng) -> (evd,(na,dom,rng))
+ | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ
+ | _ -> error_not_product_loc Loc.ghost env evd typ in
+ let avoid = ids_of_named_context (evar_context evi) in
+ let id =
+ next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in
+ let newenv = push_named (LocalAssum (id, dom)) evenv in
+ let filter = Filter.extend 1 (evar_filter evi) in
+ let src = evar_source evk evd1 in
+ let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in
+ let lam = mkLambda (Name id, dom, subst_var id body) in
+ Evd.define evk lam evd2, lam
+
+let define_evar_as_lambda env evd (evk,args) =
+ let evd,lam = define_pure_evar_as_lambda env evd evk in
+ (* Quick way to compute the instantiation of evk with args *)
+ let na,dom,body = destLambda lam in
+ let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
+ let evbody = mkEvar (fst (destEvar body), evbodyargs) in
+ evd,mkLambda (na, dom, evbody)
+
+let rec evar_absorb_arguments env evd (evk,args as ev) = function
+ | [] -> evd,ev
+ | a::l ->
+ (* TODO: optimize and avoid introducing intermediate evars *)
+ let evd,lam = define_pure_evar_as_lambda env evd evk in
+ let _,_,body = destLambda lam in
+ let evk = fst (destEvar body) in
+ evar_absorb_arguments env evd (evk, Array.cons a args) l
+
+(* Refining an evar to a sort *)
+
+let define_evar_as_sort env evd (ev,args) =
+ let evd, u = new_univ_variable univ_rigid evd in
+ let evi = Evd.find_undefined evd ev in
+ let s = Type u in
+ let concl = Reductionops.whd_betadeltaiota (evar_env evi) evd evi.evar_concl in
+ let sort = destSort concl in
+ let evd' = Evd.define ev (mkSort s) evd in
+ Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s
+
+(* Propagation of constraints through application and abstraction:
+ Given a type constraint on a functional term, returns the type
+ constraint on its domain and codomain. If the input constraint is
+ an evar instantiate it with the product of 2 new evars. *)
+
+let split_tycon loc env evd tycon =
+ let rec real_split evd c =
+ let t = Reductionops.whd_betadeltaiota env evd c in
+ match kind_of_term t with
+ | Prod (na,dom,rng) -> evd, (na, dom, rng)
+ | Evar ev (* ev is undefined because of whd_betadeltaiota *) ->
+ let (evd',prod) = define_evar_as_product evd ev in
+ let (_,dom,rng) = destProd prod in
+ evd',(Anonymous, dom, rng)
+ | App (c,args) when isEvar c ->
+ let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in
+ real_split evd' (mkApp (lam,args))
+ | _ -> error_not_product_loc loc env evd c
+ in
+ match tycon with
+ | None -> evd,(Anonymous,None,None)
+ | Some c ->
+ let evd', (n, dom, rng) = real_split evd c in
+ evd', (n, mk_tycon dom, mk_tycon rng)
+
+let valcon_of_tycon x = x
+let lift_tycon n = Option.map (lift n)
+
+let pr_tycon env = function
+ None -> str "None"
+ | Some t -> Termops.print_constr_env env t
diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli
new file mode 100644
index 000000000..07b0e69d9
--- /dev/null
+++ b/pretyping/evardefine.mli
@@ -0,0 +1,46 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Names
+open Term
+open Evd
+open Environ
+
+val env_nf_evar : evar_map -> env -> env
+val env_nf_betaiotaevar : evar_map -> env -> env
+
+type type_constraint = types option
+type val_constraint = constr option
+
+val empty_tycon : type_constraint
+val mk_tycon : constr -> type_constraint
+val empty_valcon : val_constraint
+val mk_valcon : constr -> val_constraint
+
+(** Instantiate an evar by as many lambda's as needed so that its arguments
+ are moved to the evar substitution (i.e. turn [?x[vars1:=args1] args] into
+ [?y[vars1:=args1,vars:=args]] with
+ [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *)
+val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
+ evar_map * existential
+
+val split_tycon :
+ Loc.t -> env -> evar_map -> type_constraint ->
+ evar_map * (Name.t * type_constraint * type_constraint)
+
+val valcon_of_tycon : type_constraint -> val_constraint
+val lift_tycon : int -> type_constraint -> type_constraint
+
+val define_evar_as_product : evar_map -> existential -> evar_map * types
+val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types
+val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
+
+(** {6 debug pretty-printer:} *)
+
+val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
+
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 8329de2ee..a765d3091 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -36,6 +36,7 @@ open Typeops
open Globnames
open Nameops
open Evarutil
+open Evardefine
open Pretype_errors
open Glob_term
open Glob_ops
@@ -234,6 +235,23 @@ let check_extra_evars_are_solved env current_sigma pending =
| _ ->
error_unsolvable_implicit loc env current_sigma evk None) pending
+(* [check_evars] fails if some unresolved evar remains *)
+
+let check_evars env initial_sigma sigma c =
+ let rec proc_rec c =
+ match kind_of_term c with
+ | Evar (evk,_ as ev) ->
+ (match existential_opt_value sigma ev with
+ | Some c -> proc_rec c
+ | None ->
+ if not (Evd.mem initial_sigma evk) then
+ let (loc,k) = evar_source evk sigma in
+ match k with
+ | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
+ | _ -> Pretype_errors.error_unsolvable_implicit loc env sigma evk None)
+ | _ -> Constr.iter proc_rec c
+ in proc_rec c
+
let check_evars_are_solved env current_sigma frozen pending =
check_typeclasses_instances_are_solved env current_sigma frozen;
check_problems_are_solved env current_sigma;
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 40745ed09..4c4c535d8 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -130,6 +130,10 @@ val solve_remaining_evars : inference_flags ->
val check_evars_are_solved :
env -> (* current map: *) evar_map -> (* map to check: *) pending -> unit
+(** [check_evars env initial_sigma extended_sigma c] fails if some
+ new unresolved evar remains in [c] *)
+val check_evars : env -> evar_map -> evar_map -> constr -> unit
+
(**/**)
(** Internal of Pretyping... *)
val pretype :
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index b59589bda..c8b3307d7 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -1,4 +1,5 @@
Locusops
+Pretype_errors
Reductionops
Inductiveops
Vnorm
@@ -6,9 +7,8 @@ Arguments_renaming
Nativenorm
Retyping
Cbv
-Pretype_errors
Find_subterm
-Evarutil
+Evardefine
Evarsolve
Recordops
Evarconv
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 935e18d8d..7f4249c5b 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -594,9 +594,7 @@ let pr_state (tm,sk) =
(*** Reduction Functions Operators ***)
(*************************************)
-let safe_evar_value sigma ev =
- try Some (Evd.existential_value sigma ev)
- with NotInstantiatedEvar | Not_found -> None
+let safe_evar_value = Evarutil.safe_evar_value
let safe_meta_value sigma ev =
try Some (Evd.meta_value sigma ev)
@@ -1183,30 +1181,8 @@ let whd_zeta c = Stack.zip (local_whd_state_gen zeta Evd.empty (c,Stack.empty))
(****************************************************************************)
(* Replacing defined evars for error messages *)
-let rec whd_evar sigma c =
- match kind_of_term c with
- | Evar ev ->
- let (evk, args) = ev in
- let args = Array.map (fun c -> whd_evar sigma c) args in
- (match safe_evar_value sigma (evk, args) with
- Some c -> whd_evar sigma c
- | None -> c)
- | Sort (Type u) ->
- let u' = Evd.normalize_universe sigma u in
- if u' == u then c else mkSort (Sorts.sort_of_univ u')
- | Const (c', u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkConstU (c', u')
- | Ind (i, u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkIndU (i, u')
- | Construct (co, u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkConstructU (co, u')
- | _ -> c
-
-let nf_evar =
- local_strong whd_evar
+let whd_evar = Evarutil.whd_evar
+let nf_evar = Evarutil.nf_evar
(* lazy reduction functions. The infos must be created for each term *)
(* Note by HH [oct 08] : why would it be the job of clos_norm_flags to add
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 0faa35c87..3a5796fe1 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -490,7 +490,7 @@ let is_instance = function
Nota: we will only check the resolvability status of undefined evars.
*)
-let resolvable = Store.field ()
+let resolvable = Proofview.Unsafe.typeclass_resolvable
let set_resolvable s b =
if b then Store.remove s resolvable
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 5347d965b..52afa7f83 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -39,7 +39,7 @@ let e_type_judgment env evdref j =
match kind_of_term (whd_betadeltaiota env !evdref j.uj_type) with
| Sort s -> {utj_val = j.uj_val; utj_type = s }
| Evar ev ->
- let (evd,s) = Evarutil.define_evar_as_sort env !evdref ev in
+ let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in
evdref := evd; { utj_val = j.uj_val; utj_type = s }
| _ -> error_not_type env j
@@ -61,7 +61,7 @@ let e_judge_of_apply env evdref funj argjv =
else
error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv
| Evar ev ->
- let (evd',t) = Evarutil.define_evar_as_product !evdref ev in
+ let (evd',t) = Evardefine.define_evar_as_product !evdref ev in
evdref := evd';
let (_,_,c2) = destProd t in
apply_rec (n+1) (subst1 hj.uj_val c2) restjl
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index a7b415552..a4a386530 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -19,6 +19,7 @@ open Evd
open Reduction
open Reductionops
open Evarutil
+open Evardefine
open Evarsolve
open Pretype_errors
open Retyping
diff --git a/printing/printer.ml b/printing/printer.ml
index b89005887..2e67fa5ff 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -50,7 +50,7 @@ let pr_lconstr_core goal_concl_style env sigma t =
let pr_lconstr_env env = pr_lconstr_core false env
let pr_constr_env env = pr_constr_core false env
-let _ = Hook.set Proofview.Refine.pr_constr pr_constr_env
+let _ = Hook.set Refine.pr_constr pr_constr_env
let pr_lconstr_goal_style_env env = pr_lconstr_core true env
let pr_constr_goal_style_env env = pr_constr_core true env
diff --git a/proofs/proof.ml b/proofs/proof.ml
index b604fde4e..86af420dc 100644
--- a/proofs/proof.ml
+++ b/proofs/proof.ml
@@ -334,22 +334,24 @@ let compact p =
(*** Tactics ***)
let run_tactic env tac pr =
+ let open Proofview.Notations in
let sp = pr.proofview in
- let (_,tacticced_proofview,(status,to_shelve,give_up),info_trace) =
- Proofview.apply env tac sp
+ let undef sigma l = List.filter (fun g -> Evd.is_undefined sigma g) l in
+ let tac =
+ tac >>= fun () ->
+ Proofview.tclEVARMAP >>= fun sigma ->
+ (* Already solved goals are not to be counted as shelved. Nor are
+ they to be marked as unresolvable. *)
+ let retrieved = undef sigma (List.rev (Evd.future_goals sigma)) in
+ let sigma = List.fold_left Proofview.Unsafe.mark_as_goal sigma retrieved in
+ Proofview.Unsafe.tclEVARS sigma >>= fun () ->
+ Proofview.tclUNIT retrieved
in
- let sigma = Proofview.return tacticced_proofview in
- (* Already solved goals are not to be counted as shelved. Nor are
- they to be marked as unresolvable. *)
- let undef l = List.filter (fun g -> Evd.is_undefined sigma g) l in
- let retrieved = undef (List.rev (Evd.future_goals sigma)) in
- let shelf = (undef pr.shelf)@retrieved@(undef to_shelve) in
- let proofview =
- List.fold_left
- Proofview.Unsafe.mark_as_goal
- tacticced_proofview
- retrieved
+ let (retrieved,proofview,(status,to_shelve,give_up),info_trace) =
+ Proofview.apply env tac sp
in
+ let sigma = Proofview.return proofview in
+ let shelf = (undef sigma pr.shelf)@retrieved@(undef sigma to_shelve) in
let given_up = pr.given_up@give_up in
let proofview = Proofview.Unsafe.reset_future_goals proofview in
{ pr with proofview ; shelf ; given_up },(status,info_trace)
diff --git a/proofs/proofs.mllib b/proofs/proofs.mllib
index 08556d62e..9130a186b 100644
--- a/proofs/proofs.mllib
+++ b/proofs/proofs.mllib
@@ -4,7 +4,7 @@ Evar_refiner
Proof_using
Proof_errors
Logic
-Proofview
+Refine
Proof
Proof_global
Redexpr
diff --git a/proofs/refine.ml b/proofs/refine.ml
new file mode 100644
index 000000000..db0b26f7e
--- /dev/null
+++ b/proofs/refine.ml
@@ -0,0 +1,122 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Util
+open Proofview_monad
+open Sigma.Notations
+open Proofview.Notations
+open Context.Named.Declaration
+
+let extract_prefix env info =
+ let ctx1 = List.rev (Environ.named_context env) in
+ let ctx2 = List.rev (Evd.evar_context info) in
+ let rec share l1 l2 accu = match l1, l2 with
+ | d1 :: l1, d2 :: l2 ->
+ if d1 == d2 then share l1 l2 (d1 :: accu)
+ else (accu, d2 :: l2)
+ | _ -> (accu, l2)
+ in
+ share ctx1 ctx2 []
+
+let typecheck_evar ev env sigma =
+ let info = Evd.find sigma ev in
+ (** Typecheck the hypotheses. *)
+ let type_hyp (sigma, env) decl =
+ let t = get_type decl in
+ let evdref = ref sigma in
+ let _ = Typing.e_sort_of env evdref t in
+ let () = match decl with
+ | LocalAssum _ -> ()
+ | LocalDef (_,body,_) -> Typing.e_check env evdref body t
+ in
+ (!evdref, Environ.push_named decl env)
+ in
+ let (common, changed) = extract_prefix env info in
+ let env = Environ.reset_with_named_context (Environ.val_of_named_context common) env in
+ let (sigma, env) = List.fold_left type_hyp (sigma, env) changed in
+ (** Typecheck the conclusion *)
+ let evdref = ref sigma in
+ let _ = Typing.e_sort_of env evdref (Evd.evar_concl info) in
+ !evdref
+
+let typecheck_proof c concl env sigma =
+ let evdref = ref sigma in
+ let () = Typing.e_check env evdref c concl in
+ !evdref
+
+let (pr_constrv,pr_constr) =
+ Hook.make ~default:(fun _env _sigma _c -> Pp.str"<constr>") ()
+
+let refine ?(unsafe = true) f = Proofview.Goal.enter { enter = begin fun gl ->
+ let gl = Proofview.Goal.assume gl in
+ let sigma = Proofview.Goal.sigma gl in
+ let sigma = Sigma.to_evar_map sigma in
+ let env = Proofview.Goal.env gl in
+ let concl = Proofview.Goal.concl gl in
+ (** Save the [future_goals] state to restore them after the
+ refinement. *)
+ let prev_future_goals = Evd.future_goals sigma in
+ let prev_principal_goal = Evd.principal_future_goal sigma in
+ (** Create the refinement term *)
+ let (c, sigma) = Sigma.run (Evd.reset_future_goals sigma) f in
+ let evs = Evd.future_goals sigma in
+ let evkmain = Evd.principal_future_goal sigma in
+ (** Check that the introduced evars are well-typed *)
+ let fold accu ev = typecheck_evar ev env accu in
+ let sigma = if unsafe then sigma else CList.fold_left fold sigma evs in
+ (** Check that the refined term is typesafe *)
+ let sigma = if unsafe then sigma else typecheck_proof c concl env sigma in
+ (** Check that the goal itself does not appear in the refined term *)
+ let self = Proofview.Goal.goal gl in
+ let _ =
+ if not (Evarutil.occur_evar_upto sigma self c) then ()
+ else Pretype_errors.error_occur_check env sigma self c
+ in
+ (** Proceed to the refinement *)
+ let sigma = match evkmain with
+ | None -> Evd.define self c sigma
+ | Some evk ->
+ let id = Evd.evar_ident self sigma in
+ let sigma = Evd.define self c sigma in
+ match id with
+ | None -> sigma
+ | Some id -> Evd.rename evk id sigma
+ in
+ (** Restore the [future goals] state. *)
+ let sigma = Evd.restore_future_goals sigma prev_future_goals prev_principal_goal in
+ (** Select the goals *)
+ let comb = CList.map_filter (Proofview.Unsafe.advance sigma) (CList.rev evs) in
+ let sigma = CList.fold_left Proofview.Unsafe.mark_as_goal sigma comb in
+ let trace () = Pp.(hov 2 (str"refine"++spc()++ Hook.get pr_constrv env sigma c)) in
+ Proofview.Trace.name_tactic trace (Proofview.tclUNIT ()) >>= fun () ->
+ Proofview.Unsafe.tclEVARS sigma >>= fun () ->
+ Proofview.Unsafe.tclSETGOALS comb
+end }
+
+(** Useful definitions *)
+
+let with_type env evd c t =
+ let my_type = Retyping.get_type_of env evd c in
+ let j = Environ.make_judge c my_type in
+ let (evd,j') =
+ Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t
+ in
+ evd , j'.Environ.uj_val
+
+let refine_casted ?unsafe f = Proofview.Goal.enter { enter = begin fun gl ->
+ let gl = Proofview.Goal.assume gl in
+ let concl = Proofview.Goal.concl gl in
+ let env = Proofview.Goal.env gl in
+ let f = { run = fun h ->
+ let Sigma (c, h, p) = f.run h in
+ let sigma, c = with_type env (Sigma.to_evar_map h) c concl in
+ Sigma (c, Sigma.Unsafe.of_evar_map sigma, p)
+ } in
+ refine ?unsafe f
+end }
diff --git a/proofs/refine.mli b/proofs/refine.mli
new file mode 100644
index 000000000..17c7e28ca
--- /dev/null
+++ b/proofs/refine.mli
@@ -0,0 +1,37 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Proofview
+
+(** {6 The refine tactic} *)
+
+(** Printer used to print the constr which refine refines. *)
+val pr_constr :
+ (Environ.env -> Evd.evar_map -> Term.constr -> Pp.std_ppcmds) Hook.t
+
+(** {7 Refinement primitives} *)
+
+val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
+(** In [refine ?unsafe t], [t] is a term with holes under some
+ [evar_map] context. The term [t] is used as a partial solution
+ for the current goal (refine is a goal-dependent tactic), the
+ new holes created by [t] become the new subgoals. Exceptions
+ raised during the interpretation of [t] are caught and result in
+ tactic failures. If [unsafe] is [false] (default is [true]) [t] is
+ type-checked beforehand. *)
+
+(** {7 Helper functions} *)
+
+val with_type : Environ.env -> Evd.evar_map ->
+ Term.constr -> Term.types -> Evd.evar_map * Term.constr
+(** [with_type env sigma c t] ensures that [c] is of type [t]
+ inserting a coercion if needed. *)
+
+val refine_casted : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
+(** Like {!refine} except the refined term is coerced to the conclusion of the
+ current goal. *)
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index 0cc796886..23aa8dcb4 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -370,7 +370,7 @@ let refine_tac ist simple c =
let expected_type = Pretyping.OfType concl in
let c = Tacinterp.type_uconstr ~flags ~expected_type ist c in
let update = { run = fun sigma -> c.delayed env sigma } in
- let refine = Proofview.Refine.refine ~unsafe:false update in
+ let refine = Refine.refine ~unsafe:false update in
if simple then refine
else refine <*>
Tactics.New.reduce_after_refine <*>
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 730da147a..e5abad686 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -1072,7 +1072,7 @@ let prepare_hint check (poly,local) env init (sigma,c) =
subst := (evar,mkVar id)::!subst;
mkNamedLambda id t (iter (replace_term evar (mkVar id) c)) in
let c' = iter c in
- if check then Evarutil.check_evars (Global.env()) Evd.empty sigma c';
+ if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c';
let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in
if poly then IsConstr (c', diff)
else if local then IsConstr (c', diff)
diff --git a/tactics/inv.ml b/tactics/inv.ml
index 6841ab0ec..89c6beb32 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -459,7 +459,7 @@ let raw_inversion inv_kind id status names =
in
let refined id =
let prf = mkApp (mkVar id, args) in
- Proofview.Refine.refine { run = fun h -> Sigma (prf, h, Sigma.refl) }
+ Refine.refine { run = fun h -> Sigma (prf, h, Sigma.refl) }
in
let neqns = List.length realargs in
let as_mode = names != None in
diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml
index 67d21886b..fb04bee07 100644
--- a/tactics/rewrite.ml
+++ b/tactics/rewrite.ml
@@ -1539,7 +1539,7 @@ let assert_replacing id newt tac =
| d :: rem -> insert_dependent env (LocalAssum (get_id d, newt)) [] after @ rem
in
let env' = Environ.reset_with_named_context (val_of_named_context nc) env in
- Proofview.Refine.refine ~unsafe:false { run = begin fun sigma ->
+ Refine.refine ~unsafe:false { run = begin fun sigma ->
let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma concl in
let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma newt in
let map d =
@@ -1568,7 +1568,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
let gls = List.rev (Evd.fold_undefined fold undef []) in
match clause, prf with
| Some id, Some p ->
- let tac = Proofview.Refine.refine ~unsafe:false { run = fun h -> Sigma (p, h, Sigma.refl) } <*> Proofview.Unsafe.tclNEWGOALS gls in
+ let tac = Refine.refine ~unsafe:false { run = fun h -> Sigma (p, h, Sigma.refl) } <*> Proofview.Unsafe.tclNEWGOALS gls in
Proofview.Unsafe.tclEVARS undef <*>
assert_replacing id newt tac
| Some id, None ->
@@ -1582,7 +1582,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma newt in
Sigma (mkApp (p, [| ev |]), sigma, q)
end } in
- Proofview.Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls
+ Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls
end }
| None, None ->
Proofview.Unsafe.tclEVARS undef <*>
@@ -1880,7 +1880,7 @@ let build_morphism_signature m =
let morph = e_app_poly env evd PropGlobal.proper_type [| t; sig_; m |] in
let evd = solve_constraints env !evd in
let m = Evarutil.nf_evar evd morph in
- Evarutil.check_evars env Evd.empty evd m; m
+ Pretyping.check_evars env Evd.empty evd m; m
let default_morphism sign m =
let env = Global.env () in
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index f725a0654..7ae178af5 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -163,7 +163,7 @@ let _ =
does not check anything. *)
let unsafe_intro env store decl b =
let open Context.Named.Declaration in
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
let ctx = named_context_val env in
let nctx = push_named_context_val decl ctx in
let inst = List.map (mkVar % get_id) (named_context env) in
@@ -199,7 +199,7 @@ let convert_concl ?(check=true) ty k =
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
let conclty = Proofview.Goal.raw_concl gl in
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
let Sigma ((), sigma, p) =
if check then begin
let sigma = Sigma.to_evar_map sigma in
@@ -222,7 +222,7 @@ let convert_hyp ?(check=true) d =
let store = Proofview.Goal.extra gl in
let sign = convert_hyp check (named_context_val env) sigma d in
let env = reset_with_named_context sign env in
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar env sigma ~principal:true ~store ty
end }
end }
@@ -345,7 +345,7 @@ let rename_hyp repl =
let nconcl = subst concl in
let nctx = Environ.val_of_named_context nhyps in
let instance = List.map (mkVar % get_id) hyps in
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar_instance nctx sigma nconcl ~store instance
end }
end }
@@ -1070,7 +1070,7 @@ let cut c =
let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_of_hyps gl) in
(** Backward compat: normalize [c]. *)
let c = if normalize_cut then local_strong whd_betaiota sigma c else c in
- Proofview.Refine.refine ~unsafe:true { run = begin fun h ->
+ Refine.refine ~unsafe:true { run = begin fun h ->
let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in
let Sigma (x, h, q) = Evarutil.new_evar env h c in
let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in
@@ -1736,7 +1736,7 @@ let cut_and_apply c =
| Prod (_,c1,c2) when not (dependent (mkRel 1) c2) ->
let concl = Proofview.Goal.concl gl in
let env = Tacmach.New.pf_env gl in
- Proofview.Refine.refine { run = begin fun sigma ->
+ Refine.refine { run = begin fun sigma ->
let typ = mkProd (Anonymous, c2, concl) in
let Sigma (f, sigma, p) = Evarutil.new_evar env sigma typ in
let Sigma (x, sigma, q) = Evarutil.new_evar env sigma c1 in
@@ -1757,7 +1757,7 @@ let cut_and_apply c =
(* let refine_no_check = Profile.profile2 refine_no_checkkey refine_no_check *)
let new_exact_no_check c =
- Proofview.Refine.refine ~unsafe:true { run = fun h -> Sigma.here c h }
+ Refine.refine ~unsafe:true { run = fun h -> Sigma.here c h }
let exact_check c =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
@@ -1808,7 +1808,7 @@ let assumption =
in
if is_same_type then
(Proofview.Unsafe.tclEVARS sigma) <*>
- Proofview.Refine.refine ~unsafe:true { run = fun h -> Sigma.here (mkVar (get_id decl)) h }
+ Refine.refine ~unsafe:true { run = fun h -> Sigma.here (mkVar (get_id decl)) h }
else arec gl only_eq rest
in
let assumption_tac = { enter = begin fun gl ->
@@ -1893,7 +1893,7 @@ let clear_body ids =
check_is_type env concl msg
in
check_hyps <*> check_concl <*>
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar env sigma ~principal:true concl
end }
end }
@@ -1950,7 +1950,7 @@ let apply_type newcl args =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- Proofview.Refine.refine { run = begin fun sigma ->
+ Refine.refine { run = begin fun sigma ->
let newcl = nf_betaiota (Sigma.to_evar_map sigma) newcl (* As in former Logic.refine *) in
let Sigma (ev, sigma, p) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
@@ -1971,7 +1971,7 @@ let bring_hyps hyps =
let concl = Tacmach.New.pf_nf_concl gl in
let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in
let args = Array.of_list (Context.Named.to_instance hyps) in
- Proofview.Refine.refine { run = begin fun sigma ->
+ Refine.refine { run = begin fun sigma ->
let Sigma (ev, sigma, p) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
Sigma (mkApp (ev, args), sigma, p)
@@ -2671,7 +2671,7 @@ let new_generalize_gen_let lconstr =
0 lconstr (concl, sigma, [])
in
let tac =
- Proofview.Refine.refine { run = begin fun sigma ->
+ Refine.refine { run = begin fun sigma ->
let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true newcl in
Sigma ((applist (ev, args)), sigma, p)
end }
@@ -3325,7 +3325,7 @@ let mk_term_eq env sigma ty t ty' t' =
let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
let open Context.Rel.Declaration in
- Proofview.Refine.refine { run = begin fun sigma ->
+ Refine.refine { run = begin fun sigma ->
let eqslen = List.length eqs in
(* Abstract by the "generalized" hypothesis equality proof if necessary. *)
let abshypeq, abshypt =
@@ -4095,6 +4095,10 @@ let check_enough_applied env sigma elim =
(* Last argument is supposed to be the induction argument *)
check_expected_type env sigma elimc elimt
+let guard_no_unifiable = Proofview.guard_no_unifiable >>= function
+| None -> Proofview.tclUNIT ()
+| Some l -> Proofview.tclZERO (RefinerError (UnresolvedBindings l))
+
let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
id ((pending,(c0,lbind)),(eqname,names)) t0 inhyps cls tac =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
@@ -4122,14 +4126,14 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
(* and destruct has side conditions first *)
Tacticals.New.tclTHENLAST)
(Tacticals.New.tclTHENLIST [
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
let b = not with_evars && with_eq != None in
let Sigma (c, sigma, p) = use_bindings env sigma elim b (c0,lbind) t0 in
let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in
let Sigma (ans, sigma, q) = mkletin_goal env sigma store with_eq false (id,lastlhyp,ccl,c) (Some t) in
Sigma (ans, sigma, p +> q)
end };
- Proofview.(if with_evars then shelve_unifiable else guard_no_unifiable);
+ if with_evars then Proofview.shelve_unifiable else guard_no_unifiable;
if is_arg_pure_hyp
then Tacticals.New.tclTRY (Proofview.V82.tactic (thin [destVar c0]))
else Proofview.tclUNIT ();
@@ -4146,7 +4150,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
let env = reset_with_named_context sign env in
let tac =
Tacticals.New.tclTHENLIST [
- Proofview.Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
mkletin_goal env sigma store with_eq true (id,lastlhyp,ccl,c) None
end };
tac
@@ -4791,6 +4795,6 @@ module New = struct
{onhyps=None; concl_occs=AllOccurrences }
let refine ?unsafe c =
- Proofview.Refine.refine ?unsafe c <*>
+ Refine.refine ?unsafe c <*>
reduce_after_refine
end
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 26ea01769..4c4a96ec0 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -431,7 +431,7 @@ end
module New : sig
val refine : ?unsafe:bool -> constr Sigma.run -> unit Proofview.tactic
- (** [refine ?unsafe c] is [Proofview.Refine.refine ?unsafe c]
+ (** [refine ?unsafe c] is [Refine.refine ?unsafe c]
followed by beta-iota-reduction of the conclusion. *)
val reduce_after_refine : unit Proofview.tactic
diff --git a/toplevel/classes.ml b/toplevel/classes.ml
index 0a83c49c8..2fc0f5ff1 100644
--- a/toplevel/classes.ml
+++ b/toplevel/classes.ml
@@ -191,7 +191,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
nf t
in
- Evarutil.check_evars env Evd.empty !evars termtype;
+ Pretyping.check_evars env Evd.empty !evars termtype;
let pl, ctx = Evd.universe_context ?names:pl !evars in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id
(ParameterEntry
@@ -287,7 +287,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
let evm, nf = Evarutil.nf_evar_map_universes !evars in
let termtype = nf termtype in
let _ = (* Check that the type is free of evars now. *)
- Evarutil.check_evars env Evd.empty evm termtype
+ Pretyping.check_evars env Evd.empty evm termtype
in
let term = Option.map nf term in
if not (Evd.has_undefined evm) && not (Option.is_empty term) then
@@ -330,7 +330,7 @@ let new_instance ?(abstract=false) ?(global=false) poly ctx (instid, bk, cl) pro
if not (Option.is_empty term) then
let init_refine =
Tacticals.New.tclTHENLIST [
- Proofview.Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) };
+ Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) };
Proofview.Unsafe.tclNEWGOALS gls;
Tactics.New.reduce_after_refine;
]
@@ -361,7 +361,7 @@ let context poly l =
let _, ((env', fullctx), impls) = interp_context_evars env evars l in
let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
let fullctx = Context.Rel.map subst fullctx in
- let ce t = Evarutil.check_evars env Evd.empty !evars t in
+ let ce t = Pretyping.check_evars env Evd.empty !evars t in
let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml
index 4ee1537c2..ad848ccfc 100644
--- a/toplevel/himsg.ml
+++ b/toplevel/himsg.ml
@@ -76,6 +76,15 @@ let rec contract3' env a b c = function
let y,x = contract3' env a b c x in
y,CannotSolveConstraint ((pb,env,t,u),x)
+(** Ad-hoc reductions *)
+
+let j_nf_betaiotaevar sigma j =
+ { uj_val = Evarutil.nf_evar sigma j.uj_val;
+ uj_type = Reductionops.nf_betaiota sigma j.uj_type }
+
+let jv_nf_betaiotaevar sigma jl =
+ Array.map (j_nf_betaiotaevar sigma) jl
+
(** Printers *)
let pr_lconstr c = quote (pr_lconstr c)
@@ -322,7 +331,7 @@ let explain_unification_error env sigma p1 p2 = function
let explain_actual_type env sigma j t reason =
let env = make_all_name_different env in
- let j = Evarutil.j_nf_betaiotaevar sigma j in
+ let j = j_nf_betaiotaevar sigma j in
let t = Reductionops.nf_betaiota sigma t in
(** Actually print *)
let pe = pr_ne_context_of (str "In environment") env sigma in
@@ -337,7 +346,7 @@ let explain_actual_type env sigma j t reason =
ppreason ++ str ".")
let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
- let randl = Evarutil.jv_nf_betaiotaevar sigma randl in
+ let randl = jv_nf_betaiotaevar sigma randl in
let exptyp = Evarutil.nf_evar sigma exptyp in
let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
let rator = Evarutil.j_nf_evar sigma rator in
@@ -779,7 +788,7 @@ let explain_unsatisfiable_constraints env sigma constr comp =
explain_typeclass_resolution env sigma info k ++ fnl () ++ cstr
let explain_pretype_error env sigma err =
- let env = Evarutil.env_nf_betaiotaevar sigma env in
+ let env = Evardefine.env_nf_betaiotaevar sigma env in
let env = make_all_name_different env in
match err with
| CantFindCaseType c -> explain_cant_find_case_type env sigma c
diff --git a/toplevel/record.ml b/toplevel/record.ml
index db82c205c..96cafb072 100644
--- a/toplevel/record.ml
+++ b/toplevel/record.ml
@@ -156,7 +156,7 @@ let typecheck_params_and_fields def id pl t ps nots fs =
let evars, nf = Evarutil.nf_evars_and_universes evars in
let newps = Context.Rel.map nf newps in
let newfs = Context.Rel.map nf newfs in
- let ce t = Evarutil.check_evars env0 Evd.empty evars t in
+ let ce t = Pretyping.check_evars env0 Evd.empty evars t in
List.iter (iter_constr ce) (List.rev newps);
List.iter (iter_constr ce) (List.rev newfs);
Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs