aboutsummaryrefslogtreecommitdiffhomepage
path: root/engine
diff options
context:
space:
mode:
Diffstat (limited to 'engine')
-rw-r--r--engine/eConstr.ml797
-rw-r--r--engine/eConstr.mli294
-rw-r--r--engine/engine.mllib7
-rw-r--r--engine/evarutil.ml291
-rw-r--r--engine/evarutil.mli72
-rw-r--r--engine/evd.ml419
-rw-r--r--engine/evd.mli42
-rw-r--r--engine/ftactic.ml18
-rw-r--r--engine/ftactic.mli13
-rw-r--r--engine/logic_monad.ml2
-rw-r--r--engine/namegen.ml180
-rw-r--r--engine/namegen.mli61
-rw-r--r--engine/proofview.ml152
-rw-r--r--engine/proofview.mli62
-rw-r--r--engine/sigma.ml117
-rw-r--r--engine/sigma.mli131
-rw-r--r--engine/termops.ml867
-rw-r--r--engine/termops.mli226
-rw-r--r--engine/uState.ml141
-rw-r--r--engine/universes.ml1140
-rw-r--r--engine/universes.mli231
21 files changed, 3732 insertions, 1531 deletions
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
new file mode 100644
index 000000000..7b879a803
--- /dev/null
+++ b/engine/eConstr.ml
@@ -0,0 +1,797 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CErrors
+open Util
+open Names
+open Term
+open Context
+open Evd
+
+module API :
+sig
+module ESorts :
+sig
+type t
+val make : Sorts.t -> t
+val kind : Evd.evar_map -> t -> Sorts.t
+val unsafe_to_sorts : t -> Sorts.t
+end
+module EInstance :
+sig
+type t
+val make : Univ.Instance.t -> t
+val kind : Evd.evar_map -> t -> Univ.Instance.t
+val empty : t
+val is_empty : t -> bool
+val unsafe_to_instance : t -> Univ.Instance.t
+end
+type t
+val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term
+val kind_upto : Evd.evar_map -> constr -> (constr, types, Sorts.t, Univ.Instance.t) Constr.kind_of_term
+val kind_of_type : Evd.evar_map -> t -> (t, t) kind_of_type
+val whd_evar : Evd.evar_map -> t -> t
+val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t
+val of_constr : Constr.t -> t
+val to_constr : evar_map -> t -> Constr.t
+val unsafe_to_constr : t -> Constr.t
+val unsafe_eq : (t, Constr.t) eq
+val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, t) Context.Named.Declaration.pt
+val unsafe_to_named_decl : (t, t) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt
+val unsafe_to_rel_decl : (t, t) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, t) Context.Rel.Declaration.pt
+val to_rel_decl : Evd.evar_map -> (t, t) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+end =
+struct
+
+module ESorts =
+struct
+ type t = Sorts.t
+ let make s = s
+ let kind sigma = function
+ | Type u -> sort_of_univ (Evd.normalize_universe sigma u)
+ | s -> s
+ let unsafe_to_sorts s = s
+end
+
+module EInstance =
+struct
+ type t = Univ.Instance.t
+ let make i = i
+ let kind sigma i =
+ if Univ.Instance.is_empty i then i
+ else Evd.normalize_universe_instance sigma i
+ let empty = Univ.Instance.empty
+ let is_empty = Univ.Instance.is_empty
+ let unsafe_to_instance t = t
+end
+
+type t = Constr.t
+
+let safe_evar_value sigma ev =
+ try Some (Evd.existential_value sigma ev)
+ with NotInstantiatedEvar | Not_found -> None
+
+let rec whd_evar sigma c =
+ match Constr.kind c with
+ | Evar ev ->
+ begin match safe_evar_value sigma ev with
+ | Some c -> whd_evar sigma c
+ | None -> c
+ end
+ | App (f, args) when isEvar f ->
+ (** Enforce smart constructor invariant on applications *)
+ let ev = destEvar f in
+ begin match safe_evar_value sigma ev with
+ | None -> c
+ | Some f -> whd_evar sigma (mkApp (f, args))
+ end
+ | Cast (c0, k, t) when isEvar c0 ->
+ (** Enforce smart constructor invariant on casts. *)
+ let ev = destEvar c0 in
+ begin match safe_evar_value sigma ev with
+ | None -> c
+ | Some c -> whd_evar sigma (mkCast (c, k, t))
+ end
+ | _ -> c
+
+let kind sigma c = Constr.kind (whd_evar sigma c)
+let kind_upto = kind
+let kind_of_type sigma c = Term.kind_of_type (whd_evar sigma c)
+let of_kind = Constr.of_kind
+let of_constr c = c
+let unsafe_to_constr c = c
+let unsafe_eq = Refl
+
+let rec to_constr sigma c = match Constr.kind c with
+| Evar ev ->
+ begin match safe_evar_value sigma ev with
+ | Some c -> to_constr sigma c
+ | None -> Constr.map (fun c -> to_constr sigma c) c
+ end
+| 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')
+| _ -> Constr.map (fun c -> to_constr sigma c) c
+
+let of_named_decl d = d
+let unsafe_to_named_decl d = d
+let of_rel_decl d = d
+let unsafe_to_rel_decl d = d
+let to_rel_decl sigma d = Context.Rel.Declaration.map_constr (to_constr sigma) d
+
+end
+
+include API
+
+type types = t
+type constr = t
+type existential = t pexistential
+type fixpoint = (t, t) pfixpoint
+type cofixpoint = (t, t) pcofixpoint
+type unsafe_judgment = (constr, types) Environ.punsafe_judgment
+type unsafe_type_judgment = types Environ.punsafe_type_judgment
+type named_declaration = (constr, types) Context.Named.Declaration.pt
+type rel_declaration = (constr, types) Context.Rel.Declaration.pt
+type named_context = (constr, types) Context.Named.pt
+type rel_context = (constr, types) Context.Rel.pt
+
+let in_punivs a = (a, EInstance.empty)
+
+let mkProp = of_kind (Sort (ESorts.make Sorts.prop))
+let mkSet = of_kind (Sort (ESorts.make Sorts.set))
+let mkType u = of_kind (Sort (ESorts.make (Sorts.Type u)))
+let mkRel n = of_kind (Rel n)
+let mkVar id = of_kind (Var id)
+let mkMeta n = of_kind (Meta n)
+let mkEvar e = of_kind (Evar e)
+let mkSort s = of_kind (Sort (ESorts.make s))
+let mkCast (b, k, t) = of_kind (Cast (b, k, t))
+let mkProd (na, t, u) = of_kind (Prod (na, t, u))
+let mkLambda (na, t, c) = of_kind (Lambda (na, t, c))
+let mkLetIn (na, b, t, c) = of_kind (LetIn (na, b, t, c))
+let mkApp (f, arg) = of_kind (App (f, arg))
+let mkConstU pc = of_kind (Const pc)
+let mkConst c = of_kind (Const (in_punivs c))
+let mkIndU pi = of_kind (Ind pi)
+let mkInd i = of_kind (Ind (in_punivs i))
+let mkConstructU pc = of_kind (Construct pc)
+let mkConstruct c = of_kind (Construct (in_punivs c))
+let mkConstructUi ((ind,u),i) = of_kind (Construct ((ind,i),u))
+let mkCase (ci, c, r, p) = of_kind (Case (ci, c, r, p))
+let mkFix f = of_kind (Fix f)
+let mkCoFix f = of_kind (CoFix f)
+let mkProj (p, c) = of_kind (Proj (p, c))
+let mkArrow t1 t2 = of_kind (Prod (Anonymous, t1, t2))
+
+let applist (f, arg) = mkApp (f, Array.of_list arg)
+
+let isRel sigma c = match kind sigma c with Rel _ -> true | _ -> false
+let isVar sigma c = match kind sigma c with Var _ -> true | _ -> false
+let isInd sigma c = match kind sigma c with Ind _ -> true | _ -> false
+let isEvar sigma c = match kind sigma c with Evar _ -> true | _ -> false
+let isMeta sigma c = match kind sigma c with Meta _ -> true | _ -> false
+let isSort sigma c = match kind sigma c with Sort _ -> true | _ -> false
+let isCast sigma c = match kind sigma c with Cast _ -> true | _ -> false
+let isApp sigma c = match kind sigma c with App _ -> true | _ -> false
+let isLambda sigma c = match kind sigma c with Lambda _ -> true | _ -> false
+let isLetIn sigma c = match kind sigma c with LetIn _ -> true | _ -> false
+let isProd sigma c = match kind sigma c with Prod _ -> true | _ -> false
+let isConst sigma c = match kind sigma c with Const _ -> true | _ -> false
+let isConstruct sigma c = match kind sigma c with Construct _ -> true | _ -> false
+let isFix sigma c = match kind sigma c with Fix _ -> true | _ -> false
+let isCoFix sigma c = match kind sigma c with CoFix _ -> true | _ -> false
+let isCase sigma c = match kind sigma c with Case _ -> true | _ -> false
+let isProj sigma c = match kind sigma c with Proj _ -> true | _ -> false
+let isVarId sigma id c =
+ match kind sigma c with Var id' -> Id.equal id id' | _ -> false
+let isRelN sigma n c =
+ match kind sigma c with Rel n' -> Int.equal n n' | _ -> false
+
+let destRel sigma c = match kind sigma c with
+| Rel p -> p
+| _ -> raise DestKO
+
+let destVar sigma c = match kind sigma c with
+| Var p -> p
+| _ -> raise DestKO
+
+let destInd sigma c = match kind sigma c with
+| Ind p -> p
+| _ -> raise DestKO
+
+let destEvar sigma c = match kind sigma c with
+| Evar p -> p
+| _ -> raise DestKO
+
+let destMeta sigma c = match kind sigma c with
+| Meta p -> p
+| _ -> raise DestKO
+
+let destSort sigma c = match kind sigma c with
+| Sort p -> p
+| _ -> raise DestKO
+
+let destCast sigma c = match kind sigma c with
+| Cast (c, k, t) -> (c, k, t)
+| _ -> raise DestKO
+
+let destApp sigma c = match kind sigma c with
+| App (f, a) -> (f, a)
+| _ -> raise DestKO
+
+let destLambda sigma c = match kind sigma c with
+| Lambda (na, t, c) -> (na, t, c)
+| _ -> raise DestKO
+
+let destLetIn sigma c = match kind sigma c with
+| LetIn (na, b, t, c) -> (na, b, t, c)
+| _ -> raise DestKO
+
+let destProd sigma c = match kind sigma c with
+| Prod (na, t, c) -> (na, t, c)
+| _ -> raise DestKO
+
+let destConst sigma c = match kind sigma c with
+| Const p -> p
+| _ -> raise DestKO
+
+let destConstruct sigma c = match kind sigma c with
+| Construct p -> p
+| _ -> raise DestKO
+
+let destFix sigma c = match kind sigma c with
+| Fix p -> p
+| _ -> raise DestKO
+
+let destCoFix sigma c = match kind sigma c with
+| CoFix p -> p
+| _ -> raise DestKO
+
+let destCase sigma c = match kind sigma c with
+| Case (ci, t, c, p) -> (ci, t, c, p)
+| _ -> raise DestKO
+
+let destProj sigma c = match kind sigma c with
+| Proj (p, c) -> (p, c)
+| _ -> raise DestKO
+
+let decompose_app sigma c =
+ match kind sigma c with
+ | App (f,cl) -> (f, Array.to_list cl)
+ | _ -> (c,[])
+
+let decompose_lam sigma c =
+ let rec lamdec_rec l c = match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c
+ | Cast (c,_,_) -> lamdec_rec l c
+ | _ -> l,c
+ in
+ lamdec_rec [] c
+
+let decompose_lam_assum sigma c =
+ let open Rel.Declaration in
+ let rec lamdec_rec l c =
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
+ | Cast (c,_,_) -> lamdec_rec l c
+ | _ -> l,c
+ in
+ lamdec_rec Context.Rel.empty c
+
+let decompose_lam_n_assum sigma n c =
+ let open Rel.Declaration in
+ if n < 0 then
+ user_err Pp.(str "decompose_lam_n_assum: integer parameter must be positive");
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> user_err Pp.(str "decompose_lam_n_assum: not enough abstractions")
+ in
+ lamdec_rec Context.Rel.empty n c
+
+let decompose_lam_n_decls sigma n =
+ let open Rel.Declaration in
+ if n < 0 then
+ user_err Pp.(str "decompose_lam_n_decls: integer parameter must be positive");
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> user_err Pp.(str "decompose_lam_n_decls: not enough abstractions")
+ in
+ lamdec_rec Context.Rel.empty n
+
+let lamn n env b =
+ let rec lamrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> lamrec (n-1, l, mkLambda (v,t,b))
+ | _ -> assert false
+ in
+ lamrec (n,env,b)
+
+let compose_lam l b = lamn (List.length l) l b
+
+let rec to_lambda sigma n prod =
+ if Int.equal n 0 then
+ prod
+ else
+ match kind sigma prod with
+ | Prod (na,ty,bd) -> mkLambda (na,ty,to_lambda sigma (n-1) bd)
+ | Cast (c,_,_) -> to_lambda sigma n c
+ | _ -> user_err ~hdr:"to_lambda" (Pp.mt ())
+
+let decompose_prod sigma c =
+ let rec proddec_rec l c = match kind sigma c with
+ | Prod (x,t,c) -> proddec_rec ((x,t)::l) c
+ | Cast (c,_,_) -> proddec_rec l c
+ | _ -> l,c
+ in
+ proddec_rec [] c
+
+let decompose_prod_assum sigma c =
+ let open Rel.Declaration in
+ let rec proddec_rec l c =
+ match kind sigma c with
+ | Prod (x,t,c) -> proddec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> proddec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
+ | Cast (c,_,_) -> proddec_rec l c
+ | _ -> l,c
+ in
+ proddec_rec Context.Rel.empty c
+
+let decompose_prod_n_assum sigma n c =
+ let open Rel.Declaration in
+ if n < 0 then
+ user_err Pp.(str "decompose_prod_n_assum: integer parameter must be positive");
+ let rec prodec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> prodec_rec l n c
+ | c -> user_err Pp.(str "decompose_prod_n_assum: not enough assumptions")
+ in
+ prodec_rec Context.Rel.empty n c
+
+let existential_type sigma (evk, args) =
+ of_constr (existential_type sigma (evk, Array.map unsafe_to_constr args))
+
+let map sigma f c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c
+ | Cast (b,k,t) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkCast (b', k, t')
+ | Prod (na,t,b) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkProd (na, t', b')
+ | Lambda (na,t,b) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkLambda (na, t', b')
+ | LetIn (na,b,t,k) ->
+ let b' = f b in
+ let t' = f t in
+ let k' = f k in
+ if b'==b && t' == t && k'==k then c
+ else mkLetIn (na, b', t', k')
+ | App (b,l) ->
+ let b' = f b in
+ let l' = Array.smartmap f l in
+ if b'==b && l'==l then c
+ else mkApp (b', l')
+ | Proj (p,t) ->
+ let t' = f t in
+ if t' == t then c
+ else mkProj (p, t')
+ | Evar (e,l) ->
+ let l' = Array.smartmap f l in
+ if l'==l then c
+ else mkEvar (e, l')
+ | Case (ci,p,b,bl) ->
+ let b' = f b in
+ let p' = f p in
+ let bl' = Array.smartmap f bl in
+ if b'==b && p'==p && bl'==bl then c
+ else mkCase (ci, p', b', bl')
+ | Fix (ln,(lna,tl,bl)) ->
+ let tl' = Array.smartmap f tl in
+ let bl' = Array.smartmap f bl in
+ if tl'==tl && bl'==bl then c
+ else mkFix (ln,(lna,tl',bl'))
+ | CoFix(ln,(lna,tl,bl)) ->
+ let tl' = Array.smartmap f tl in
+ let bl' = Array.smartmap f bl in
+ if tl'==tl && bl'==bl then c
+ else mkCoFix (ln,(lna,tl',bl'))
+
+let map_with_binders sigma g f l c0 = match kind sigma c0 with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c0
+ | Cast (c, k, t) ->
+ let c' = f l c in
+ let t' = f l t in
+ if c' == c && t' == t then c0
+ else mkCast (c', k, t')
+ | Prod (na, t, c) ->
+ let t' = f l t in
+ let c' = f (g l) c in
+ if t' == t && c' == c then c0
+ else mkProd (na, t', c')
+ | Lambda (na, t, c) ->
+ let t' = f l t in
+ let c' = f (g l) c in
+ if t' == t && c' == c then c0
+ else mkLambda (na, t', c')
+ | LetIn (na, b, t, c) ->
+ let b' = f l b in
+ let t' = f l t in
+ let c' = f (g l) c in
+ if b' == b && t' == t && c' == c then c0
+ else mkLetIn (na, b', t', c')
+ | App (c, al) ->
+ let c' = f l c in
+ let al' = CArray.Fun1.smartmap f l al in
+ if c' == c && al' == al then c0
+ else mkApp (c', al')
+ | Proj (p, t) ->
+ let t' = f l t in
+ if t' == t then c0
+ else mkProj (p, t')
+ | Evar (e, al) ->
+ let al' = CArray.Fun1.smartmap f l al in
+ if al' == al then c0
+ else mkEvar (e, al')
+ | Case (ci, p, c, bl) ->
+ let p' = f l p in
+ let c' = f l c in
+ let bl' = CArray.Fun1.smartmap f l bl in
+ if p' == p && c' == c && bl' == bl then c0
+ else mkCase (ci, p', c', bl')
+ | Fix (ln, (lna, tl, bl)) ->
+ let tl' = CArray.Fun1.smartmap f l tl in
+ let l' = iterate g (Array.length tl) l in
+ let bl' = CArray.Fun1.smartmap f l' bl in
+ if tl' == tl && bl' == bl then c0
+ else mkFix (ln,(lna,tl',bl'))
+ | CoFix(ln,(lna,tl,bl)) ->
+ let tl' = CArray.Fun1.smartmap f l tl in
+ let l' = iterate g (Array.length tl) l in
+ let bl' = CArray.Fun1.smartmap f l' bl in
+ mkCoFix (ln,(lna,tl',bl'))
+
+let iter sigma f c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,_,t) -> f c; f t
+ | Prod (_,t,c) -> f t; f c
+ | Lambda (_,t,c) -> f t; f c
+ | LetIn (_,b,t,c) -> f b; f t; f c
+ | App (c,l) -> f c; Array.iter f l
+ | Proj (p,c) -> f c
+ | Evar (_,l) -> Array.iter f l
+ | Case (_,p,c,bl) -> f p; f c; Array.iter f bl
+ | Fix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
+ | CoFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
+
+let iter_with_full_binders sigma g f n c =
+ let open Context.Rel.Declaration in
+ match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,_,t) -> f n c; f n t
+ | Prod (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c
+ | Lambda (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c
+ | LetIn (na,b,t,c) -> f n b; f n t; f (g (LocalDef (na, b, t)) n) c
+ | App (c,l) -> f n c; CArray.Fun1.iter f n l
+ | Evar (_,l) -> CArray.Fun1.iter f n l
+ | Case (_,p,c,bl) -> f n p; f n c; CArray.Fun1.iter f n bl
+ | Proj (p,c) -> f n c
+ | Fix (_,(lna,tl,bl)) ->
+ Array.iter (f n) tl;
+ let n' = Array.fold_left2 (fun n na t -> g (LocalAssum (na,t)) n) n lna tl in
+ Array.iter (f n') bl
+ | CoFix (_,(lna,tl,bl)) ->
+ Array.iter (f n) tl;
+ let n' = Array.fold_left2 (fun n na t -> g (LocalAssum (na,t)) n) n lna tl in
+ Array.iter (f n') bl
+
+let iter_with_binders sigma g f n c =
+ iter_with_full_binders sigma (fun _ acc -> g acc) f n c
+
+let fold sigma f acc c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> acc
+ | Cast (c,_,t) -> f (f acc c) t
+ | Prod (_,t,c) -> f (f acc t) c
+ | Lambda (_,t,c) -> f (f acc t) c
+ | LetIn (_,b,t,c) -> f (f (f acc b) t) c
+ | App (c,l) -> Array.fold_left f (f acc c) l
+ | Proj (p,c) -> f acc c
+ | Evar (_,l) -> Array.fold_left f acc l
+ | Case (_,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
+ | Fix (_,(lna,tl,bl)) ->
+ Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl
+ | CoFix (_,(lna,tl,bl)) ->
+ Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl
+
+let compare_gen k eq_inst eq_sort eq_constr c1 c2 =
+ (c1 == c2) || Constr.compare_head_gen_with k k eq_inst eq_sort eq_constr c1 c2
+
+let eq_constr sigma c1 c2 =
+ let kind c = kind_upto sigma c in
+ let rec eq_constr c1 c2 =
+ compare_gen kind (fun _ -> Univ.Instance.equal) Sorts.equal eq_constr c1 c2
+ in
+ eq_constr (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+let eq_constr_nounivs sigma c1 c2 =
+ let kind c = kind_upto sigma c in
+ let rec eq_constr c1 c2 =
+ compare_gen kind (fun _ _ _ -> true) (fun _ _ -> true) eq_constr c1 c2
+ in
+ eq_constr (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+let compare_constr sigma cmp c1 c2 =
+ let kind c = kind_upto sigma c in
+ let cmp c1 c2 = cmp (of_constr c1) (of_constr c2) in
+ compare_gen kind (fun _ -> Univ.Instance.equal) Sorts.equal cmp (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+(** TODO: factorize with universes.ml *)
+let test_constr_universes sigma leq m n =
+ let open Universes in
+ let kind c = kind_upto sigma c in
+ if m == n then Some Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,UEq,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let leq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,ULe,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n = compare_gen kind eq_universes eq_sorts eq_constr' m n in
+ let res =
+ if leq then
+ let rec compare_leq m n =
+ Constr.compare_head_gen_leq_with kind kind eq_universes leq_sorts eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ compare_leq m n
+ else
+ Constr.compare_head_gen_with kind kind eq_universes eq_sorts eq_constr' m n
+ in
+ if res then Some !cstrs else None
+
+let eq_constr_universes sigma m n =
+ test_constr_universes sigma false (unsafe_to_constr m) (unsafe_to_constr n)
+let leq_constr_universes sigma m n =
+ test_constr_universes sigma true (unsafe_to_constr m) (unsafe_to_constr n)
+
+let compare_head_gen_proj env sigma equ eqs eqc' m n =
+ let kind c = kind_upto sigma c in
+ match kind_upto sigma m, kind_upto sigma n with
+ | Proj (p, c), App (f, args)
+ | App (f, args), Proj (p, c) ->
+ (match kind_upto sigma f with
+ | Const (p', u) when Constant.equal (Projection.constant p) p' ->
+ let pb = Environ.lookup_projection p env in
+ let npars = pb.Declarations.proj_npars in
+ if Array.length args == npars + 1 then
+ eqc' c args.(npars)
+ else false
+ | _ -> false)
+ | _ -> Constr.compare_head_gen_with kind kind equ eqs eqc' m n
+
+let eq_constr_universes_proj env sigma m n =
+ let open Universes in
+ if m == n then Some Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1, UEq, Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen_proj env sigma eq_universes eq_sorts eq_constr' m n
+ in
+ let res = eq_constr' (unsafe_to_constr m) (unsafe_to_constr n) in
+ if res then Some !cstrs else None
+
+open Context
+open Environ
+
+let cast_list : type a b. (a,b) eq -> a list -> b list =
+ fun Refl x -> x
+
+let cast_list_snd : type a b. (a,b) eq -> ('c * a) list -> ('c * b) list =
+ fun Refl x -> x
+
+let cast_rel_decl :
+ type a b. (a,b) eq -> (a, a) Rel.Declaration.pt -> (b, b) Rel.Declaration.pt =
+ fun Refl x -> x
+
+let cast_rel_context :
+ type a b. (a,b) eq -> (a, a) Rel.pt -> (b, b) Rel.pt =
+ fun Refl x -> x
+
+let cast_named_decl :
+ type a b. (a,b) eq -> (a, a) Named.Declaration.pt -> (b, b) Named.Declaration.pt =
+ fun Refl x -> x
+
+let cast_named_context :
+ type a b. (a,b) eq -> (a, a) Named.pt -> (b, b) Named.pt =
+ fun Refl x -> x
+
+
+module Vars =
+struct
+exception LocalOccur
+let to_constr = unsafe_to_constr
+let to_rel_decl = unsafe_to_rel_decl
+
+type substl = t list
+
+(** Operations that commute with evar-normalization *)
+let lift n c = of_constr (Vars.lift n (to_constr c))
+let liftn n m c = of_constr (Vars.liftn n m (to_constr c))
+
+let substnl subst n c = of_constr (Vars.substnl (cast_list unsafe_eq subst) n (to_constr c))
+let substl subst c = of_constr (Vars.substl (cast_list unsafe_eq subst) (to_constr c))
+let subst1 c r = of_constr (Vars.subst1 (to_constr c) (to_constr r))
+
+let substnl_decl subst n d = of_rel_decl (Vars.substnl_decl (cast_list unsafe_eq subst) n (to_rel_decl d))
+let substl_decl subst d = of_rel_decl (Vars.substl_decl (cast_list unsafe_eq subst) (to_rel_decl d))
+let subst1_decl c d = of_rel_decl (Vars.subst1_decl (to_constr c) (to_rel_decl d))
+
+let replace_vars subst c =
+ of_constr (Vars.replace_vars (cast_list_snd unsafe_eq subst) (to_constr c))
+let substn_vars n subst c = of_constr (Vars.substn_vars n subst (to_constr c))
+let subst_vars subst c = of_constr (Vars.subst_vars subst (to_constr c))
+let subst_var subst c = of_constr (Vars.subst_var subst (to_constr c))
+
+let subst_univs_level_constr subst c =
+ of_constr (Vars.subst_univs_level_constr subst (to_constr c))
+(** Operations that dot NOT commute with evar-normalization *)
+let noccurn sigma n term =
+ let rec occur_rec n c = match kind sigma c with
+ | Rel m -> if Int.equal m n then raise LocalOccur
+ | _ -> iter_with_binders sigma succ occur_rec n c
+ in
+ try occur_rec n term; true with LocalOccur -> false
+
+let noccur_between sigma n m term =
+ let rec occur_rec n c = match kind sigma c with
+ | Rel p -> if n<=p && p<n+m then raise LocalOccur
+ | _ -> iter_with_binders sigma succ occur_rec n c
+ in
+ try occur_rec n term; true with LocalOccur -> false
+
+let closedn sigma n c =
+ let rec closed_rec n c = match kind sigma c with
+ | Rel m -> if m>n then raise LocalOccur
+ | _ -> iter_with_binders sigma succ closed_rec n c
+ in
+ try closed_rec n c; true with LocalOccur -> false
+
+let closed0 sigma c = closedn sigma 0 c
+
+let subst_of_rel_context_instance ctx subst =
+ cast_list (sym unsafe_eq)
+ (Vars.subst_of_rel_context_instance (cast_rel_context unsafe_eq ctx) (cast_list unsafe_eq subst))
+
+end
+
+let rec isArity sigma c =
+ match kind sigma c with
+ | Prod (_,_,c) -> isArity sigma c
+ | LetIn (_,b,_,c) -> isArity sigma (Vars.subst1 b c)
+ | Cast (c,_,_) -> isArity sigma c
+ | Sort _ -> true
+ | _ -> false
+
+let mkProd_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkLambda_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkLambda (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedProd id typ c = mkProd (Name id, typ, Vars.subst_var id c)
+let mkNamedLambda id typ c = mkLambda (Name id, typ, Vars.subst_var id c)
+let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, Vars.subst_var id c2)
+
+let mkNamedProd_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
+
+let mkNamedLambda_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedLambda id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
+
+let it_mkProd_or_LetIn t ctx = List.fold_left (fun c d -> mkProd_or_LetIn d c) t ctx
+let it_mkLambda_or_LetIn t ctx = List.fold_left (fun c d -> mkLambda_or_LetIn d c) t ctx
+
+let push_rel d e = push_rel (cast_rel_decl unsafe_eq d) e
+let push_rel_context d e = push_rel_context (cast_rel_context unsafe_eq d) e
+let push_named d e = push_named (cast_named_decl unsafe_eq d) e
+let push_named_context d e = push_named_context (cast_named_context unsafe_eq d) e
+let push_named_context_val d e = push_named_context_val (cast_named_decl unsafe_eq d) e
+
+let rel_context e = cast_rel_context (sym unsafe_eq) (rel_context e)
+let named_context e = cast_named_context (sym unsafe_eq) (named_context e)
+
+let val_of_named_context e = val_of_named_context (cast_named_context unsafe_eq e)
+let named_context_of_val e = cast_named_context (sym unsafe_eq) (named_context_of_val e)
+
+let lookup_rel i e = cast_rel_decl (sym unsafe_eq) (lookup_rel i e)
+let lookup_named n e = cast_named_decl (sym unsafe_eq) (lookup_named n e)
+let lookup_named_val n e = cast_named_decl (sym unsafe_eq) (lookup_named_val n e)
+
+let fresh_global ?loc ?rigid ?names env sigma reference =
+ let (evd,t) = Evd.fresh_global ?loc ?rigid ?names env sigma reference in
+ evd, of_constr t
+
+let is_global sigma gr c =
+ Globnames.is_global gr (to_constr sigma c)
+
+module Unsafe =
+struct
+let to_sorts = ESorts.unsafe_to_sorts
+let to_instance = EInstance.unsafe_to_instance
+let to_constr = unsafe_to_constr
+let to_rel_decl = unsafe_to_rel_decl
+let to_named_decl = unsafe_to_named_decl
+let eq = unsafe_eq
+end
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
new file mode 100644
index 000000000..4dbf6c18a
--- /dev/null
+++ b/engine/eConstr.mli
@@ -0,0 +1,294 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CSig
+open Names
+open Constr
+open Environ
+
+type t
+(** Type of incomplete terms. Essentially a wrapper around {!Constr.t} ensuring
+ that {!Constr.kind} does not observe defined evars. *)
+
+type types = t
+type constr = t
+type existential = t pexistential
+type fixpoint = (t, t) pfixpoint
+type cofixpoint = (t, t) pcofixpoint
+type unsafe_judgment = (constr, types) Environ.punsafe_judgment
+type unsafe_type_judgment = types Environ.punsafe_type_judgment
+type named_declaration = (constr, types) Context.Named.Declaration.pt
+type rel_declaration = (constr, types) Context.Rel.Declaration.pt
+type named_context = (constr, types) Context.Named.pt
+type rel_context = (constr, types) Context.Rel.pt
+
+(** {5 Universe variables} *)
+
+module ESorts :
+sig
+ type t
+ (** Type of sorts up-to universe unification. Essentially a wrapper around
+ Sorts.t so that normalization is ensured statically. *)
+
+ val make : Sorts.t -> t
+ (** Turn a sort into an up-to sort. *)
+
+ val kind : Evd.evar_map -> t -> Sorts.t
+ (** Returns the view into the current sort. Note that the kind of a variable
+ may change if the unification state of the evar map changes. *)
+
+end
+
+module EInstance :
+sig
+ type t
+ (** Type of universe instances up-to universe unification. Similar to
+ {ESorts.t} for {Univ.Instance.t}. *)
+
+ val make : Univ.Instance.t -> t
+ val kind : Evd.evar_map -> t -> Univ.Instance.t
+ val empty : t
+ val is_empty : t -> bool
+end
+
+(** {5 Destructors} *)
+
+val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term
+(** Same as {!Constr.kind} except that it expands evars and normalizes
+ universes on the fly. *)
+
+val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t, Sorts.t, Univ.Instance.t) Constr.kind_of_term
+
+val to_constr : Evd.evar_map -> t -> Constr.t
+(** Returns the evar-normal form of the argument. See {!Evarutil.nf_evar}. *)
+
+val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type
+
+(** {5 Constructors} *)
+
+val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t
+(** Construct a term from a view. *)
+
+val of_constr : Constr.t -> t
+(** Translate a kernel term into an incomplete term in O(1). *)
+
+(** {5 Insensitive primitives}
+
+ Evar-insensitive versions of the corresponding functions. See the {!Constr}
+ module for more information.
+
+*)
+
+(** {6 Constructors} *)
+
+val mkRel : int -> t
+val mkVar : Id.t -> t
+val mkMeta : metavariable -> t
+val mkEvar : t pexistential -> t
+val mkSort : Sorts.t -> t
+val mkProp : t
+val mkSet : t
+val mkType : Univ.universe -> t
+val mkCast : t * cast_kind * t -> t
+val mkProd : Name.t * t * t -> t
+val mkLambda : Name.t * t * t -> t
+val mkLetIn : Name.t * t * t * t -> t
+val mkApp : t * t array -> t
+val mkConst : constant -> t
+val mkConstU : constant * EInstance.t -> t
+val mkProj : (projection * t) -> t
+val mkInd : inductive -> t
+val mkIndU : inductive * EInstance.t -> t
+val mkConstruct : constructor -> t
+val mkConstructU : constructor * EInstance.t -> t
+val mkConstructUi : (inductive * EInstance.t) * int -> t
+val mkCase : case_info * t * t * t array -> t
+val mkFix : (t, t) pfixpoint -> t
+val mkCoFix : (t, t) pcofixpoint -> t
+val mkArrow : t -> t -> t
+
+val applist : t * t list -> t
+
+val mkProd_or_LetIn : rel_declaration -> t -> t
+val mkLambda_or_LetIn : rel_declaration -> t -> t
+val it_mkProd_or_LetIn : t -> rel_context -> t
+val it_mkLambda_or_LetIn : t -> rel_context -> t
+
+val mkNamedLambda : Id.t -> types -> constr -> constr
+val mkNamedLetIn : Id.t -> constr -> types -> constr -> constr
+val mkNamedProd : Id.t -> types -> types -> types
+val mkNamedLambda_or_LetIn : named_declaration -> types -> types
+val mkNamedProd_or_LetIn : named_declaration -> types -> types
+
+(** {6 Simple case analysis} *)
+
+val isRel : Evd.evar_map -> t -> bool
+val isVar : Evd.evar_map -> t -> bool
+val isInd : Evd.evar_map -> t -> bool
+val isEvar : Evd.evar_map -> t -> bool
+val isMeta : Evd.evar_map -> t -> bool
+val isSort : Evd.evar_map -> t -> bool
+val isCast : Evd.evar_map -> t -> bool
+val isApp : Evd.evar_map -> t -> bool
+val isLambda : Evd.evar_map -> t -> bool
+val isLetIn : Evd.evar_map -> t -> bool
+val isProd : Evd.evar_map -> t -> bool
+val isConst : Evd.evar_map -> t -> bool
+val isConstruct : Evd.evar_map -> t -> bool
+val isFix : Evd.evar_map -> t -> bool
+val isCoFix : Evd.evar_map -> t -> bool
+val isCase : Evd.evar_map -> t -> bool
+val isProj : Evd.evar_map -> t -> bool
+val isArity : Evd.evar_map -> t -> bool
+val isVarId : Evd.evar_map -> Id.t -> t -> bool
+val isRelN : Evd.evar_map -> int -> t -> bool
+
+val destRel : Evd.evar_map -> t -> int
+val destMeta : Evd.evar_map -> t -> metavariable
+val destVar : Evd.evar_map -> t -> Id.t
+val destSort : Evd.evar_map -> t -> ESorts.t
+val destCast : Evd.evar_map -> t -> t * cast_kind * t
+val destProd : Evd.evar_map -> t -> Name.t * types * types
+val destLambda : Evd.evar_map -> t -> Name.t * types * t
+val destLetIn : Evd.evar_map -> t -> Name.t * t * types * t
+val destApp : Evd.evar_map -> t -> t * t array
+val destConst : Evd.evar_map -> t -> constant * EInstance.t
+val destEvar : Evd.evar_map -> t -> t pexistential
+val destInd : Evd.evar_map -> t -> inductive * EInstance.t
+val destConstruct : Evd.evar_map -> t -> constructor * EInstance.t
+val destCase : Evd.evar_map -> t -> case_info * t * t * t array
+val destProj : Evd.evar_map -> t -> projection * t
+val destFix : Evd.evar_map -> t -> (t, t) pfixpoint
+val destCoFix : Evd.evar_map -> t -> (t, t) pcofixpoint
+
+val decompose_app : Evd.evar_map -> t -> t * t list
+
+val decompose_lam : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_lam_assum : Evd.evar_map -> t -> rel_context * t
+val decompose_lam_n_assum : Evd.evar_map -> int -> t -> rel_context * t
+val decompose_lam_n_decls : Evd.evar_map -> int -> t -> rel_context * t
+
+val compose_lam : (Name.t * t) list -> t -> t
+val to_lambda : Evd.evar_map -> int -> t -> t
+
+val decompose_prod : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_prod_assum : Evd.evar_map -> t -> rel_context * t
+val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t
+
+val existential_type : Evd.evar_map -> existential -> types
+val whd_evar : Evd.evar_map -> constr -> constr
+
+(** {6 Equality} *)
+
+val eq_constr : Evd.evar_map -> t -> t -> bool
+val eq_constr_nounivs : Evd.evar_map -> t -> t -> bool
+val eq_constr_universes : Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val leq_constr_universes : Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val eq_constr_universes_proj : Environ.env -> Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val compare_constr : Evd.evar_map -> (t -> t -> bool) -> t -> t -> bool
+
+(** {6 Iterators} *)
+
+val map : Evd.evar_map -> (t -> t) -> t -> t
+val map_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> t) -> 'a -> t -> t
+val iter : Evd.evar_map -> (t -> unit) -> t -> unit
+val iter_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> unit) -> 'a -> t -> unit
+val iter_with_full_binders : Evd.evar_map -> (rel_declaration -> 'a -> 'a) -> ('a -> t -> unit) -> 'a -> t -> unit
+val fold : Evd.evar_map -> ('a -> t -> 'a) -> 'a -> t -> 'a
+
+(** {6 Substitutions} *)
+
+module Vars :
+sig
+
+(** See vars.mli for the documentation of the functions below *)
+
+type substl = t list
+
+val lift : int -> t -> t
+val liftn : int -> int -> t -> t
+val substnl : substl -> int -> t -> t
+val substl : substl -> t -> t
+val subst1 : t -> t -> t
+
+val substnl_decl : substl -> int -> rel_declaration -> rel_declaration
+val substl_decl : substl -> rel_declaration -> rel_declaration
+val subst1_decl : t -> rel_declaration -> rel_declaration
+
+val replace_vars : (Id.t * t) list -> t -> t
+val substn_vars : int -> Id.t list -> t -> t
+val subst_vars : Id.t list -> t -> t
+val subst_var : Id.t -> t -> t
+
+val noccurn : Evd.evar_map -> int -> t -> bool
+val noccur_between : Evd.evar_map -> int -> int -> t -> bool
+
+val closedn : Evd.evar_map -> int -> t -> bool
+val closed0 : Evd.evar_map -> t -> bool
+
+val subst_univs_level_constr : Univ.universe_level_subst -> t -> t
+val subst_of_rel_context_instance : rel_context -> t list -> t list
+
+
+end
+
+(** {5 Environment handling} *)
+
+val push_rel : rel_declaration -> env -> env
+val push_rel_context : rel_context -> env -> env
+
+val push_named : named_declaration -> env -> env
+val push_named_context : named_context -> env -> env
+val push_named_context_val : named_declaration -> named_context_val -> named_context_val
+
+val rel_context : env -> rel_context
+val named_context : env -> named_context
+
+val val_of_named_context : named_context -> named_context_val
+val named_context_of_val : named_context_val -> named_context
+
+val lookup_rel : int -> env -> rel_declaration
+val lookup_named : variable -> env -> named_declaration
+val lookup_named_val : variable -> named_context_val -> named_declaration
+
+(* XXX Missing Sigma proxy *)
+val fresh_global :
+ ?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:Univ.Instance.t -> Environ.env ->
+ Evd.evar_map -> Globnames.global_reference -> Evd.evar_map * t
+
+val is_global : Evd.evar_map -> Globnames.global_reference -> t -> bool
+
+(** {5 Extra} *)
+
+val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, types) Context.Named.Declaration.pt
+val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, types) Context.Rel.Declaration.pt
+
+val to_rel_decl : Evd.evar_map -> (t, types) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+
+(** {5 Unsafe operations} *)
+
+module Unsafe :
+sig
+ val to_constr : t -> Constr.t
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_rel_decl : (t, types) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_named_decl : (t, types) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_sorts : ESorts.t -> Sorts.t
+ (** Physical identity. Does not care for normalization. *)
+
+ val to_instance : EInstance.t -> Univ.Instance.t
+ (** Physical identity. Does not care for normalization. *)
+
+ val eq : (t, Constr.t) eq
+ (** Use for transparent cast between types. *)
+end
diff --git a/engine/engine.mllib b/engine/engine.mllib
index 9ce5af819..afc02d7f6 100644
--- a/engine/engine.mllib
+++ b/engine/engine.mllib
@@ -1,9 +1,10 @@
Logic_monad
-Termops
-Namegen
+Universes
UState
Evd
-Sigma
+EConstr
+Namegen
+Termops
Proofview_monad
Evarutil
Proofview
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 3d08585a0..2afc12cd3 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -10,17 +10,14 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Termops
open Namegen
open Pre_env
open Environ
open Evd
-open Sigma.Notations
-let safe_evar_info sigma evk =
- try Some (Evd.find sigma evk)
- with Not_found -> None
+module RelDecl = Context.Rel.Declaration
+module NamedDecl = Context.Named.Declaration
let safe_evar_value sigma ev =
try Some (Evd.existential_value sigma ev)
@@ -44,10 +41,11 @@ let evd_comb2 f evdref x y =
z
let e_new_global evdref x =
- evd_comb1 (Evd.fresh_global (Global.env())) evdref x
+ EConstr.of_constr (evd_comb1 (Evd.fresh_global (Global.env())) evdref x)
let new_global evd x =
- Sigma.fresh_global (Global.env()) evd x
+ let (evd, c) = Evd.fresh_global (Global.env()) evd x in
+ (evd, EConstr.of_constr c)
(****************************************************)
(* Expanding/testing/exposing existential variables *)
@@ -65,34 +63,15 @@ let rec flush_and_check_evars sigma c =
| Some c -> flush_and_check_evars sigma c)
| _ -> map_constr (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 flush_and_check_evars sigma c =
+ flush_and_check_evars sigma (EConstr.Unsafe.to_constr c)
-let rec whd_evar sigma c =
- match kind_of_term c with
- | Evar (evk, args) ->
- begin match safe_evar_info sigma evk with
- | Some ({ evar_body = Evar_defined c } as info) ->
- let args = Array.map (fun c -> whd_evar sigma c) args in
- let c = instantiate_evar_array info c args in
- whd_evar sigma c
- | _ -> c
- end
- | 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)
+(** Term exploration up to instantiation. *)
+let kind_of_term_upto = EConstr.kind_upto
+
+let nf_evar0 sigma t = EConstr.to_constr sigma (EConstr.of_constr t)
+let whd_evar = EConstr.whd_evar
+let nf_evar sigma c = EConstr.of_constr (EConstr.to_constr sigma c)
let j_nf_evar sigma j =
{ uj_val = nf_evar sigma j.uj_val;
@@ -117,23 +96,23 @@ let e_nf_evars_and_universes evdref =
let nf_evar_map_universes evm =
let evm = Evd.nf_constraints evm in
let subst = Evd.universe_subst evm in
- if Univ.LMap.is_empty subst then evm, nf_evar evm
+ if Univ.LMap.is_empty subst then evm, nf_evar0 evm
else
let f = nf_evars_universes evm in
Evd.raw_map (fun _ -> map_evar_info f) evm, f
let nf_named_context_evar sigma ctx =
- Context.Named.map (nf_evar sigma) ctx
+ Context.Named.map (nf_evar0 sigma) ctx
let nf_rel_context_evar sigma ctx =
Context.Rel.map (nf_evar sigma) ctx
let nf_env_evar sigma env =
let nc' = nf_named_context_evar sigma (Environ.named_context env) in
- let rel' = nf_rel_context_evar sigma (Environ.rel_context env) in
- push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
+ let rel' = nf_rel_context_evar sigma (EConstr.rel_context env) in
+ EConstr.push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
-let nf_evar_info evc info = map_evar_info (nf_evar evc) info
+let nf_evar_info evc info = map_evar_info (nf_evar0 evc) info
let nf_evar_map evm =
Evd.raw_map (fun _ evi -> nf_evar_info evm evi) evm
@@ -145,21 +124,11 @@ let nf_evar_map_undefined evm =
(* Auxiliary functions for the conversion algorithms modulo evars
*)
-(* A probably faster though more approximative variant of
- [has_undefined (nf_evar c)]: instances are not substituted and
- maybe an evar occurs in an instance and it would disappear by
- instantiation *)
-
let has_undefined_evars evd t =
let rec has_ev t =
- match kind_of_term t with
- | Evar (ev,args) ->
- (match evar_body (Evd.find evd ev) with
- | Evar_defined c ->
- has_ev c; Array.iter has_ev args
- | Evar_empty ->
- raise NotInstantiatedEvar)
- | _ -> iter_constr has_ev t in
+ match EConstr.kind evd t with
+ | Evar _ -> raise NotInstantiatedEvar
+ | _ -> EConstr.iter evd has_ev t in
try let _ = has_ev t in false
with (Not_found | NotInstantiatedEvar) -> true
@@ -167,13 +136,11 @@ let is_ground_term evd t =
not (has_undefined_evars evd t)
let is_ground_env evd env =
- let open Context.Rel.Declaration in
let is_ground_rel_decl = function
- | LocalDef (_,b,_) -> is_ground_term evd b
+ | RelDecl.LocalDef (_,b,_) -> is_ground_term evd (EConstr.of_constr b)
| _ -> true in
- let open Context.Named.Declaration in
let is_ground_named_decl = function
- | LocalDef (_,b,_) -> is_ground_term evd b
+ | NamedDecl.LocalDef (_,b,_) -> is_ground_term evd (EConstr.of_constr b)
| _ -> true in
List.for_all is_ground_rel_decl (rel_context env) &&
List.for_all is_ground_named_decl (named_context env)
@@ -192,7 +159,9 @@ let is_ground_env = memo is_ground_env
exception NoHeadEvar
-let head_evar =
+let head_evar sigma c =
+ (** FIXME: this breaks if using evar-insensitive code *)
+ let c = EConstr.Unsafe.to_constr c in
let rec hrec c = match kind_of_term c with
| Evar (evk,_) -> evk
| Case (_,_,c,_) -> hrec c
@@ -201,33 +170,24 @@ let head_evar =
| Proj (p, c) -> hrec c
| _ -> raise NoHeadEvar
in
- hrec
+ hrec c
(* Expand head evar if any (currently consider only applications but I
guess it should consider Case too) *)
let whd_head_evar_stack sigma c =
- let rec whrec (c, l as s) =
- match kind_of_term c with
- | Evar (evk,args as ev) ->
- let v =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar | Not_found -> None in
- begin match v with
- | None -> s
- | Some c -> whrec (c, l)
- end
+ let rec whrec (c, l) =
+ match EConstr.kind sigma c with
| Cast (c,_,_) -> whrec (c, l)
| App (f,args) -> whrec (f, args :: l)
- | _ -> s
+ | c -> (EConstr.of_kind c, l)
in
whrec (c, [])
let whd_head_evar sigma c =
+ let open EConstr in
let (f, args) = whd_head_evar_stack sigma c in
- (** optim: don't reallocate if empty/singleton *)
match args with
- | [] -> f
| [arg] -> mkApp (f, arg)
| _ -> mkApp (f, Array.concat args)
@@ -242,7 +202,7 @@ let new_meta =
let meta_ctr = Summary.ref 0 ~name:meta_counter_summary_name in
fun () -> incr meta_ctr; !meta_ctr
-let mk_new_meta () = mkMeta(new_meta())
+let mk_new_meta () = EConstr.mkMeta(new_meta())
(* The list of non-instantiated existential declarations (order is important) *)
@@ -255,12 +215,11 @@ let non_instantiated sigma =
(************************)
let make_pure_subst evi args =
- let open Context.Named.Declaration in
snd (List.fold_right
(fun decl (args,l) ->
match args with
- | a::rest -> (rest, (get_id decl, a)::l)
- | _ -> anomaly (Pp.str "Instance does not match its signature"))
+ | a::rest -> (rest, (NamedDecl.get_id decl, a)::l)
+ | _ -> anomaly (Pp.str "Instance does not match its signature."))
(evar_filtered_context evi) (Array.rev_to_list args,[]))
(*------------------------------------*
@@ -297,17 +256,20 @@ let make_pure_subst evi args =
*)
let csubst_subst (k, s) c =
+ (** Safe because this is a substitution *)
+ let c = EConstr.Unsafe.to_constr c in
let rec subst n c = match Constr.kind c with
| Rel m ->
if m <= n then c
- else if m - n <= k then Int.Map.find (k - m + n) s
+ else if m - n <= k then EConstr.Unsafe.to_constr (Int.Map.find (k - m + n) s)
else mkRel (m - k)
| _ -> Constr.map_with_binders succ subst n c
in
- if k = 0 then c else subst 0 c
+ let c = if k = 0 then c else subst 0 c in
+ EConstr.of_constr c
let subst2 subst vsubst c =
- csubst_subst subst (replace_vars vsubst c)
+ csubst_subst subst (EConstr.Vars.replace_vars vsubst c)
let next_ident_away id avoid =
let avoid id = Id.Set.mem id avoid in
@@ -318,34 +280,36 @@ let next_name_away na avoid =
let id = match na with Name id -> id | Anonymous -> default_non_dependent_ident in
next_ident_away_from id avoid
-type csubst = int * Constr.t Int.Map.t
+type csubst = int * EConstr.t Int.Map.t
let empty_csubst = (0, Int.Map.empty)
type ext_named_context =
- csubst * (Id.t * Constr.constr) list *
- Id.Set.t * Context.Named.t
+ csubst * (Id.t * EConstr.constr) list *
+ Id.Set.t * EConstr.named_context
let push_var id (n, s) =
- let s = Int.Map.add n (mkVar id) s in
+ let s = Int.Map.add n (EConstr.mkVar id) s in
(succ n, s)
-let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
- let open Context.Named.Declaration in
+let push_rel_decl_to_named_context sigma decl (subst, vsubst, avoid, nc) =
+ let open EConstr in
+ let open Vars in
+ let map_decl f d =
+ NamedDecl.map_constr f d
+ in
let replace_var_named_declaration id0 id decl =
- let id' = get_id decl in
+ let id' = NamedDecl.get_id decl in
let id' = if Id.equal id0 id' then id else id' in
let vsubst = [id0 , mkVar id] in
- decl |> set_id id' |> map_constr (replace_vars vsubst)
+ decl |> NamedDecl.set_id id' |> map_decl (replace_vars vsubst)
in
let extract_if_neq id = function
| Anonymous -> None
| Name id' when id_ord id id' = 0 -> None
| Name id' -> Some id'
in
- let open Context.Rel.Declaration in
- let (na, c, t) = to_tuple decl in
- let open Context.Named.Declaration in
+ let na = RelDecl.get_name decl in
let id =
(* ppedrot: we want to infer nicer names for the refine tactic, but
keeping at the same time backward compatibility in other code
@@ -356,7 +320,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
else
(** id_of_name_using_hdchar only depends on the rel context which is empty
here *)
- next_ident_away (id_of_name_using_hdchar empty_env t na) avoid
+ next_ident_away (id_of_name_using_hdchar empty_env sigma (RelDecl.get_type decl) na) avoid
in
match extract_if_neq id na with
| Some id0 when not (is_section_variable id0) ->
@@ -366,10 +330,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
context. Unless [id] is a section variable. *)
let subst = (fst subst, Int.Map.map (replace_vars [id0,mkVar id]) (snd subst)) in
let vsubst = (id0,mkVar id)::vsubst in
- let d = match c with
- | None -> LocalAssum (id0, subst2 subst vsubst t)
- | Some c -> LocalDef (id0, subst2 subst vsubst c, subst2 subst vsubst t)
- in
+ let d = decl |> NamedDecl.of_rel_decl (fun _ -> id0) |> map_decl (subst2 subst vsubst) in
let nc = List.map (replace_var_named_declaration id0 id) nc in
(push_var id0 subst, vsubst, Id.Set.add id avoid, d :: nc)
| _ ->
@@ -377,15 +338,13 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
incorrect. We revert to a less robust behaviour where
the new binder has name [id]. Which amounts to the same
behaviour than when [id=id0]. *)
- let d = match c with
- | None -> LocalAssum (id, subst2 subst vsubst t)
- | Some c -> LocalDef (id, subst2 subst vsubst c, subst2 subst vsubst t)
- in
+ let d = decl |> NamedDecl.of_rel_decl (fun _ -> id) |> map_decl (subst2 subst vsubst) in
(push_var id subst, vsubst, Id.Set.add id avoid, d :: nc)
-let push_rel_context_to_named_context env typ =
+let push_rel_context_to_named_context env sigma typ =
(* compute the instances relative to the named context and rel_context *)
let open Context.Named.Declaration in
+ let open EConstr in
let ids = List.map get_id (named_context env) in
let inst_vars = List.map mkVar ids in
if List.is_empty (Environ.rel_context env) then
@@ -397,7 +356,7 @@ let push_rel_context_to_named_context env typ =
(* with vars of the rel context *)
(* We do keep the instances corresponding to local definition (see above) *)
let (subst, vsubst, _, env) =
- Context.Rel.fold_outside push_rel_decl_to_named_context
+ Context.Rel.fold_outside (fun d acc -> push_rel_decl_to_named_context sigma d acc)
(rel_context env) ~init:(empty_csubst, [], avoid, named_context env) in
(val_of_named_context env, subst2 subst vsubst typ, inst_rels@inst_vars, subst, vsubst)
@@ -405,23 +364,31 @@ let push_rel_context_to_named_context env typ =
* Entry points to define new evars *
*------------------------------------*)
-let default_source = (Loc.ghost,Evar_kinds.InternalHole)
+let default_source = Loc.tag @@ Evar_kinds.InternalHole
-let restrict_evar evd evk filter candidates =
- let evd = Sigma.to_evar_map evd in
- let evd, evk' = Evd.restrict evk filter ?candidates evd in
- Sigma.Unsafe.of_pair (evk', Evd.declare_future_goal evk' evd)
+let restrict_evar evd evk filter ?src candidates =
+ let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in
+ let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in
+ Evd.declare_future_goal evk' evd, evk'
let new_pure_evar_full evd evi =
- let evd = Sigma.to_evar_map evd in
let (evd, evk) = Evd.new_evar evd evi in
let evd = Evd.declare_future_goal evk evd in
- Sigma.Unsafe.of_pair (evk, evd)
+ (evd, evk)
let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?candidates ?(store = Store.empty) ?naming ?(principal=false) typ =
- let evd = Sigma.to_evar_map evd in
+ let typ = EConstr.Unsafe.to_constr typ in
+ let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in
let default_naming = Misctypes.IntroAnonymous in
let naming = Option.default default_naming naming in
+ let name = match naming with
+ | Misctypes.IntroAnonymous -> None
+ | Misctypes.IntroIdentifier id -> Some id
+ | Misctypes.IntroFresh id ->
+ let has_name id = try let _ = Evd.evar_key id evd in true with Not_found -> false in
+ let id = Namegen.next_ident_away_from id has_name in
+ Some id
+ in
let evi = {
evar_hyps = sign;
evar_concl = typ;
@@ -431,70 +398,67 @@ let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?ca
evar_candidates = candidates;
evar_extra = store; }
in
- let (evd, newevk) = Evd.new_evar evd ~naming evi in
+ let (evd, newevk) = Evd.new_evar evd ?name evi in
let evd =
if principal then Evd.declare_principal_goal newevk evd
else Evd.declare_future_goal newevk evd
in
- Sigma.Unsafe.of_pair (newevk, evd)
+ (evd, newevk)
let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?principal instance =
+ let open EConstr in
assert (not !Flags.debug ||
List.distinct (ids_of_named_context (named_context_of_val sign)));
- let Sigma (newevk, evd, p) = new_pure_evar sign evd ?src ?filter ?candidates ?store ?naming ?principal typ in
- Sigma (mkEvar (newevk,Array.of_list instance), evd, p)
+ let (evd, newevk) = new_pure_evar sign evd ?src ?filter ?candidates ?store ?naming ?principal typ in
+ evd, mkEvar (newevk,Array.of_list instance)
(* [new_evar] declares a new existential in an env env with type typ *)
(* Converting the env into the sign of the evar to define *)
let new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ =
- let sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env typ in
- let candidates = Option.map (List.map (subst2 subst vsubst)) candidates in
+ let sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env evd typ in
+ let map c = subst2 subst vsubst c in
+ let candidates = Option.map (fun l -> List.map map l) candidates in
let instance =
match filter with
| None -> instance
| Some filter -> Filter.filter_list filter instance in
new_evar_instance sign evd typ' ?src ?filter ?candidates ?store ?naming ?principal instance
-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 new_type_evar env evd ?src ?filter ?naming ?principal rigid =
- let Sigma (s, evd', p) = Sigma.new_sort_variable rigid evd in
- let Sigma (e, evd', q) = new_evar env evd' ?src ?filter ?naming ?principal (mkSort s) in
- Sigma ((e, s), evd', p +> q)
+ let (evd', s) = new_sort_variable rigid evd in
+ let (evd', e) = new_evar env evd' ?src ?filter ?naming ?principal (EConstr.mkSort s) in
+ evd', (e, s)
let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid =
- let sigma = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma (c, sigma, _) = new_type_evar env sigma ?src ?filter ?naming ?principal rigid in
- let sigma = Sigma.to_evar_map sigma in
- evdref := sigma;
+ let (evd, c) = new_type_evar env !evdref ?src ?filter ?naming ?principal rigid in
+ evdref := evd;
c
let new_Type ?(rigid=Evd.univ_flexible) env evd =
- let Sigma (s, sigma, p) = Sigma.new_sort_variable rigid evd in
- Sigma (mkSort s, sigma, p)
+ let open EConstr in
+ let (evd, s) = new_sort_variable rigid evd in
+ (evd, mkSort s)
let e_new_Type ?(rigid=Evd.univ_flexible) env evdref =
let evd', s = new_sort_variable rigid !evdref in
- evdref := evd'; mkSort s
+ evdref := evd'; EConstr.mkSort s
(* The same using side-effect *)
let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?naming ?principal ty =
- let (evd',ev) = new_evar_unsafe env !evdref ~src:src ?filter ?candidates ?store ?naming ?principal ty in
+ let (evd',ev) = new_evar env !evdref ~src:src ?filter ?candidates ?store ?naming ?principal ty in
evdref := evd';
ev
(* This assumes an evar with identity instance and generalizes it over only
- the De Bruijn part of the context *)
+ the de Bruijn part of the context *)
let generalize_evar_over_rels sigma (ev,args) =
+ let open EConstr in
let evi = Evd.find sigma ev in
let sign = named_context_of_val evi.evar_hyps in
List.fold_left2
(fun (c,inst as x) a d ->
- if isRel a then (mkNamedProd_or_LetIn d c,a::inst) else x)
- (evi.evar_concl,[]) (Array.to_list args) sign
+ if isRel sigma a then (mkNamedProd_or_LetIn d c,a::inst) else x)
+ (EConstr.of_constr evi.evar_concl,[]) (Array.to_list args) sign
(************************************)
(* Removing a dependency in an evar *)
@@ -532,7 +496,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
| Evar (evk,l as ev) ->
if Evd.is_defined !evdref evk then
(* If evk is already defined we replace it by its definition *)
- let nc = whd_evar !evdref c in
+ let nc = Evd.existential_value !evdref ev in
(check_and_clear_in_constr env evdref err ids global nc)
else
(* We check for dependencies to elements of ids in the
@@ -542,6 +506,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
removed *)
let evi = Evd.find_undefined !evdref evk in
let ctxt = Evd.evar_filtered_context evi in
+ let ctxt = List.map (fun d -> map_named_decl EConstr.of_constr d) ctxt in
let (rids,filter) =
List.fold_right2
(fun h a (ri,filter) ->
@@ -549,19 +514,18 @@ let rec check_and_clear_in_constr env evdref err ids global c =
(* Check if some id to clear occurs in the instance
a of rid in ev and remember the dependency *)
let check id = if Id.Set.mem id ids then raise (Depends id) in
- let () = Id.Set.iter check (collect_vars a) in
+ let () = Id.Set.iter check (collect_vars !evdref (EConstr.of_constr a)) in
(* Check if some rid to clear in the context of ev
has dependencies in another hyp of the context of ev
and transitively remember the dependency *)
let check id _ =
- if occur_var_in_decl (Global.env ()) id h
+ if occur_var_in_decl (Global.env ()) !evdref id h
then raise (Depends id)
in
let () = Id.Map.iter check ri in
(* No dependency at all, we can keep this ev's context hyp *)
(ri, true::filter)
- with Depends id -> let open Context.Named.Declaration in
- (Id.Map.add (get_id h) id ri, false::filter))
+ with Depends id -> (Id.Map.add (NamedDecl.get_id h) id ri, false::filter))
ctxt (Array.to_list l) (Id.Map.empty,[]) in
(* Check if some rid to clear in the context of ev has dependencies
in the type of ev and adjust the source of the dependency *)
@@ -577,9 +541,8 @@ let rec check_and_clear_in_constr env evdref err ids global c =
else
let origfilter = Evd.evar_filter evi in
let filter = Evd.Filter.apply_subfilter origfilter filter in
- let evd = Sigma.Unsafe.of_evar_map !evdref in
- let Sigma (_, evd, _) = restrict_evar evd evk filter None in
- let evd = Sigma.to_evar_map evd in
+ let evd = !evdref in
+ let (evd,_) = restrict_evar evd evk filter None in
evdref := evd;
(* spiwack: hacking session to mark the old [evk] as having been "cleared" *)
let evi = Evd.find !evdref evk in
@@ -588,7 +551,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
let evi' = { evi with evar_extra = extra' } in
evdref := Evd.add !evdref evk evi' ;
(* spiwack: /hacking session *)
- whd_evar !evdref c
+ Evd.existential_value !evdref ev
| _ -> map_constr (check_and_clear_in_constr env evdref err ids global) c
@@ -596,14 +559,14 @@ let clear_hyps_in_evi_main env evdref hyps terms ids =
(* clear_hyps_in_evi erases hypotheses ids in hyps, checking if some
hypothesis does not depend on a element of ids, and erases ids in
the contexts of the evars occurring in evi *)
+ let terms = List.map EConstr.Unsafe.to_constr terms in
let global = Id.Set.exists is_section_variable ids in
let terms =
List.map (check_and_clear_in_constr env evdref (OccurHypInSimpleClause None) ids global) terms in
let nhyps =
- let open Context.Named.Declaration in
let check_context decl =
- let err = OccurHypInSimpleClause (Some (get_id decl)) in
- map_constr (check_and_clear_in_constr env evdref err ids global) decl
+ let err = OccurHypInSimpleClause (Some (NamedDecl.get_id decl)) in
+ NamedDecl.map_constr (check_and_clear_in_constr env evdref err ids global) decl
in
let check_value vk = match force_lazy_val vk with
| None -> vk
@@ -618,7 +581,7 @@ let clear_hyps_in_evi_main env evdref hyps terms ids =
in
remove_hyps ids check_context check_value hyps
in
- (nhyps,terms)
+ (nhyps,List.map EConstr.of_constr terms)
let clear_hyps_in_evi env evdref hyps concl ids =
match clear_hyps_in_evi_main env evdref hyps [concl] ids with
@@ -642,8 +605,8 @@ let process_dependent_evar q acc evm is_dependent e =
hypotheses), they are all dependent. *)
queue_term q true evi.evar_concl;
List.iter begin fun decl ->
- let open Context.Named.Declaration in
- queue_term q true (get_type decl);
+ let open NamedDecl in
+ queue_term q true (NamedDecl.get_type decl);
match decl with
| LocalAssum _ -> ()
| LocalDef (_,b,_) -> queue_term q true b
@@ -707,30 +670,26 @@ let rec advance sigma evk =
let undefined_evars_of_term evd t =
let rec evrec acc c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (n, l) ->
- let acc = Array.fold_left evrec acc l in
- (try match (Evd.find evd n).evar_body with
- | Evar_empty -> Evar.Set.add n acc
- | Evar_defined c -> evrec acc c
- with Not_found -> anomaly ~label:"undefined_evars_of_term" (Pp.str "evar not found"))
- | _ -> fold_constr evrec acc c
+ let acc = Evar.Set.add n acc in
+ Array.fold_left evrec acc l
+ | _ -> EConstr.fold evd evrec acc c
in
evrec Evar.Set.empty t
let undefined_evars_of_named_context evd nc =
- let open Context.Named.Declaration in
Context.Named.fold_outside
- (fold (fun c s -> Evar.Set.union s (undefined_evars_of_term evd c)))
+ (NamedDecl.fold_constr (fun c s -> Evar.Set.union s (undefined_evars_of_term evd (EConstr.of_constr c))))
nc
~init:Evar.Set.empty
let undefined_evars_of_evar_info evd evi =
- Evar.Set.union (undefined_evars_of_term evd evi.evar_concl)
+ Evar.Set.union (undefined_evars_of_term evd (EConstr.of_constr evi.evar_concl))
(Evar.Set.union
(match evi.evar_body with
| Evar_empty -> Evar.Set.empty
- | Evar_defined b -> undefined_evars_of_term evd b)
+ | Evar_defined b -> undefined_evars_of_term evd (EConstr.of_constr b))
(undefined_evars_of_named_context evd
(named_context_of_val evi.evar_hyps)))
@@ -739,6 +698,7 @@ let undefined_evars_of_evar_info evd evi =
[evar_map]. If unification only need to check superficially, tactics
do not have this luxury, and need the more complete version. *)
let occur_evar_upto sigma n c =
+ let c = EConstr.Unsafe.to_constr c in
let rec occur_rec c = match kind_of_term c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar e -> Option.iter occur_rec (existential_opt_value sigma e)
@@ -750,8 +710,9 @@ let occur_evar_upto sigma n c =
any type has type Type. May cause some trouble, but not so far... *)
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)
+ let open EConstr in
+ let (evd', s) = new_univ_variable univ_rigid evd in
+ (evd', { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) })
let subterm_source evk (loc,k) =
let evk = match k with
@@ -760,10 +721,6 @@ let subterm_source evk (loc,k) =
(loc,Evar_kinds.SubEvar evk)
-(** Term exploration up to instantiation. *)
-let kind_of_term_upto 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
universes. The term [t] is interpreted in [sigma1] while [u] is
@@ -784,5 +741,5 @@ let eq_constr_univs_test sigma1 sigma2 t u =
in
match ans with None -> false | Some _ -> true
-type type_constraint = types option
-type val_constraint = constr option
+type type_constraint = EConstr.types option
+type val_constraint = EConstr.constr option
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index 585b9cedf..a8b6b5861 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -10,6 +10,7 @@ open Names
open Term
open Evd
open Environ
+open EConstr
(** This module provides useful higher-level functions for evar manipulation. *)
@@ -21,22 +22,22 @@ val mk_new_meta : unit -> constr
(** {6 Creating a fresh evar given their type and context} *)
val new_evar :
- env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ env -> evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?candidates:constr list -> ?store:Store.t ->
?naming:Misctypes.intro_pattern_naming_expr ->
- ?principal:bool -> types -> (constr, 'r) Sigma.sigma
+ ?principal:bool -> types -> evar_map * EConstr.t
val new_pure_evar :
- named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ named_context_val -> evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?candidates:constr list -> ?store:Store.t ->
?naming:Misctypes.intro_pattern_naming_expr ->
- ?principal:bool -> types -> (evar, 'r) Sigma.sigma
+ ?principal:bool -> types -> evar_map * evar
-val new_pure_evar_full : 'r Sigma.t -> evar_info -> (evar, 'r) Sigma.sigma
+val new_pure_evar_full : evar_map -> evar_info -> evar_map * evar
(** the same with side-effects *)
val e_new_evar :
- env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ env -> evar_map ref -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?candidates:constr list -> ?store:Store.t ->
?naming:Misctypes.intro_pattern_naming_expr ->
?principal:bool -> types -> constr
@@ -44,23 +45,23 @@ val e_new_evar :
(** Create a new Type existential variable, as we keep track of
them during type-checking and unification. *)
val new_type_evar :
- env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ env -> evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid ->
- (constr * sorts, 'r) Sigma.sigma
+ evar_map * (constr * sorts)
val e_new_type_evar : env -> evar_map ref ->
- ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t ->
+ ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts
-val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (constr, 'r) Sigma.sigma
+val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr
val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr
-val restrict_evar : 'r Sigma.t -> existential_key -> Filter.t ->
- constr list option -> (existential_key, 'r) Sigma.sigma
+val restrict_evar : evar_map -> existential_key -> Filter.t ->
+ ?src:Evar_kinds.t Loc.located -> constr list option -> evar_map * existential_key
(** Polymorphic constants *)
-val new_global : 'r Sigma.t -> Globnames.global_reference -> (constr, 'r) Sigma.sigma
+val new_global : evar_map -> Globnames.global_reference -> evar_map * constr
val e_new_global : evar_map ref -> Globnames.global_reference -> constr
(** Create a fresh evar in a context different from its definition context:
@@ -70,15 +71,15 @@ val e_new_global : evar_map ref -> Globnames.global_reference -> constr
of [inst] are typed in the occurrence context and their type (seen
as a telescope) is [sign] *)
val new_evar_instance :
- named_context_val -> 'r Sigma.t -> types ->
- ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list ->
+ named_context_val -> evar_map -> types ->
+ ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t -> ?candidates:constr list ->
?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr ->
?principal:bool ->
- constr list -> (constr, 'r) Sigma.sigma
+ constr list -> evar_map * constr
-val make_pure_subst : evar_info -> constr array -> (Id.t * constr) list
+val make_pure_subst : evar_info -> 'a array -> (Id.t * 'a) list
-val safe_evar_value : evar_map -> existential -> constr option
+val safe_evar_value : evar_map -> Constr.existential -> Constr.constr option
(** {6 Evars/Metas switching...} *)
@@ -88,7 +89,7 @@ val non_instantiated : evar_map -> evar_info Evar.Map.t
(** [head_evar c] returns the head evar of [c] if any *)
exception NoHeadEvar
-val head_evar : constr -> existential_key (** may raise NoHeadEvar *)
+val head_evar : evar_map -> constr -> existential_key (** may raise NoHeadEvar *)
(* Expand head evar if any *)
val whd_head_evar : evar_map -> constr -> constr
@@ -128,11 +129,11 @@ val undefined_evars_of_evar_info : evar_map -> evar_info -> Evar.Set.t
(** [occur_evar_upto sigma k c] returns [true] if [k] appears in
[c]. It looks up recursively in [sigma] for the value of existential
variables. *)
-val occur_evar_upto : evar_map -> Evar.t -> Constr.t -> bool
+val occur_evar_upto : evar_map -> Evar.t -> constr -> bool
(** {6 Value/Type constraints} *)
-val judge_of_new_Type : 'r Sigma.t -> (unsafe_judgment, 'r) Sigma.sigma
+val judge_of_new_Type : evar_map -> evar_map * unsafe_judgment
(***********************************************************)
@@ -150,7 +151,7 @@ val tj_nf_evar :
evar_map -> unsafe_type_judgment -> unsafe_type_judgment
val nf_named_context_evar : evar_map -> Context.Named.t -> Context.Named.t
-val nf_rel_context_evar : evar_map -> Context.Rel.t -> Context.Rel.t
+val nf_rel_context_evar : evar_map -> rel_context -> rel_context
val nf_env_evar : evar_map -> env -> env
val nf_evar_info : evar_map -> evar_info -> evar_info
@@ -159,39 +160,40 @@ val nf_evar_map_undefined : evar_map -> evar_map
(** Presenting terms without solved evars *)
-val nf_evars_universes : evar_map -> constr -> constr
+val nf_evars_universes : evar_map -> Constr.constr -> Constr.constr
-val nf_evars_and_universes : evar_map -> evar_map * (constr -> constr)
-val e_nf_evars_and_universes : evar_map ref -> (constr -> constr) * Universes.universe_opt_subst
+val nf_evars_and_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
+val e_nf_evars_and_universes : evar_map ref -> (Constr.constr -> Constr.constr) * Universes.universe_opt_subst
(** Normalize the evar map w.r.t. universes, after simplification of constraints.
Return the substitution function for constrs as well. *)
-val nf_evar_map_universes : evar_map -> evar_map * (constr -> constr)
+val nf_evar_map_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
(** Replacing all evars, possibly raising [Uninstantiated_evar] *)
exception Uninstantiated_evar of existential_key
-val flush_and_check_evars : evar_map -> constr -> constr
+val flush_and_check_evars : evar_map -> constr -> Constr.constr
(** {6 Term manipulation up to instantiation} *)
(** Like {!Constr.kind} except that [kind_of_term sigma t] exposes [t]
as an evar [e] only if [e] is uninstantiated in [sigma]. Otherwise the
value of [e] in [sigma] is (recursively) used. *)
-val kind_of_term_upto : evar_map -> constr -> (constr,types) kind_of_term
+val kind_of_term_upto : evar_map -> Constr.constr ->
+ (Constr.constr, Constr.types, Sorts.t, Univ.Instance.t) kind_of_term
(** [eq_constr_univs_test sigma1 sigma2 t u] tests equality of [t] and
[u] up to existential variable instantiation and equalisable
universes. The term [t] is interpreted in [sigma1] while [u] is
interpreted in [sigma2]. The universe constraints in [sigma2] are
assumed to be an extention of those in [sigma1]. *)
-val eq_constr_univs_test : evar_map -> evar_map -> constr -> constr -> bool
+val eq_constr_univs_test : evar_map -> evar_map -> Constr.constr -> Constr.constr -> bool
(** {6 Removing hyps in evars'context}
raise OccurHypInSimpleClause if the removal breaks dependencies *)
type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
-| EvarTypingBreak of existential
+| EvarTypingBreak of Constr.existential
exception ClearDependencyError of Id.t * clear_dependency_error
@@ -208,16 +210,16 @@ val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> ty
type csubst
val empty_csubst : csubst
-val csubst_subst : csubst -> Constr.t -> Constr.t
+val csubst_subst : csubst -> constr -> constr
type ext_named_context =
- csubst * (Id.t * Constr.constr) list *
- Id.Set.t * Context.Named.t
+ csubst * (Id.t * constr) list *
+ Id.Set.t * named_context
val push_rel_decl_to_named_context :
- Context.Rel.Declaration.t -> ext_named_context -> ext_named_context
+ evar_map -> rel_declaration -> ext_named_context -> ext_named_context
-val push_rel_context_to_named_context : Environ.env -> types ->
+val push_rel_context_to_named_context : Environ.env -> evar_map -> types ->
named_context_val * types * constr list * csubst * (identifier*constr) list
val generalize_evar_over_rels : evar_map -> existential -> types * constr list
diff --git a/engine/evd.ml b/engine/evd.ml
index 6ffa4eeb7..cfc9aa635 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -13,10 +13,10 @@ open Names
open Nameops
open Term
open Vars
-open Termops
open Environ
-open Globnames
-open Context.Named.Declaration
+
+(* module RelDecl = Context.Rel.Declaration *)
+module NamedDecl = Context.Named.Declaration
(** Generic filters *)
module Filter :
@@ -149,13 +149,13 @@ let make_evar hyps ccl = {
evar_hyps = hyps;
evar_body = Evar_empty;
evar_filter = Filter.identity;
- evar_source = (Loc.ghost,Evar_kinds.InternalHole);
+ evar_source = Loc.tag @@ Evar_kinds.InternalHole;
evar_candidates = None;
evar_extra = Store.empty
}
let instance_mismatch () =
- anomaly (Pp.str "Signature and its instance do not match")
+ anomaly (Pp.str "Signature and its instance do not match.")
let evar_concl evi = evi.evar_concl
@@ -226,7 +226,7 @@ let evar_instance_array test_id info args =
if i < len then
let c = Array.unsafe_get args i in
if test_id d c then instrec filter ctxt (succ i)
- else (get_id d, c) :: instrec filter ctxt (succ i)
+ else (NamedDecl.get_id d, c) :: instrec filter ctxt (succ i)
else instance_mismatch ()
| _ -> instance_mismatch ()
in
@@ -235,7 +235,7 @@ let evar_instance_array test_id info args =
let map i d =
if (i < len) then
let c = Array.unsafe_get args i in
- if test_id d c then None else Some (get_id d, c)
+ if test_id d c then None else Some (NamedDecl.get_id d, c)
else instance_mismatch ()
in
List.map_filter_i map (evar_context info)
@@ -243,7 +243,7 @@ let evar_instance_array test_id info args =
instrec filter (evar_context info) 0
let make_evar_instance_array info args =
- evar_instance_array (isVarId % get_id) info args
+ evar_instance_array (NamedDecl.get_id %> isVarId) info args
let instantiate_evar_array info c args =
let inst = make_evar_instance_array info args in
@@ -256,7 +256,6 @@ type evar_universe_context = UState.t
type 'a in_evar_universe_context = 'a * evar_universe_context
let empty_evar_universe_context = UState.empty
-let is_empty_evar_universe_context = UState.is_empty
let union_evar_universe_context = UState.union
let evar_universe_context_set = UState.context_set
let evar_universe_context_constraints = UState.constraints
@@ -284,7 +283,7 @@ let metavars_of c =
let rec collrec acc c =
match kind_of_term c with
| Meta mv -> Int.Set.add mv acc
- | _ -> fold_constr collrec acc c
+ | _ -> Term.fold_constr collrec acc c
in
collrec Int.Set.empty c
@@ -359,12 +358,10 @@ module EvMap = Evar.Map
module EvNames :
sig
-open Misctypes
-
type t
val empty : t
-val add_name_undefined : intro_pattern_naming_expr -> Evar.t -> evar_info -> t -> t
+val add_name_undefined : Id.t option -> Evar.t -> evar_info -> t -> t
val remove_name_defined : Evar.t -> t -> t
val rename : Evar.t -> Id.t -> t -> t
val reassign_name_defined : Evar.t -> Evar.t -> t -> t
@@ -378,21 +375,13 @@ type t = Id.t EvMap.t * existential_key Idmap.t
let empty = (EvMap.empty, Idmap.empty)
-let add_name_newly_undefined naming evk evi (evtoid, idtoev as names) =
- let id = match naming with
- | Misctypes.IntroAnonymous -> None
- | Misctypes.IntroIdentifier id ->
- if Idmap.mem id idtoev then
- user_err_loc
- (Loc.ghost,"",str "Already an existential evar of name " ++ pr_id id);
- Some id
- | Misctypes.IntroFresh id ->
- let id = Namegen.next_ident_away_from id (fun id -> Idmap.mem id idtoev) in
- Some id
- in
+let add_name_newly_undefined id evk evi (evtoid, idtoev as names) =
match id with
| None -> names
- | Some id -> (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
+ | Some id ->
+ if Idmap.mem id idtoev then
+ user_err (str "Already an existential evar of name " ++ pr_id id);
+ (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
let add_name_undefined naming evk evi (evtoid,idtoev as evar_names) =
if EvMap.mem evk evtoid then
@@ -411,7 +400,7 @@ let rename evk id (evtoid, idtoev) =
match id' with
| None -> (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
| Some id' ->
- if Idmap.mem id idtoev then anomaly (str "Evar name already in use");
+ if Idmap.mem id idtoev then anomaly (str "Evar name already in use.");
(EvMap.update evk id evtoid (* overwrite old name *), Idmap.add id evk (Idmap.remove id' idtoev))
let reassign_name_defined evk evk' (evtoid, idtoev as names) =
@@ -462,9 +451,9 @@ type evar_map = {
let rename evk id evd =
{ evd with evar_names = EvNames.rename evk id evd.evar_names }
-let add_with_name ?(naming = Misctypes.IntroAnonymous) d e i = match i.evar_body with
+let add_with_name ?name d e i = match i.evar_body with
| Evar_empty ->
- let evar_names = EvNames.add_name_undefined naming e i d.evar_names in
+ let evar_names = EvNames.add_name_undefined name e i d.evar_names in
{ d with undf_evars = EvMap.add e i d.undf_evars; evar_names }
| Evar_defined _ ->
let evar_names = EvNames.remove_name_defined e d.evar_names in
@@ -481,9 +470,9 @@ let new_untyped_evar =
let evar_ctr = Summary.ref 0 ~name:evar_counter_summary_name in
fun () -> incr evar_ctr; Evar.unsafe_of_int !evar_ctr
-let new_evar evd ?naming evi =
+let new_evar evd ?name evi =
let evk = new_untyped_evar () in
- let evd = add_with_name evd ?naming evk evi in
+ let evd = add_with_name evd ?name evk evi in
(evd, evk)
let remove d e =
@@ -504,15 +493,6 @@ let find_undefined d e = EvMap.find e d.undf_evars
let mem d e = EvMap.mem e d.undf_evars || EvMap.mem e d.defn_evars
-(* spiwack: this function loses information from the original evar_map
- it might be an idea not to export it. *)
-let to_list d =
- (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
- let l = ref [] in
- EvMap.iter (fun evk x -> l := (evk,x)::!l) d.defn_evars;
- EvMap.iter (fun evk x -> l := (evk,x)::!l) d.undf_evars;
- !l
-
let undefined_map d = d.undf_evars
let drop_all_defined d = { d with defn_evars = EvMap.empty }
@@ -573,7 +553,7 @@ let existential_type d (n, args) =
let info =
try find d n
with Not_found ->
- anomaly (str "Evar " ++ str (string_of_existential n) ++ str " was not declared") in
+ anomaly (str "Evar " ++ str (string_of_existential n) ++ str " was not declared.") in
instantiate_evar_array info info.evar_concl args
let add_constraints d c =
@@ -655,9 +635,9 @@ let define_aux def undef evk body =
try EvMap.find evk undef
with Not_found ->
if EvMap.mem evk def then
- anomaly ~label:"Evd.define" (Pp.str "cannot define an evar twice")
+ anomaly ~label:"Evd.define" (Pp.str "cannot define an evar twice.")
else
- anomaly ~label:"Evd.define" (Pp.str "cannot define undeclared evar")
+ anomaly ~label:"Evd.define" (Pp.str "cannot define undeclared evar.")
in
let () = assert (oldinfo.evar_body == Evar_empty) in
let newinfo = { oldinfo with evar_body = Evar_defined body } in
@@ -673,19 +653,19 @@ let define evk body evd =
let evar_names = EvNames.remove_name_defined evk evd.evar_names in
{ evd with defn_evars; undf_evars; last_mods; evar_names }
-let restrict evk filter ?candidates evd =
+let restrict evk filter ?candidates ?src evd =
let evk' = new_untyped_evar () in
let evar_info = EvMap.find evk evd.undf_evars in
let evar_info' =
{ evar_info with evar_filter = filter;
evar_candidates = candidates;
- evar_extra = Store.empty } in
+ evar_source = (match src with None -> evar_info.evar_source | Some src -> src) } in
let last_mods = match evd.conv_pbs with
| [] -> evd.last_mods
| _ -> Evar.Set.add evk evd.last_mods in
let evar_names = EvNames.reassign_name_defined evk evk' evd.evar_names in
let ctxt = Filter.filter_list filter (evar_context evar_info) in
- let id_inst = Array.map_of_list (mkVar % get_id) ctxt in
+ let id_inst = Array.map_of_list (NamedDecl.get_id %> mkVar) ctxt in
let body = mkEvar(evk',id_inst) in
let (defn_evars, undf_evars) = define_aux evd.defn_evars evd.undf_evars evk body in
{ evd with undf_evars = EvMap.add evk' evar_info' undf_evars;
@@ -724,34 +704,26 @@ let loc_of_conv_pb evd (pbty,env,t1,t2) =
| _ ->
match kind_of_term (fst (decompose_app t2)) with
| Evar (evk2,_) -> fst (evar_source evk2 evd)
- | _ -> Loc.ghost
+ | _ -> None
(** The following functions return the set of evars immediately
contained in the object *)
(* excluding defined evars *)
-let evar_list c =
- let rec evrec acc c =
- match kind_of_term c with
- | Evar (evk, _ as ev) -> ev :: acc
- | _ -> fold_constr evrec acc c in
- evrec [] c
-
let evars_of_term c =
let rec evrec acc c =
match kind_of_term c with
| Evar (n, l) -> Evar.Set.add n (Array.fold_left evrec acc l)
- | _ -> fold_constr evrec acc c
+ | _ -> Term.fold_constr evrec acc c
in
evrec Evar.Set.empty c
let evars_of_named_context nc =
- List.fold_right (fun decl s ->
- Option.fold_left (fun s t ->
- Evar.Set.union s (evars_of_term t))
- (Evar.Set.union s (evars_of_term (get_type decl))) (get_value decl))
- nc Evar.Set.empty
+ Context.Named.fold_outside
+ (NamedDecl.fold_constr (fun constr s -> Evar.Set.union s (evars_of_term constr)))
+ nc
+ ~init:Evar.Set.empty
let evars_of_filtered_evar_info evi =
Evar.Set.union (evars_of_term evi.evar_concl)
@@ -776,7 +748,6 @@ let evar_universe_context d = d.universes
let universe_context_set d = UState.context_set d.universes
-let pr_uctx_level = UState.pr_uctx_level
let universe_context ?names evd = UState.universe_context ?names evd.universes
let restrict_universe_context evd vars =
@@ -820,7 +791,7 @@ let make_evar_universe_context e l =
| Some us ->
List.fold_left
(fun uctx (loc,id) ->
- fst (UState.new_univ_variable ~loc univ_rigid (Some (Id.to_string id)) uctx))
+ fst (UState.new_univ_variable ?loc univ_rigid (Some (Id.to_string id)) uctx))
uctx us
(****************************************)
@@ -959,37 +930,8 @@ let universes evd = UState.ugraph evd.universes
let update_sigma_env evd env =
{ evd with universes = UState.update_sigma_env evd.universes env }
-(* Conversion w.r.t. an evar map and its local universes. *)
-
-let test_conversion_gen env evd pb t u =
- match pb with
- | Reduction.CONV ->
- Reduction.conv env
- ~evars:((existential_opt_value evd), (UState.ugraph evd.universes))
- t u
- | Reduction.CUMUL -> Reduction.conv_leq env
- ~evars:((existential_opt_value evd), (UState.ugraph evd.universes))
- t u
-
-let test_conversion env d pb t u =
- try test_conversion_gen env d pb t u; true
- with _ -> false
-
exception UniversesDiffer = UState.UniversesDiffer
-let eq_constr_univs evd t u =
- let fold cstr sigma =
- try Some (add_universe_constraints sigma cstr)
- with Univ.UniverseInconsistency _ | UniversesDiffer -> None
- in
- match Universes.eq_constr_univs_infer (UState.ugraph evd.universes) fold t u evd with
- | None -> evd, false
- | Some evd -> evd, true
-
-let e_eq_constr_univs evdref t u =
- let evd, b = eq_constr_univs !evdref t u in
- evdref := evd; b
-
(**********************************************************)
(* Side effects *)
@@ -1011,7 +953,7 @@ let declare_principal_goal evk evd =
| None -> { evd with
future_goals = evk::evd.future_goals;
principal_future_goal=Some evk; }
- | Some _ -> CErrors.error "Only one main subgoal per instantiation."
+ | Some _ -> CErrors.user_err Pp.(str "Only one main subgoal per instantiation.")
let future_goals evd = evd.future_goals
@@ -1080,7 +1022,7 @@ let try_meta_fvalue evd mv =
let meta_fvalue evd mv =
try try_meta_fvalue evd mv
- with Not_found -> anomaly ~label:"meta_fvalue" (Pp.str "meta has no value")
+ with Not_found -> anomaly ~label:"meta_fvalue" (Pp.str "meta has no value.")
let meta_value evd mv =
(fst (try_meta_fvalue evd mv)).rebus
@@ -1099,7 +1041,7 @@ let meta_declare mv v ?(name=Anonymous) evd =
let meta_assign mv (v, pb) evd =
let modify _ = function
| Cltyp (na, ty) -> Clval (na, (mk_freelisted v, pb), ty)
- | _ -> anomaly ~label:"meta_assign" (Pp.str "already defined")
+ | _ -> anomaly ~label:"meta_assign" (Pp.str "already defined.")
in
let metas = Metamap.modify mv modify evd.metas in
set_metas evd metas
@@ -1107,7 +1049,7 @@ let meta_assign mv (v, pb) evd =
let meta_reassign mv (v, pb) evd =
let modify _ = function
| Clval(na, _, ty) -> Clval (na, (mk_freelisted v, pb), ty)
- | _ -> anomaly ~label:"meta_reassign" (Pp.str "not yet defined")
+ | _ -> anomaly ~label:"meta_reassign" (Pp.str "not yet defined.")
in
let metas = Metamap.modify mv modify evd.metas in
set_metas evd metas
@@ -1141,14 +1083,14 @@ let retract_coercible_metas evd =
let evar_source_of_meta mv evd =
match meta_name evd mv with
- | Anonymous -> (Loc.ghost,Evar_kinds.GoalEvar)
- | Name id -> (Loc.ghost,Evar_kinds.VarInstance id)
+ | Anonymous -> Loc.tag Evar_kinds.GoalEvar
+ | Name id -> Loc.tag @@ Evar_kinds.VarInstance id
let dependent_evar_ident ev evd =
let evi = find evd ev in
match evi.evar_source with
| (_,Evar_kinds.VarInstance id) -> id
- | _ -> anomaly (str "Not an evar resulting of a dependent binding")
+ | _ -> anomaly (str "Not an evar resulting of a dependent binding.")
(**********************************************************)
(* Extra data *)
@@ -1158,10 +1100,6 @@ let set_extra_data extras evd = { evd with extras }
(*******************************************************************)
-type pending = (* before: *) evar_map * (* after: *) evar_map
-
-type pending_constr = pending * constr
-
type open_constr = evar_map * constr
(*******************************************************************)
@@ -1225,280 +1163,3 @@ module Monad =
(* Failure explanation *)
type unsolvability_explanation = SeveralInstancesFound of int
-
-(**********************************************************)
-(* Pretty-printing *)
-
-let pr_evar_suggested_name evk sigma =
- let base_id evk' evi =
- match evar_ident evk' sigma with
- | Some id -> id
- | None -> match evi.evar_source with
- | _,Evar_kinds.ImplicitArg (c,(n,Some id),b) -> id
- | _,Evar_kinds.VarInstance id -> id
- | _,Evar_kinds.GoalEvar -> Id.of_string "Goal"
- | _ ->
- let env = reset_with_named_context evi.evar_hyps (Global.env()) in
- Namegen.id_of_name_using_hdchar env evi.evar_concl Anonymous
- in
- let names = EvMap.mapi base_id sigma.undf_evars in
- let id = EvMap.find evk names in
- let fold evk' id' (seen, n) =
- if seen then (seen, n)
- else if Evar.equal evk evk' then (true, n)
- else if Id.equal id id' then (seen, succ n)
- else (seen, n)
- in
- let (_, n) = EvMap.fold fold names (false, 0) in
- if n = 0 then id else Nameops.add_suffix id (string_of_int (pred n))
-
-let pr_existential_key sigma evk = match evar_ident evk sigma with
-| None ->
- str "?" ++ pr_id (pr_evar_suggested_name evk sigma)
-| Some id ->
- str "?" ++ pr_id id
-
-let pr_instance_status (sc,typ) =
- begin match sc with
- | IsSubType -> str " [or a subtype of it]"
- | IsSuperType -> str " [or a supertype of it]"
- | Conv -> mt ()
- end ++
- begin match typ with
- | CoerceToType -> str " [up to coercion]"
- | TypeNotProcessed -> mt ()
- | TypeProcessed -> str " [type is checked]"
- end
-
-let protect f x =
- try f x
- with e -> str "EXCEPTION: " ++ str (Printexc.to_string e)
-
-let print_constr a = protect print_constr a
-
-let pr_meta_map mmap =
- let pr_name = function
- Name id -> str"[" ++ pr_id id ++ str"]"
- | _ -> mt() in
- let pr_meta_binding = function
- | (mv,Cltyp (na,b)) ->
- hov 0
- (pr_meta mv ++ pr_name na ++ str " : " ++
- print_constr b.rebus ++ fnl ())
- | (mv,Clval(na,(b,s),t)) ->
- hov 0
- (pr_meta mv ++ pr_name na ++ str " := " ++
- print_constr b.rebus ++
- str " : " ++ print_constr t.rebus ++
- spc () ++ pr_instance_status s ++ fnl ())
- in
- prlist pr_meta_binding (metamap_to_list mmap)
-
-let pr_decl (decl,ok) =
- let id = get_id decl in
- match get_value decl with
- | None -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
- | Some c -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
- print_constr c ++ str (if ok then ")" else "}")
-
-let pr_evar_source = function
- | Evar_kinds.QuestionMark _ -> str "underscore"
- | Evar_kinds.CasesType false -> str "pattern-matching return predicate"
- | Evar_kinds.CasesType true ->
- str "subterm of pattern-matching return predicate"
- | Evar_kinds.BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
- | Evar_kinds.BinderType Anonymous -> str "type of anonymous binder"
- | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
- let id = Option.get ido in
- str "parameter " ++ pr_id id ++ spc () ++ str "of" ++
- spc () ++ print_constr (printable_constr_of_global c)
- | Evar_kinds.InternalHole -> str "internal placeholder"
- | Evar_kinds.TomatchTypeParameter (ind,n) ->
- pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind)
- | Evar_kinds.GoalEvar -> str "goal evar"
- | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause"
- | Evar_kinds.MatchingVar _ -> str "matching variable"
- | Evar_kinds.VarInstance id -> str "instance of " ++ pr_id id
- | Evar_kinds.SubEvar evk ->
- str "subterm of " ++ str (string_of_existential evk)
-
-let pr_evar_info evi =
- let phyps =
- try
- let decls = match Filter.repr (evar_filter evi) with
- | None -> List.map (fun c -> (c, true)) (evar_context evi)
- | Some filter -> List.combine (evar_context evi) filter
- in
- prlist_with_sep spc pr_decl (List.rev decls)
- with Invalid_argument _ -> str "Ill-formed filtered context" in
- let pty = print_constr evi.evar_concl in
- let pb =
- match evi.evar_body with
- | Evar_empty -> mt ()
- | Evar_defined c -> spc() ++ str"=> " ++ print_constr c
- in
- let candidates =
- match evi.evar_body, evi.evar_candidates with
- | Evar_empty, Some l ->
- spc () ++ str "{" ++
- prlist_with_sep (fun () -> str "|") print_constr l ++ str "}"
- | _ ->
- mt ()
- in
- let src = str "(" ++ pr_evar_source (snd evi.evar_source) ++ str ")" in
- hov 2
- (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
- candidates ++ spc() ++ src)
-
-let compute_evar_dependency_graph (sigma : evar_map) =
- (* Compute the map binding ev to the evars whose body depends on ev *)
- let fold evk evi acc =
- let fold_ev evk' acc =
- let tab =
- try EvMap.find evk' acc
- with Not_found -> Evar.Set.empty
- in
- EvMap.add evk' (Evar.Set.add evk tab) acc
- in
- match evar_body evi with
- | Evar_empty -> assert false
- | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc
- in
- EvMap.fold fold sigma.defn_evars EvMap.empty
-
-let evar_dependency_closure n sigma =
- (** Create the DAG of depth [n] representing the recursive dependencies of
- undefined evars. *)
- let graph = compute_evar_dependency_graph sigma in
- let rec aux n curr accu =
- if Int.equal n 0 then Evar.Set.union curr accu
- else
- let fold evk accu =
- try
- let deps = EvMap.find evk graph in
- Evar.Set.union deps accu
- with Not_found -> accu
- in
- (** Consider only the newly added evars *)
- let ncurr = Evar.Set.fold fold curr Evar.Set.empty in
- (** Merge the others *)
- let accu = Evar.Set.union curr accu in
- aux (n - 1) ncurr accu
- in
- let undef = EvMap.domain (undefined_map sigma) in
- aux n undef Evar.Set.empty
-
-let evar_dependency_closure n sigma =
- let deps = evar_dependency_closure n sigma in
- let map = EvMap.bind (fun ev -> find sigma ev) deps in
- EvMap.bindings map
-
-let has_no_evar sigma =
- EvMap.is_empty sigma.defn_evars && EvMap.is_empty sigma.undf_evars
-
-let pr_evd_level evd = pr_uctx_level evd.universes
-
-let pr_evar_universe_context ctx =
- let prl = pr_uctx_level ctx in
- if is_empty_evar_universe_context ctx then mt ()
- else
- (str"UNIVERSES:"++brk(0,1)++
- h 0 (Univ.pr_universe_context_set prl (evar_universe_context_set ctx)) ++ fnl () ++
- str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
- h 0 (Univ.LSet.pr prl (UState.algebraics ctx)) ++ fnl() ++
- str"UNDEFINED UNIVERSES:"++brk(0,1)++
- h 0 (Universes.pr_universe_opt_subst (UState.subst ctx)) ++ fnl())
-
-let print_env_short env =
- let pr_body n = function
- | None -> pr_name n
- | Some b -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")" in
- let pr_named_decl decl = pr_body (Name (get_id decl)) (get_value decl) in
- let pr_rel_decl decl = let open Context.Rel.Declaration in
- pr_body (get_name decl) (get_value decl) in
- let nc = List.rev (named_context env) in
- let rc = List.rev (rel_context env) in
- str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
- str "[" ++ pr_sequence pr_rel_decl rc ++ str "]"
-
-let pr_evar_constraints pbs =
- let pr_evconstr (pbty, env, t1, t2) =
- let env =
- (** We currently allow evar instances to refer to anonymous de
- Bruijn indices, so we protect the error printing code in this
- case by giving names to every de Bruijn variable in the
- rel_context of the conversion problem. MS: we should rather
- stop depending on anonymous variables, they can be used to
- indicate independency. Also, this depends on a strategy for
- naming/renaming. *)
- Namegen.make_all_name_different env
- in
- print_env_short env ++ spc () ++ str "|-" ++ spc () ++
- print_constr_env env t1 ++ spc () ++
- str (match pbty with
- | Reduction.CONV -> "=="
- | Reduction.CUMUL -> "<=") ++
- spc () ++ print_constr_env env t2
- in
- prlist_with_sep fnl pr_evconstr pbs
-
-let pr_evar_map_gen with_univs pr_evars sigma =
- let { universes = uvs } = sigma in
- let evs = if has_no_evar sigma then mt () else pr_evars sigma ++ fnl ()
- and svs = if with_univs then pr_evar_universe_context uvs else mt ()
- and cstrs =
- if List.is_empty sigma.conv_pbs then mt ()
- else
- str "CONSTRAINTS:" ++ brk (0, 1) ++
- pr_evar_constraints sigma.conv_pbs ++ fnl ()
- and metas =
- if Metamap.is_empty sigma.metas then mt ()
- else
- str "METAS:" ++ brk (0, 1) ++ pr_meta_map sigma.metas
- in
- evs ++ svs ++ cstrs ++ metas
-
-let pr_evar_list sigma l =
- let pr (ev, evi) =
- h 0 (str (string_of_existential ev) ++
- str "==" ++ pr_evar_info evi ++
- (if evi.evar_body == Evar_empty
- then str " {" ++ pr_existential_key sigma ev ++ str "}"
- else mt ()))
- in
- h 0 (prlist_with_sep fnl pr l)
-
-let pr_evar_by_depth depth sigma = match depth with
-| None ->
- (* Print all evars *)
- str"EVARS:"++brk(0,1)++pr_evar_list sigma (to_list sigma)++fnl()
-| Some n ->
- (* Print all evars *)
- str"UNDEFINED EVARS:"++
- (if Int.equal n 0 then mt() else str" (+level "++int n++str" closure):")++
- brk(0,1)++
- pr_evar_list sigma (evar_dependency_closure n sigma)++fnl()
-
-let pr_evar_by_filter filter sigma =
- let defined = Evar.Map.filter filter sigma.defn_evars in
- let undefined = Evar.Map.filter filter sigma.undf_evars in
- let prdef =
- if Evar.Map.is_empty defined then mt ()
- else str "DEFINED EVARS:" ++ brk (0, 1) ++
- pr_evar_list sigma (Evar.Map.bindings defined)
- in
- let prundef =
- if Evar.Map.is_empty undefined then mt ()
- else str "UNDEFINED EVARS:" ++ brk (0, 1) ++
- pr_evar_list sigma (Evar.Map.bindings undefined)
- in
- prdef ++ prundef
-
-let pr_evar_map ?(with_univs=true) depth sigma =
- pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_depth depth sigma) sigma
-
-let pr_evar_map_filter ?(with_univs=true) filter sigma =
- pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_filter filter sigma) sigma
-
-let pr_metaset metas =
- str "[" ++ pr_sequence pr_meta (Metaset.elements metas) ++ str "]"
diff --git a/engine/evd.mli b/engine/evd.mli
index 780ee5ff3..3f00a3b0b 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -149,7 +149,7 @@ val has_undefined : evar_map -> bool
there are uninstantiated evars in [sigma]. *)
val new_evar : evar_map ->
- ?naming:Misctypes.intro_pattern_naming_expr -> evar_info -> evar_map * evar
+ ?name:Id.t -> evar_info -> evar_map * evar
(** Creates a fresh evar mapping to the given information. *)
val add : evar_map -> evar -> evar_info -> evar_map
@@ -240,7 +240,7 @@ val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool ->
(** {6 Misc} *)
val restrict : evar -> Filter.t -> ?candidates:constr list ->
- evar_map -> evar_map * evar
+ ?src:Evar_kinds.t located -> evar_map -> evar_map * evar
(** Restrict an undefined evar into a new evar by filtering context and
possibly limiting the instances to a set of candidates *)
@@ -414,16 +414,12 @@ val extract_changed_conv_pbs : evar_map ->
(Evar.Set.t -> evar_constraint -> bool) ->
evar_map * evar_constraint list
val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list
-val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t
+val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t option
(** The following functions return the set of evars immediately
contained in the object; need the term to be evar-normal otherwise
defined evars are returned too. *)
-val evar_list : constr -> existential list
- (** excluding evars in instances of evars and collected with
- redundancies from right to left (used by tactic "instantiate") *)
-
val evars_of_term : constr -> Evar.Set.t
(** including evars in instances of evars *)
@@ -586,27 +582,10 @@ val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor ->
val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> env ->
evar_map -> Globnames.global_reference -> evar_map * constr
-(********************************************************************
- Conversion w.r.t. an evar map, not unifying universes. See
- [Reductionops.infer_conv] for conversion up-to universes. *)
-
-val test_conversion : env -> evar_map -> conv_pb -> constr -> constr -> bool
-(** WARNING: This does not allow unification of universes *)
-
-val eq_constr_univs : evar_map -> constr -> constr -> evar_map * bool
-(** Syntactic equality up to universes, recording the associated constraints *)
-
-val e_eq_constr_univs : evar_map ref -> constr -> constr -> bool
-(** Syntactic equality up to universes. *)
-
(********************************************************************)
(* constr with holes and pending resolution of classes, conversion *)
(* problems, candidates, etc. *)
-type pending = (* before: *) evar_map * (* after: *) evar_map
-
-type pending_constr = pending * constr
-
type open_constr = evar_map * constr (* Special case when before is empty *)
(** Partially constructed constrs. *)
@@ -614,21 +593,6 @@ type open_constr = evar_map * constr (* Special case when before is empty *)
type unsolvability_explanation = SeveralInstancesFound of int
(** Failure explanation. *)
-val pr_existential_key : evar_map -> evar -> Pp.std_ppcmds
-
-val pr_evar_suggested_name : existential_key -> evar_map -> Id.t
-
-(** {5 Debug pretty-printers} *)
-
-val pr_evar_info : evar_info -> Pp.std_ppcmds
-val pr_evar_constraints : evar_constraint list -> Pp.std_ppcmds
-val pr_evar_map : ?with_univs:bool -> int option -> evar_map -> Pp.std_ppcmds
-val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> evar_info -> bool) ->
- evar_map -> Pp.std_ppcmds
-val pr_metaset : Metaset.t -> Pp.std_ppcmds
-val pr_evar_universe_context : evar_universe_context -> Pp.std_ppcmds
-val pr_evd_level : evar_map -> Univ.Level.t -> Pp.std_ppcmds
-
(** {5 Deprecated functions} *)
val create_evar_defs : evar_map -> evar_map
diff --git a/engine/ftactic.ml b/engine/ftactic.ml
index efe828ad5..8e4c5f220 100644
--- a/engine/ftactic.ml
+++ b/engine/ftactic.ml
@@ -53,31 +53,17 @@ let bind (type a) (type b) (m : a t) (f : a -> b t) : b t = m >>= function
Proofview.tclUNIT (Depends filtered)
let goals = Proofview.Goal.goals >>= fun l -> Proofview.tclUNIT (Depends l)
-let set_sigma r =
- let Sigma.Sigma (ans, sigma, _) = r in
- Proofview.Unsafe.tclEVARS (Sigma.to_evar_map sigma) >>= fun () -> ans
let nf_enter f =
bind goals
(fun gl ->
gl >>= fun gl ->
Proofview.Goal.normalize gl >>= fun nfgl ->
- Proofview.V82.wrap_exceptions (fun () -> f.enter nfgl))
-
-let nf_s_enter f =
- bind goals
- (fun gl ->
- gl >>= fun gl ->
- Proofview.Goal.normalize gl >>= fun nfgl ->
- Proofview.V82.wrap_exceptions (fun () -> set_sigma (f.s_enter nfgl)))
+ Proofview.V82.wrap_exceptions (fun () -> f nfgl))
let enter f =
bind goals
- (fun gl -> gl >>= fun gl -> Proofview.V82.wrap_exceptions (fun () -> f.enter gl))
-
-let s_enter f =
- bind goals
- (fun gl -> gl >>= fun gl -> Proofview.V82.wrap_exceptions (fun () -> set_sigma (f.s_enter gl)))
+ (fun gl -> gl >>= fun gl -> Proofview.V82.wrap_exceptions (fun () -> f gl))
let with_env t =
t >>= function
diff --git a/engine/ftactic.mli b/engine/ftactic.mli
index 7d9b760a7..c108c0c2e 100644
--- a/engine/ftactic.mli
+++ b/engine/ftactic.mli
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Proofview.Notations
-
(** This module defines potentially focussing tactics. They are used by Ltac to
emulate the historical behaviour of always-focussed tactics while still
allowing to remain global when the goal is not needed. *)
@@ -41,20 +39,13 @@ val run : 'a t -> ('a -> unit Proofview.tactic) -> unit Proofview.tactic
(** {5 Focussing} *)
-val nf_enter : ([ `NF ], 'a t) enter -> 'a t
+val nf_enter : ([ `NF ] Proofview.Goal.t -> 'a t) -> 'a t
(** Enter a goal. The resulting tactic is focussed. *)
-val enter : ([ `LZ ], 'a t) enter -> 'a t
+val enter : ([ `LZ ] Proofview.Goal.t -> 'a t) -> 'a t
(** Enter a goal, without evar normalization. The resulting tactic is
focussed. *)
-val s_enter : ([ `LZ ], 'a t) s_enter -> 'a t
-(** Enter a goal and put back an evarmap. The resulting tactic is focussed. *)
-
-val nf_s_enter : ([ `NF ], 'a t) s_enter -> 'a t
-(** Enter a goal, without evar normalization and put back an evarmap. The
- resulting tactic is focussed. *)
-
val with_env : 'a t -> (Environ.env*'a) t
(** [with_env t] returns, in addition to the return type of [t], an
environment, which is the global environment if [t] does not focus on
diff --git a/engine/logic_monad.ml b/engine/logic_monad.ml
index 4fc248911..bf1b3e0e8 100644
--- a/engine/logic_monad.ml
+++ b/engine/logic_monad.ml
@@ -34,7 +34,7 @@ exception Timeout
exception TacticFailure of exn
let _ = CErrors.register_handler begin function
- | Timeout -> CErrors.errorlabstrm "Some timeout function" (Pp.str"Timeout!")
+ | Timeout -> CErrors.user_err ~hdr:"Some timeout function" (Pp.str"Timeout!")
| Exception e -> CErrors.print e
| TacticFailure e -> CErrors.print e
| _ -> Pervasives.raise CErrors.Unhandled
diff --git a/engine/namegen.ml b/engine/namegen.ml
index 6be5c6b32..a75fe721f 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -15,15 +15,17 @@
open Util
open Names
open Term
+open Environ
+open EConstr
open Vars
open Nametab
open Nameops
open Libnames
open Globnames
-open Environ
-open Termops
open Context.Rel.Declaration
+module RelDecl = Context.Rel.Declaration
+
(**********************************************************************)
(* Conventional names *)
@@ -76,16 +78,27 @@ let is_constructor id =
with Not_found ->
false
+let is_section_variable id =
+ try let _ = Global.lookup_named id in true
+ with Not_found -> false
+
(**********************************************************************)
(* Generating "intuitive" names from its type *)
-let head_name c = (* Find the head constant of a constr if any *)
+let global_of_constr = function
+| Const (c, _) -> ConstRef c
+| Ind (i, _) -> IndRef i
+| Construct (c, _) -> ConstructRef c
+| Var id -> VarRef id
+| _ -> assert false
+
+let head_name sigma c = (* Find the head constant of a constr if any *)
let rec hdrec c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c)
| Cast (c,_,_) | App (c,_) -> hdrec c
| Proj (kn,_) -> Some (Label.to_id (con_label (Projection.constant kn)))
- | Const _ | Ind _ | Construct _ | Var _ ->
+ | Const _ | Ind _ | Construct _ | Var _ as c ->
Some (basename_of_global (global_of_constr c))
| Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
Some (match lna.(i) with Name id -> id | _ -> assert false)
@@ -100,9 +113,9 @@ let sort_hdchar = function
| Prop(_) -> "P"
| Type(_) -> "T"
-let hdchar env c =
+let hdchar env sigma c =
let rec hdrec k c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c) -> hdrec (k+1) c
| Cast (c,_,_) | App (c,_) -> hdrec k c
| Proj (kn,_) -> lowercase_first_char (Label.to_id (con_label (Projection.constant kn)))
@@ -110,13 +123,13 @@ let hdchar env c =
| Ind (x,_) -> lowercase_first_char (basename_of_global (IndRef x))
| Construct (x,_) -> lowercase_first_char (basename_of_global (ConstructRef x))
| Var id -> lowercase_first_char id
- | Sort s -> sort_hdchar s
+ | Sort s -> sort_hdchar (ESorts.kind sigma s)
| Rel n ->
(if n<=k then "p" (* the initial term is flexible product/function *)
else
- try match Environ.lookup_rel (n-k) env |> to_tuple with
- | (Name id,_,_) -> lowercase_first_char id
- | (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
+ try match lookup_rel (n-k) env with
+ | LocalAssum (Name id,_) | LocalDef (Name id,_,_) -> lowercase_first_char id
+ | LocalAssum (Anonymous,t) | LocalDef (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
with Not_found -> "y")
| Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
let id = match lna.(i) with Name id -> id | _ -> assert false in
@@ -126,41 +139,41 @@ let hdchar env c =
in
hdrec 0 c
-let id_of_name_using_hdchar env a = function
- | Anonymous -> Id.of_string (hdchar env a)
+let id_of_name_using_hdchar env sigma a = function
+ | Anonymous -> Id.of_string (hdchar env sigma a)
| Name id -> id
-let named_hd env a = function
- | Anonymous -> Name (Id.of_string (hdchar env a))
+let named_hd env sigma a = function
+ | Anonymous -> Name (Id.of_string (hdchar env sigma a))
| x -> x
-let mkProd_name env (n,a,b) = mkProd (named_hd env a n, a, b)
-let mkLambda_name env (n,a,b) = mkLambda (named_hd env a n, a, b)
+let mkProd_name env sigma (n,a,b) = mkProd (named_hd env sigma a n, a, b)
+let mkLambda_name env sigma (n,a,b) = mkLambda (named_hd env sigma a n, a, b)
let lambda_name = mkLambda_name
let prod_name = mkProd_name
-let prod_create env (a,b) = mkProd (named_hd env a Anonymous, a, b)
-let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous, a, b)
+let prod_create env sigma (a,b) = mkProd (named_hd env sigma a Anonymous, a, b)
+let lambda_create env sigma (a,b) = mkLambda (named_hd env sigma a Anonymous, a, b)
-let name_assumption env = function
- | LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
- | LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+let name_assumption env sigma = function
+ | LocalAssum (na,t) -> LocalAssum (named_hd env sigma t na, t)
+ | LocalDef (na,c,t) -> LocalDef (named_hd env sigma c na, c, t)
-let name_context env hyps =
+let name_context env sigma hyps =
snd
(List.fold_left
(fun (env,hyps) d ->
- let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
+ let d' = name_assumption env sigma d in (push_rel d' env, d' :: hyps))
(env,[]) (List.rev hyps))
-let mkProd_or_LetIn_name env b d = mkProd_or_LetIn (name_assumption env d) b
-let mkLambda_or_LetIn_name env b d = mkLambda_or_LetIn (name_assumption env d)b
+let mkProd_or_LetIn_name env sigma b d = mkProd_or_LetIn (name_assumption env sigma d) b
+let mkLambda_or_LetIn_name env sigma b d = mkLambda_or_LetIn (name_assumption env sigma d) b
-let it_mkProd_or_LetIn_name env b hyps =
- it_mkProd_or_LetIn b (name_context env hyps)
-let it_mkLambda_or_LetIn_name env b hyps =
- it_mkLambda_or_LetIn b (name_context env hyps)
+let it_mkProd_or_LetIn_name env sigma b hyps =
+ it_mkProd_or_LetIn b (name_context env sigma hyps)
+let it_mkLambda_or_LetIn_name env sigma b hyps =
+ it_mkLambda_or_LetIn b (name_context env sigma hyps)
(**********************************************************************)
(* Fresh names *)
@@ -168,7 +181,7 @@ let it_mkLambda_or_LetIn_name env b hyps =
(* Looks for next "good" name by lifting subscript *)
let next_ident_away_from id bad =
- let rec name_rec id = if bad id then name_rec (lift_subscript id) else id in
+ let rec name_rec id = if bad id then name_rec (increment_subscript id) else id in
name_rec id
(* Restart subscript from x0 if name starts with xN, or x00 if name
@@ -180,14 +193,10 @@ let restart_subscript id =
*** make_ident id (Some 0) *** but compatibility would be lost... *)
forget_subscript id
-let rec to_avoid id = function
-| [] -> false
-| id' :: avoid -> Id.equal id id' || to_avoid id avoid
-
-let visible_ids (nenv, c) =
+let visible_ids sigma (nenv, c) =
let accu = ref (Refset_env.empty, Int.Set.empty, Id.Set.empty) in
- let rec visible_ids n c = match kind_of_term c with
- | Const _ | Ind _ | Construct _ | Var _ ->
+ let rec visible_ids n c = match EConstr.kind sigma c with
+ | Const _ | Ind _ | Construct _ | Var _ as c ->
let (gseen, vseen, ids) = !accu in
let g = global_of_constr c in
if not (Refset_env.mem g gseen) then
@@ -205,8 +214,8 @@ let visible_ids (nenv, c) =
if p > n && not (Int.Set.mem p vseen) then
let vseen = Int.Set.add p vseen in
let name =
- try Some (lookup_name_of_rel (p - n) nenv)
- with Not_found ->
+ try Some (List.nth nenv (p - n - 1))
+ with Invalid_argument _ | Failure _ ->
(* Unbound index: may happen in debug and actually also
while computing temporary implicit arguments of an
inductive type *)
@@ -217,7 +226,7 @@ let visible_ids (nenv, c) =
| _ -> ids
in
accu := (gseen, vseen, ids)
- | _ -> Constr.iter_with_binders succ visible_ids n c
+ | _ -> EConstr.iter_with_binders sigma succ visible_ids n c
in
let () = visible_ids 1 c in
let (_, _, ids) = !accu in
@@ -227,11 +236,11 @@ let visible_ids (nenv, c) =
(* 1- Looks for a fresh name for printing in cases pattern *)
-let next_name_away_in_cases_pattern env_t na avoid =
+let next_name_away_in_cases_pattern sigma env_t na avoid =
let id = match na with Name id -> id | Anonymous -> default_dependent_ident in
- let visible = visible_ids env_t in
- let bad id = to_avoid id avoid || is_constructor id
- || Id.Set.mem id visible in
+ let visible = visible_ids sigma env_t in
+ let bad id = Id.List.mem id avoid || is_constructor id
+ || Id.Set.mem id visible in
next_ident_away_from id bad
(* 2- Looks for a fresh name for introduction in goal *)
@@ -244,8 +253,8 @@ let next_name_away_in_cases_pattern env_t na avoid =
name is taken by finding a free subscript starting from 0 *)
let next_ident_away_in_goal id avoid =
- let id = if to_avoid id avoid then restart_subscript id else id in
- let bad id = to_avoid id avoid || (is_global id && not (is_section_variable id)) in
+ let id = if Id.List.mem id avoid then restart_subscript id else id in
+ let bad id = Id.List.mem id avoid || (is_global id && not (is_section_variable id)) in
next_ident_away_from id bad
let next_name_away_in_goal na avoid =
@@ -262,16 +271,16 @@ let next_name_away_in_goal na avoid =
beyond the current subscript *)
let next_global_ident_away id avoid =
- let id = if to_avoid id avoid then restart_subscript id else id in
- let bad id = to_avoid id avoid || is_global id in
+ let id = if Id.List.mem id avoid then restart_subscript id else id in
+ let bad id = Id.List.mem id avoid || is_global id in
next_ident_away_from id bad
(* 4- Looks for next fresh name outside a list; if name already used,
looks for same name with lower available subscript *)
let next_ident_away id avoid =
- if to_avoid id avoid then
- next_ident_away_from (restart_subscript id) (fun id -> to_avoid id avoid)
+ if Id.List.mem id avoid then
+ next_ident_away_from (restart_subscript id) (fun id -> Id.List.mem id avoid)
else id
let next_name_away_with_default default na avoid =
@@ -291,28 +300,31 @@ let next_name_away_with_default_using_types default na avoid t =
let next_name_away = next_name_away_with_default default_non_dependent_string
-let make_all_name_different env =
- let avoid = ref (ids_of_named_context (named_context env)) in
- process_rel_context
+let make_all_name_different env sigma =
+ (** FIXME: this is inefficient, but only used in printing *)
+ let avoid = ref (Id.Set.elements (Context.Named.to_vars (named_context env))) in
+ let sign = named_context_val env in
+ let rels = rel_context env in
+ let env0 = reset_with_named_context sign env in
+ Context.Rel.fold_outside
(fun decl newenv ->
- let (na,_,t) = to_tuple decl in
- let na = named_hd newenv t na in
+ let na = named_hd newenv sigma (RelDecl.get_type decl) (RelDecl.get_name decl) in
let id = next_name_away na !avoid in
avoid := id::!avoid;
- push_rel (set_name (Name id) decl) newenv)
- env
+ push_rel (RelDecl.set_name (Name id) decl) newenv)
+ rels ~init:env0
(* 5- Looks for next fresh name outside a list; avoids also to use names that
would clash with short name of global references; if name is already used,
looks for name of same base with lower available subscript beyond current
subscript *)
-let next_ident_away_for_default_printing env_t id avoid =
- let visible = visible_ids env_t in
- let bad id = to_avoid id avoid || Id.Set.mem id visible in
+let next_ident_away_for_default_printing sigma env_t id avoid =
+ let visible = visible_ids sigma env_t in
+ let bad id = Id.List.mem id avoid || Id.Set.mem id visible in
next_ident_away_from id bad
-let next_name_away_for_default_printing env_t na avoid =
+let next_name_away_for_default_printing sigma env_t na avoid =
let id = match na with
| Name id -> id
| Anonymous ->
@@ -320,7 +332,7 @@ let next_name_away_for_default_printing env_t na avoid =
(* taken into account by the function compute_displayed_name_in; *)
(* just in case, invent a valid name *)
default_non_dependent_ident in
- next_ident_away_for_default_printing env_t id avoid
+ next_ident_away_for_default_printing sigma env_t id avoid
(**********************************************************************)
(* Displaying terms avoiding bound variables clashes *)
@@ -345,47 +357,47 @@ type renaming_flags =
| RenamingForGoal
| RenamingElsewhereFor of (Name.t list * constr)
-let next_name_for_display flags =
+let next_name_for_display sigma flags =
match flags with
- | RenamingForCasesPattern env_t -> next_name_away_in_cases_pattern env_t
+ | RenamingForCasesPattern env_t -> next_name_away_in_cases_pattern sigma env_t
| RenamingForGoal -> next_name_away_in_goal
- | RenamingElsewhereFor env_t -> next_name_away_for_default_printing env_t
+ | RenamingElsewhereFor env_t -> next_name_away_for_default_printing sigma env_t
(* Remark: Anonymous var may be dependent in Evar's contexts *)
-let compute_displayed_name_in flags avoid na c =
+let compute_displayed_name_in sigma flags avoid na c =
match na with
- | Anonymous when noccurn 1 c ->
+ | Anonymous when noccurn sigma 1 c ->
(Anonymous,avoid)
| _ ->
- let fresh_id = next_name_for_display flags na avoid in
- let idopt = if noccurn 1 c then Anonymous else Name fresh_id in
+ let fresh_id = next_name_for_display sigma flags na avoid in
+ let idopt = if noccurn sigma 1 c then Anonymous else Name fresh_id in
(idopt, fresh_id::avoid)
-let compute_and_force_displayed_name_in flags avoid na c =
+let compute_and_force_displayed_name_in sigma flags avoid na c =
match na with
- | Anonymous when noccurn 1 c ->
+ | Anonymous when noccurn sigma 1 c ->
(Anonymous,avoid)
| _ ->
- let fresh_id = next_name_for_display flags na avoid in
+ let fresh_id = next_name_for_display sigma flags na avoid in
(Name fresh_id, fresh_id::avoid)
-let compute_displayed_let_name_in flags avoid na c =
- let fresh_id = next_name_for_display flags na avoid in
+let compute_displayed_let_name_in sigma flags avoid na c =
+ let fresh_id = next_name_for_display sigma flags na avoid in
(Name fresh_id, fresh_id::avoid)
-let rename_bound_vars_as_displayed avoid env c =
+let rename_bound_vars_as_displayed sigma avoid env c =
let rec rename avoid env c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (na,c1,c2) ->
let na',avoid' =
- compute_displayed_name_in
+ compute_displayed_name_in sigma
(RenamingElsewhereFor (env,c2)) avoid na c2 in
- mkProd (na', c1, rename avoid' (add_name na' env) c2)
+ mkProd (na', c1, rename avoid' (na' :: env) c2)
| LetIn (na,c1,t,c2) ->
let na',avoid' =
- compute_displayed_let_name_in
+ compute_displayed_let_name_in sigma
(RenamingElsewhereFor (env,c2)) avoid na c2 in
- mkLetIn (na',c1,t, rename avoid' (add_name na' env) c2)
+ mkLetIn (na',c1,t, rename avoid' (na' :: env) c2)
| Cast (c,k,t) -> mkCast (rename avoid env c, k,t)
| _ -> c
in
@@ -400,14 +412,12 @@ let rename_bound_vars_as_displayed avoid env c =
let h_based_elimination_names = ref false
-let use_h_based_elimination_names () =
- !h_based_elimination_names && Flags.version_strictly_greater Flags.V8_4
+let use_h_based_elimination_names () = !h_based_elimination_names
open Goptions
let _ = declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = true; (* remove in 8.8 *)
optname = "use of \"H\"-based proposition names in elimination tactics";
optkey = ["Standard";"Proposition";"Elimination";"Names"];
optread = (fun () -> !h_based_elimination_names);
diff --git a/engine/namegen.mli b/engine/namegen.mli
index faff2c916..14846a918 100644
--- a/engine/namegen.mli
+++ b/engine/namegen.mli
@@ -11,6 +11,8 @@
open Names
open Term
open Environ
+open Evd
+open EConstr
(*********************************************************************
Conventional default names *)
@@ -26,27 +28,27 @@ val default_dependent_ident : Id.t (* "x" *)
val lowercase_first_char : Id.t -> string
val sort_hdchar : sorts -> string
-val hdchar : env -> types -> string
-val id_of_name_using_hdchar : env -> types -> Name.t -> Id.t
-val named_hd : env -> types -> Name.t -> Name.t
-val head_name : types -> Id.t option
+val hdchar : env -> evar_map -> types -> string
+val id_of_name_using_hdchar : env -> evar_map -> types -> Name.t -> Id.t
+val named_hd : env -> evar_map -> types -> Name.t -> Name.t
+val head_name : evar_map -> types -> Id.t option
-val mkProd_name : env -> Name.t * types * types -> types
-val mkLambda_name : env -> Name.t * types * constr -> constr
+val mkProd_name : env -> evar_map -> Name.t * types * types -> types
+val mkLambda_name : env -> evar_map -> Name.t * types * constr -> constr
(** Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
-val prod_name : env -> Name.t * types * types -> types
-val lambda_name : env -> Name.t * types * constr -> constr
+val prod_name : env -> evar_map -> Name.t * types * types -> types
+val lambda_name : env -> evar_map -> Name.t * types * constr -> constr
-val prod_create : env -> types * types -> constr
-val lambda_create : env -> types * constr -> constr
-val name_assumption : env -> Context.Rel.Declaration.t -> Context.Rel.Declaration.t
-val name_context : env -> Context.Rel.t -> Context.Rel.t
+val prod_create : env -> evar_map -> types * types -> constr
+val lambda_create : env -> evar_map -> types * constr -> constr
+val name_assumption : env -> evar_map -> rel_declaration -> rel_declaration
+val name_context : env -> evar_map -> rel_context -> rel_context
-val mkProd_or_LetIn_name : env -> types -> Context.Rel.Declaration.t -> types
-val mkLambda_or_LetIn_name : env -> constr -> Context.Rel.Declaration.t -> constr
-val it_mkProd_or_LetIn_name : env -> types -> Context.Rel.t -> types
-val it_mkLambda_or_LetIn_name : env -> constr -> Context.Rel.t -> constr
+val mkProd_or_LetIn_name : env -> evar_map -> types -> rel_declaration -> types
+val mkLambda_or_LetIn_name : env -> evar_map -> constr -> rel_declaration -> constr
+val it_mkProd_or_LetIn_name : env -> evar_map -> types -> rel_context -> types
+val it_mkLambda_or_LetIn_name : env -> evar_map -> constr -> rel_context -> constr
(*********************************************************************
Fresh names *)
@@ -54,7 +56,22 @@ val it_mkLambda_or_LetIn_name : env -> constr -> Context.Rel.t -> constr
(** Avoid clashing with a name satisfying some predicate *)
val next_ident_away_from : Id.t -> (Id.t -> bool) -> Id.t
-(** Avoid clashing with a name of the given list *)
+(** [next_ident_away original_id unwanted_ids] returns a new identifier as close as possible
+ to the [original_id] while avoiding all [unwanted_ids].
+
+ In particular:
+ {ul {- if [original_id] does not appear in the list of [unwanted_ids], then [original_id] is returned.}
+ {- if [original_id] appears in the list of [unwanted_ids],
+ then this function returns a new id that:
+ {ul {- has the same {i root} as the [original_id],}
+ {- does not occur in the list of [unwanted_ids],}
+ {- has the smallest possible {i subscript}.}}}}
+
+ where by {i subscript} of some identifier we mean last part of it that is composed
+ only from (decimal) digits and by {i root} of some identifier we mean
+ the whole identifier except for the {i subscript}.
+
+ E.g. if we take [foo42], then [42] is the {i subscript}, and [foo] is the root. *)
val next_ident_away : Id.t -> Id.t list -> Id.t
(** Avoid clashing with a name already used in current module *)
@@ -83,16 +100,16 @@ type renaming_flags =
| RenamingForGoal (** avoid all globals (as in intro) *)
| RenamingElsewhereFor of (Name.t list * constr)
-val make_all_name_different : env -> env
+val make_all_name_different : env -> evar_map -> env
val compute_displayed_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_and_force_displayed_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_displayed_let_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val rename_bound_vars_as_displayed :
- Id.t list -> Name.t list -> types -> types
+ evar_map -> Id.t list -> Name.t list -> types -> types
(**********************************************************************)
(* Naming strategy for arguments in Prop when eliminating inductive types *)
diff --git a/engine/proofview.ml b/engine/proofview.ml
index cc63ee21a..c542fd976 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -16,13 +16,14 @@
open Pp
open Util
open Proofview_monad
-open Sigma.Notations
open Context.Named.Declaration
(** Main state of tactics *)
type proofview = Proofview_monad.proofview
-type entry = (Term.constr * Term.types) list
+(* The first items in pairs below are proofs (under construction).
+ The second items in the pairs below are statements that are being proved. *)
+type entry = (EConstr.constr * EConstr.types) list
(** Returns a stylised view of a proofview for use by, for instance,
ide-s. *)
@@ -35,15 +36,16 @@ let proofview p =
p.comb , p.solution
let compact el ({ solution } as pv) =
- let nf = Evarutil.nf_evar solution in
+ let nf c = Evarutil.nf_evar solution c in
+ let nf0 c = EConstr.Unsafe.to_constr (Evarutil.nf_evar solution (EConstr.of_constr c)) in
let size = Evd.fold (fun _ _ i -> i+1) solution 0 in
let new_el = List.map (fun (t,ty) -> nf t, nf ty) el in
let pruned_solution = Evd.drop_all_defined solution in
let apply_subst_einfo _ ei =
Evd.({ ei with
- evar_concl = nf ei.evar_concl;
- evar_hyps = Environ.map_named_val nf ei.evar_hyps;
- evar_candidates = Option.map (List.map nf) ei.evar_candidates }) in
+ evar_concl = nf0 ei.evar_concl;
+ evar_hyps = Environ.map_named_val nf0 ei.evar_hyps;
+ evar_candidates = Option.map (List.map nf0) ei.evar_candidates }) in
let new_solution = Evd.raw_map_undefined apply_subst_einfo pruned_solution in
let new_size = Evd.fold (fun _ _ i -> i+1) new_solution 0 in
Feedback.msg_info (Pp.str (Printf.sprintf "Evars: %d -> %d\n" size new_size));
@@ -54,7 +56,7 @@ let compact el ({ solution } as pv) =
type telescope =
| TNil of Evd.evar_map
- | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+ | TCons of Environ.env * Evd.evar_map * EConstr.types * (Evd.evar_map -> EConstr.constr -> telescope)
let typeclass_resolvable = Evd.Store.field ()
@@ -63,16 +65,14 @@ let dependent_init =
for type classes. *)
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
+ let src = Loc.tag @@ Evar_kinds.GoalEvar in
(* Main routine *)
let rec aux = function
| TNil sigma -> [], { solution = sigma; comb = []; shelf = [] }
| TCons (env, sigma, typ, t) ->
- let sigma = Sigma.Unsafe.of_evar_map sigma in
- let Sigma (econstr, sigma, _) = Evarutil.new_evar env sigma ~src ~store typ in
- let sigma = Sigma.to_evar_map sigma in
+ let (sigma, econstr) = Evarutil.new_evar env sigma ~src ~store typ in
+ let (gl, _) = EConstr.destEvar sigma econstr in
let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in
- let (gl, _) = Term.destEvar econstr in
let entry = (econstr, typ) :: ret in
entry, { solution = sol; comb = gl :: comb; shelf = [] }
in
@@ -286,7 +286,7 @@ let tclONCE = Proof.once
exception MoreThanOneSuccess
let _ = CErrors.register_handler begin function
- | MoreThanOneSuccess -> CErrors.error "This tactic has more than one success."
+ | MoreThanOneSuccess -> CErrors.user_err Pp.(str "This tactic has more than one success.")
| _ -> raise CErrors.Unhandled
end
@@ -341,7 +341,7 @@ let set_nosuchgoals_hook f = nosuchgoals_hook := f
let _ = CErrors.register_handler begin function
| NoSuchGoals n ->
let suffix = !nosuchgoals_hook n in
- CErrors.errorlabstrm ""
+ CErrors.user_err
(str "No such " ++ str (String.plural n "goal") ++ str "." ++
pr_non_empty_arg (fun x -> x) suffix)
| _ -> raise CErrors.Unhandled
@@ -421,13 +421,13 @@ let tclFOCUSID id t =
exception SizeMismatch of int*int
let _ = CErrors.register_handler begin function
- | SizeMismatch (i,_) ->
+ | SizeMismatch (i,j) ->
let open Pp in
let errmsg =
str"Incorrect number of goals" ++ spc() ++
- str"(expected "++int i++str(String.plural i " tactic") ++ str")."
+ str"(expected "++int i++str(String.plural i " tactic") ++ str", was given "++ int j++str")."
in
- CErrors.errorlabstrm "" errmsg
+ CErrors.user_err errmsg
| _ -> raise CErrors.Unhandled
end
@@ -451,6 +451,25 @@ let iter_goal i =
Solution.get >>= fun evd ->
Comb.set CList.(undefined evd (flatten (rev subgoals)))
+(** List iter but allocates a list of results *)
+let map_goal i =
+ let rev = List.rev in (* hem... Proof masks List... *)
+ let open Proof in
+ Comb.get >>= fun initial ->
+ Proof.List.fold_left begin fun (acc, subgoals as cur) goal ->
+ Solution.get >>= fun step ->
+ match Evarutil.advance step goal with
+ | None -> return cur
+ | Some goal ->
+ Comb.set [goal] >>
+ i goal >>= fun res ->
+ Proof.map (fun comb -> comb :: subgoals) Comb.get >>= fun x ->
+ return (res :: acc, x)
+ end ([],[]) initial >>= fun (results_rev, subgoals) ->
+ Solution.get >>= fun evd ->
+ Comb.set CList.(undefined evd (flatten (rev subgoals))) >>
+ return (rev results_rev)
+
(** A variant of [Monad.List.fold_left2] where the first list is the
list of focused goals. The argument tactic is executed in a focus
comprising only of the current goal, a goal which has been solved
@@ -583,7 +602,15 @@ let tclINDEPENDENT tac =
let tac = InfoL.tag (Info.DBranch) tac in
InfoL.tag (Info.Dispatch) (iter_goal (fun _ -> tac))
-
+let tclINDEPENDENTL tac =
+ let open Proof in
+ Pv.get >>= fun initial ->
+ match initial.comb with
+ | [] -> tclUNIT []
+ | [_] -> tac >>= fun x -> return [x]
+ | _ ->
+ let tac = InfoL.tag (Info.DBranch) tac in
+ InfoL.tag (Info.Dispatch) (map_goal (fun _ -> tac))
(** {7 Goal manipulation} *)
@@ -666,6 +693,12 @@ let mark_in_evm ~goal evd content =
let info =
if goal then
{ info with Evd.evar_source = match info.Evd.evar_source with
+ (* Two kinds for goal evars:
+ - GoalEvar (morally not dependent)
+ - VarInstance (morally dependent of some name).
+ This is a heuristic for naming these evars. *)
+ | loc, (Evar_kinds.QuestionMark (_,Names.Name id) |
+ Evar_kinds.ImplicitArg (_,(_,Some id),_)) -> loc, Evar_kinds.VarInstance id
| _, (Evar_kinds.VarInstance _ | Evar_kinds.GoalEvar) as x -> x
| loc,_ -> loc,Evar_kinds.GoalEvar }
else info
@@ -837,11 +870,11 @@ let tclPROGRESS t =
if not test then
tclUNIT res
else
- tclZERO (CErrors.UserError ("Proofview.tclPROGRESS" , Pp.str"Failed to progress."))
+ tclZERO (CErrors.UserError (Some "Proofview.tclPROGRESS" , Pp.str"Failed to progress."))
exception Timeout
let _ = CErrors.register_handler begin function
- | Timeout -> CErrors.errorlabstrm "Proofview.tclTIMEOUT" (Pp.str"Tactic timeout!")
+ | Timeout -> CErrors.user_err ~hdr:"Proofview.tclTIMEOUT" (Pp.str"Tactic timeout!")
| _ -> Pervasives.raise CErrors.Unhandled
end
@@ -976,31 +1009,25 @@ let catchable_exception = function
module Goal = struct
- type ('a, 'r) t = {
+ type 'a t = {
env : Environ.env;
sigma : Evd.evar_map;
- concl : Term.constr ;
+ concl : EConstr.constr ;
self : Evar.t ; (* for compatibility with old-style definitions *)
}
- type ('a, 'b) enter =
- { enter : 'r. ('a, 'r) t -> 'b }
-
- let assume (gl : ('a, 'r) t) = (gl :> ([ `NF ], 'r) t)
-
- let env { env=env } = env
- 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_extra sigma self
-
- let raw_concl { concl=concl } = concl
+ let assume (gl : 'a t) = (gl :> [ `NF ] t)
+ let env {env} = env
+ let sigma {sigma} = sigma
+ let hyps {env} = EConstr.named_context env
+ let concl {concl} = concl
+ let extra {sigma; self} = goal_extra sigma self
let gmake_with info env sigma goal =
{ env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ;
sigma = sigma ;
- concl = Evd.evar_concl info ;
+ concl = EConstr.of_constr (Evd.evar_concl info);
self = goal }
let nf_gmake env sigma goal =
@@ -1011,11 +1038,11 @@ module Goal = struct
let nf_enter f =
InfoL.tag (Info.Dispatch) begin
iter_goal begin fun goal ->
- Env.get >>= fun env ->
+ tclENV >>= fun env ->
tclEVARMAP >>= fun sigma ->
try
let (gl, sigma) = nf_gmake env sigma goal in
- tclTHEN (Unsafe.tclEVARS sigma) (InfoL.tag (Info.DBranch) (f.enter gl))
+ tclTHEN (Unsafe.tclEVARS sigma) (InfoL.tag (Info.DBranch) (f gl))
with e when catchable_exception e ->
let (e, info) = CErrors.push e in
tclZERO ~info e
@@ -1033,7 +1060,7 @@ module Goal = struct
gmake_with info env sigma goal
let enter f =
- let f gl = InfoL.tag (Info.DBranch) (f.enter gl) in
+ let f gl = InfoL.tag (Info.DBranch) (f gl) in
InfoL.tag (Info.Dispatch) begin
iter_goal begin fun goal ->
Env.get >>= fun env ->
@@ -1048,7 +1075,7 @@ module Goal = struct
exception NotExactlyOneSubgoal
let _ = CErrors.register_handler begin function
| NotExactlyOneSubgoal ->
- CErrors.errorlabstrm "" (Pp.str"Not exactly one subgoal.")
+ CErrors.user_err (Pp.str"Not exactly one subgoal.")
| _ -> raise CErrors.Unhandled
end
@@ -1058,48 +1085,13 @@ module Goal = struct
| [goal] -> begin
Env.get >>= fun env ->
tclEVARMAP >>= fun sigma ->
- try f.enter (gmake env sigma goal)
+ try f (gmake env sigma goal)
with e when catchable_exception e ->
let (e, info) = CErrors.push e in
tclZERO ~info e
end
| _ -> tclZERO NotExactlyOneSubgoal
- type ('a, 'b) s_enter =
- { s_enter : 'r. ('a, 'r) t -> ('b, 'r) Sigma.sigma }
-
- let s_enter f =
- InfoL.tag (Info.Dispatch) begin
- iter_goal begin fun goal ->
- Env.get >>= fun env ->
- tclEVARMAP >>= fun sigma ->
- try
- let gl = gmake env sigma goal in
- let Sigma (tac, sigma, _) = f.s_enter gl in
- let sigma = Sigma.to_evar_map sigma in
- tclTHEN (Unsafe.tclEVARS sigma) (InfoL.tag (Info.DBranch) tac)
- with e when catchable_exception e ->
- let (e, info) = CErrors.push e in
- tclZERO ~info e
- end
- end
-
- let nf_s_enter f =
- InfoL.tag (Info.Dispatch) begin
- iter_goal begin fun goal ->
- Env.get >>= fun env ->
- tclEVARMAP >>= fun sigma ->
- try
- let (gl, sigma) = nf_gmake env sigma goal in
- let Sigma (tac, sigma, _) = f.s_enter gl in
- let sigma = Sigma.to_evar_map sigma in
- tclTHEN (Unsafe.tclEVARS sigma) (InfoL.tag (Info.DBranch) tac)
- with e when catchable_exception e ->
- let (e, info) = CErrors.push e in
- tclZERO ~info e
- end
- end
-
let goals =
Pv.get >>= fun step ->
let sigma = step.solution in
@@ -1123,8 +1115,6 @@ module Goal = struct
(* compatibility *)
let goal { self=self } = self
- let lift (gl : ('a, 'r) t) _ = (gl :> ('a, 's) t)
-
end
@@ -1213,12 +1203,12 @@ module V82 = struct
{ Evd.it = comb ; sigma = solution }
let top_goals initial { solution=solution; } =
- let goals = CList.map (fun (t,_) -> fst (Term.destEvar t)) initial in
+ let goals = CList.map (fun (t,_) -> fst (Term.destEvar (EConstr.Unsafe.to_constr t))) initial in
{ Evd.it = goals ; sigma=solution; }
let top_evars initial =
let evars_of_initial (c,_) =
- Evar.Set.elements (Evd.evars_of_term c)
+ Evar.Set.elements (Evd.evars_of_term (EConstr.Unsafe.to_constr c))
in
CList.flatten (CList.map evars_of_initial initial)
@@ -1248,8 +1238,4 @@ module Notations = struct
let (>>=) = tclBIND
let (<*>) = tclTHEN
let (<+>) t1 t2 = tclOR t1 (fun _ -> t2)
- type ('a, 'b) enter = ('a, 'b) Goal.enter =
- { enter : 'r. ('a, 'r) Goal.t -> 'b }
- type ('a, 'b) s_enter = ('a, 'b) Goal.s_enter =
- { s_enter : 'r. ('a, 'r) Goal.t -> ('b, 'r) Sigma.sigma }
end
diff --git a/engine/proofview.mli b/engine/proofview.mli
index e8a941b08..e98f59f0f 100644
--- a/engine/proofview.mli
+++ b/engine/proofview.mli
@@ -13,7 +13,7 @@
state and returning a value of type ['a]. *)
open Util
-open Term
+open EConstr
(** Main state of tactics *)
type proofview
@@ -43,7 +43,7 @@ val compact : entry -> proofview -> entry * proofview
empty [evar_map] (indeed a proof can be triggered by an incomplete
pretyping), [init] takes an additional argument to represent the
initial [evar_map]. *)
-val init : Evd.evar_map -> (Environ.env * Term.types) list -> entry * proofview
+val init : Evd.evar_map -> (Environ.env * types) list -> entry * proofview
(** A [telescope] is a list of environment and conclusion like in
{!init}, except that each element may depend on the previous
@@ -52,7 +52,7 @@ val init : Evd.evar_map -> (Environ.env * Term.types) list -> entry * proofview
[evar_map] is threaded in state passing style. *)
type telescope =
| TNil of Evd.evar_map
- | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+ | TCons of Environ.env * Evd.evar_map * types * (Evd.evar_map -> constr -> telescope)
(** Like {!init}, but goals are allowed to be dependent on one
another. Dependencies between goals is represented with the type
@@ -292,6 +292,7 @@ val tclEXTEND : unit tactic list -> unit tactic -> unit tactic list -> unit tact
independent of backtracking in another. It is equivalent to
[tclEXTEND [] tac []]. *)
val tclINDEPENDENT : unit tactic -> unit tactic
+val tclINDEPENDENTL: 'a tactic -> 'a list tactic
(** {7 Goal manipulation} *)
@@ -468,71 +469,48 @@ module Goal : sig
data using {!assume} if you known you do not rely on the assumption of
being normalized, at your own risk.
- The second parameter is a stage indicating where the goal belongs. See
- module {!Sigma}.
*)
- type ('a, 'r) t
+ type 'a t
(** Assume that you do not need the goal to be normalized. *)
- val assume : ('a, 'r) t -> ([ `NF ], 'r) t
+ val assume : 'a t -> [ `NF ] t
(** Normalises the argument goal. *)
- val normalize : ('a, 'r) t -> ([ `NF ], 'r) t tactic
+ val normalize : 'a t -> [ `NF ] t tactic
(** [concl], [hyps], [env] and [sigma] given a goal [gl] return
respectively the conclusion of [gl], the hypotheses of [gl], the
environment of [gl] (i.e. the global environment and the
hypotheses) and the current evar map. *)
- val concl : ([ `NF ], 'r) t -> Term.constr
- val hyps : ([ `NF ], 'r) t -> Context.Named.t
- val env : ('a, 'r) t -> Environ.env
- val sigma : ('a, 'r) t -> 'r Sigma.t
- val extra : ('a, 'r) t -> Evd.Store.t
-
- (** Returns the goal's conclusion even if the goal is not
- normalised. *)
- val raw_concl : ('a, 'r) t -> Term.constr
-
- type ('a, 'b) enter =
- { enter : 'r. ('a, 'r) t -> 'b }
+ val concl : 'a t -> constr
+ val hyps : 'a t -> named_context
+ val env : 'a t -> Environ.env
+ val sigma : 'a t -> Evd.evar_map
+ val extra : 'a t -> Evd.Store.t
(** [nf_enter t] applies the goal-dependent tactic [t] in each goal
independently, in the manner of {!tclINDEPENDENT} except that
the current goal is also given as an argument to [t]. The goal
is normalised with respect to evars. *)
- val nf_enter : ([ `NF ], unit tactic) enter -> unit tactic
+ val nf_enter : ([ `NF ] t -> unit tactic) -> unit tactic
(** Like {!nf_enter}, but does not normalize the goal beforehand. *)
- val enter : ([ `LZ ], unit tactic) enter -> unit tactic
+ val enter : ([ `LZ ] t -> unit tactic) -> unit tactic
(** Like {!enter}, but assumes exactly one goal under focus, raising *)
(** an error otherwise. *)
- val enter_one : ([ `LZ ], 'a tactic) enter -> 'a tactic
-
- type ('a, 'b) s_enter =
- { s_enter : 'r. ('a, 'r) t -> ('b, 'r) Sigma.sigma }
-
- (** A variant of {!enter} allows to work with a monotonic state. The evarmap
- returned by the argument is put back into the current state before firing
- the returned tactic. *)
- val s_enter : ([ `LZ ], unit tactic) s_enter -> unit tactic
-
- (** Like {!s_enter}, but normalizes the goal beforehand. *)
- val nf_s_enter : ([ `NF ], unit tactic) s_enter -> unit tactic
+ val enter_one : ([ `LZ ] t -> 'a tactic) -> 'a tactic
(** Recover the list of current goals under focus, without evar-normalization.
FIXME: encapsulate the level in an existential type. *)
- val goals : ([ `LZ ], 'r) t tactic list tactic
+ val goals : [ `LZ ] t tactic list tactic
(** [unsolved g] is [true] if [g] is still unsolved in the current
proof state. *)
- val unsolved : ('a, 'r) t -> bool tactic
+ val unsolved : 'a t -> bool tactic
(** Compatibility: avoid if possible *)
- val goal : ([ `NF ], 'r) t -> Evar.t
-
- (** Every goal is valid at a later stage. FIXME: take a later evarmap *)
- val lift : ('a, 'r) t -> ('r, 's) Sigma.le -> ('a, 's) t
+ val goal : [ `NF ] t -> Evar.t
end
@@ -619,8 +597,4 @@ module Notations : sig
(** {!tclOR}: [t1+t2] = [tclOR t1 (fun _ -> t2)]. *)
val (<+>) : 'a tactic -> 'a tactic -> 'a tactic
- type ('a, 'b) enter = ('a, 'b) Goal.enter =
- { enter : 'r. ('a, 'r) Goal.t -> 'b }
- type ('a, 'b) s_enter = ('a, 'b) Goal.s_enter =
- { s_enter : 'r. ('a, 'r) Goal.t -> ('b, 'r) Sigma.sigma }
end
diff --git a/engine/sigma.ml b/engine/sigma.ml
deleted file mode 100644
index 9381a33af..000000000
--- a/engine/sigma.ml
+++ /dev/null
@@ -1,117 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-type 'a t = Evd.evar_map
-
-type ('a, 'b) le = unit
-
-let refl = ()
-let cons _ _ = ()
-let (+>) = fun _ _ -> ()
-
-type ('a, 'r) sigma = Sigma : 'a * 's t * ('r, 's) le -> ('a, 'r) sigma
-
-type 'a evar = Evar.t
-
-let lift_evar evk () = evk
-
-let to_evar_map evd = evd
-let to_evar evk = evk
-
-let here x s = Sigma (x, s, ())
-
-(** API *)
-
-type 'r fresh = Fresh : 's evar * 's t * ('r, 's) le -> 'r fresh
-
-let new_evar sigma ?naming info =
- let (sigma, evk) = Evd.new_evar sigma ?naming info in
- Fresh (evk, sigma, ())
-
-let define evk c sigma =
- Sigma ((), Evd.define evk c sigma, ())
-
-let new_univ_level_variable ?loc ?name rigid sigma =
- let (sigma, u) = Evd.new_univ_level_variable ?loc ?name rigid sigma in
- Sigma (u, sigma, ())
-
-let new_univ_variable ?loc ?name rigid sigma =
- let (sigma, u) = Evd.new_univ_variable ?loc ?name rigid sigma in
- Sigma (u, sigma, ())
-
-let new_sort_variable ?loc ?name rigid sigma =
- let (sigma, u) = Evd.new_sort_variable ?loc ?name rigid sigma in
- Sigma (u, sigma, ())
-
-let fresh_sort_in_family ?loc ?rigid env sigma s =
- let (sigma, s) = Evd.fresh_sort_in_family ?loc ?rigid env sigma s in
- Sigma (s, sigma, ())
-
-let fresh_constant_instance ?loc env sigma cst =
- let (sigma, cst) = Evd.fresh_constant_instance ?loc env sigma cst in
- Sigma (cst, sigma, ())
-
-let fresh_inductive_instance ?loc env sigma ind =
- let (sigma, ind) = Evd.fresh_inductive_instance ?loc env sigma ind in
- Sigma (ind, sigma, ())
-
-let fresh_constructor_instance ?loc env sigma pc =
- let (sigma, c) = Evd.fresh_constructor_instance ?loc env sigma pc in
- Sigma (c, sigma, ())
-
-let fresh_global ?loc ?rigid ?names env sigma r =
- let (sigma, c) = Evd.fresh_global ?loc ?rigid ?names env sigma r in
- Sigma (c, sigma, ())
-
-(** Run *)
-
-type 'a run = { run : 'r. 'r t -> ('a, 'r) sigma }
-
-let run sigma f : 'a * Evd.evar_map =
- let Sigma (x, sigma, ()) = f.run sigma in
- (x, sigma)
-
-(** Monotonic references *)
-
-type evdref = Evd.evar_map ref
-
-let apply evdref f =
- let Sigma (x, sigma, ()) = f.run !evdref in
- evdref := sigma;
- x
-
-let purify f =
- let f (sigma : Evd.evar_map) =
- let evdref = ref sigma in
- let ans = f evdref in
- Sigma (ans, !evdref, ())
- in
- { run = f }
-
-(** Unsafe primitives *)
-
-module Unsafe =
-struct
-
-let le = ()
-let of_evar_map sigma = sigma
-let of_evar evk = evk
-let of_ref ref = ref
-let of_pair (x, sigma) = Sigma (x, sigma, ())
-
-end
-
-module Notations =
-struct
- type ('a, 'r) sigma_ = ('a, 'r) sigma =
- Sigma : 'a * 's t * ('r, 's) le -> ('a, 'r) sigma_
-
- let (+>) = fun _ _ -> ()
-
- type 'a run_ = 'a run = { run : 'r. 'r t -> ('a, 'r) sigma }
-end
diff --git a/engine/sigma.mli b/engine/sigma.mli
deleted file mode 100644
index 7473a251b..000000000
--- a/engine/sigma.mli
+++ /dev/null
@@ -1,131 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-open Names
-open Constr
-
-(** Monotonous state enforced by typing.
-
- This module allows to constrain uses of evarmaps in a monotonous fashion,
- and in particular statically suppress evar leaks and the like. To this
- ends, it defines a type of indexed evarmaps whose phantom typing ensures
- monotonous use.
-*)
-
-(** {5 Stages} *)
-
-type ('a, 'b) le
-(** Relationship stating that stage ['a] is anterior to stage ['b] *)
-
-val refl : ('a, 'a) le
-(** Reflexivity of anteriority *)
-
-val cons : ('a, 'b) le -> ('b, 'c) le -> ('a, 'c) le
-(** Transitivity of anteriority *)
-
-val (+>) : ('a, 'b) le -> ('b, 'c) le -> ('a, 'c) le
-(** Alias for {!cons} *)
-
-(** {5 Monotonous evarmaps} *)
-
-type 'r t
-(** Stage-indexed evarmaps. This is just a plain evarmap with a phantom type. *)
-
-type ('a, 'r) sigma = Sigma : 'a * 's t * ('r, 's) le -> ('a, 'r) sigma
-(** Return values at a later stage *)
-
-type 'r evar
-(** Stage-indexed evars *)
-
-(** {5 Constructors} *)
-
-val here : 'a -> 'r t -> ('a, 'r) sigma
-(** [here x s] is a shorthand for [Sigma (x, s, refl)] *)
-
-(** {5 Postponing} *)
-
-val lift_evar : 'r evar -> ('r, 's) le -> 's evar
-(** Any evar existing at stage ['r] is also valid at any later stage. *)
-
-(** {5 Downcasting} *)
-
-val to_evar_map : 'r t -> Evd.evar_map
-val to_evar : 'r evar -> Evar.t
-
-(** {5 Monotonous API} *)
-
-type 'r fresh = Fresh : 's evar * 's t * ('r, 's) le -> 'r fresh
-
-val new_evar : 'r t -> ?naming:Misctypes.intro_pattern_naming_expr ->
- Evd.evar_info -> 'r fresh
-
-val define : 'r evar -> Constr.t -> 'r t -> (unit, 'r) sigma
-
-(** Polymorphic universes *)
-
-val new_univ_level_variable : ?loc:Loc.t -> ?name:string ->
- Evd.rigid -> 'r t -> (Univ.universe_level, 'r) sigma
-val new_univ_variable : ?loc:Loc.t -> ?name:string ->
- Evd.rigid -> 'r t -> (Univ.universe, 'r) sigma
-val new_sort_variable : ?loc:Loc.t -> ?name:string ->
- Evd.rigid -> 'r t -> (Sorts.t, 'r) sigma
-
-val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:Evd.rigid -> Environ.env ->
- 'r t -> Term.sorts_family -> (Term.sorts, 'r) sigma
-val fresh_constant_instance :
- ?loc:Loc.t -> Environ.env -> 'r t -> constant -> (pconstant, 'r) sigma
-val fresh_inductive_instance :
- ?loc:Loc.t -> Environ.env -> 'r t -> inductive -> (pinductive, 'r) sigma
-val fresh_constructor_instance : ?loc:Loc.t -> Environ.env -> 'r t -> constructor ->
- (pconstructor, 'r) sigma
-
-val fresh_global : ?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:Univ.Instance.t -> Environ.env ->
- 'r t -> Globnames.global_reference -> (constr, 'r) sigma
-
-(** FILLME *)
-
-(** {5 Run} *)
-
-type 'a run = { run : 'r. 'r t -> ('a, 'r) sigma }
-
-val run : Evd.evar_map -> 'a run -> 'a * Evd.evar_map
-
-(** {5 Imperative monotonic functions} *)
-
-type evdref
-(** Monotonic references over evarmaps *)
-
-val apply : evdref -> 'a run -> 'a
-(** Apply a monotonic function on a reference. *)
-
-val purify : (evdref -> 'a) -> 'a run
-(** Converse of {!apply}. *)
-
-(** {5 Unsafe primitives} *)
-
-module Unsafe :
-sig
- val le : ('a, 'b) le
- val of_evar_map : Evd.evar_map -> 'r t
- val of_evar : Evd.evar -> 'r evar
- val of_ref : Evd.evar_map ref -> evdref
- val of_pair : ('a * Evd.evar_map) -> ('a, 'r) sigma
-end
-
-(** {5 Notations} *)
-
-module Notations :
-sig
- type ('a, 'r) sigma_ = ('a, 'r) sigma =
- Sigma : 'a * 's t * ('r, 's) le -> ('a, 'r) sigma_
-
- type 'a run_ = 'a run = { run : 'r. 'r t -> ('a, 'r) sigma }
-
- val (+>) : ('a, 'b) le -> ('b, 'c) le -> ('a, 'c) le
- (** Alias for {!cons} *)
-end
diff --git a/engine/termops.ml b/engine/termops.ml
index 624289786..fc3291df1 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -17,6 +17,7 @@ open Environ
module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
+module CompactedDecl = Context.Compacted.Declaration
(* Sorts and sort family *)
@@ -30,10 +31,6 @@ let pr_sort_family = function
| InProp -> (str "Prop")
| InType -> (str "Type")
-let pr_name = function
- | Name id -> pr_id id
- | Anonymous -> str "_"
-
let pr_con sp = str(string_of_con sp)
let pr_fix pr_constr ((t,i),(lna,tl,bl)) =
@@ -41,7 +38,7 @@ let pr_fix pr_constr ((t,i),(lna,tl,bl)) =
hov 1
(str"fix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
- pr_name na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
+ Name.print na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
str"}")
@@ -64,10 +61,10 @@ let rec pr_constr c = match kind_of_term c with
(str"(" ++ pr_constr t ++ str " ->" ++ spc() ++
pr_constr c ++ str")")
| Lambda (na,t,c) -> hov 1
- (str"fun " ++ pr_name na ++ str":" ++
+ (str"fun " ++ Name.print na ++ str":" ++
pr_constr t ++ str" =>" ++ spc() ++ pr_constr c)
| LetIn (na,b,t,c) -> hov 0
- (str"let " ++ pr_name na ++ str":=" ++ pr_constr b ++
+ (str"let " ++ Name.print na ++ str":=" ++ pr_constr b ++
str":" ++ brk(1,2) ++ pr_constr t ++ cut() ++
pr_constr c)
| App (c,l) -> hov 1
@@ -92,24 +89,350 @@ let rec pr_constr c = match kind_of_term c with
hov 1
(str"cofix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,ty,bd) ->
- pr_name na ++ str":" ++ pr_constr ty ++
+ Name.print na ++ str":" ++ pr_constr ty ++
cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
str"}")
-let term_printer = ref (fun _ -> pr_constr)
-let print_constr_env t = !term_printer t
-let print_constr t = !term_printer (Global.env()) t
+let term_printer = ref (fun _env _sigma c -> pr_constr (EConstr.Unsafe.to_constr c))
+let print_constr_env env sigma t = !term_printer env sigma t
+let print_constr t = !term_printer (Global.env()) Evd.empty t
let set_print_constr f = term_printer := f
+module EvMap = Evar.Map
+
+let pr_evar_suggested_name evk sigma =
+ let open Evd in
+ let base_id evk' evi =
+ match evar_ident evk' sigma with
+ | Some id -> id
+ | None -> match evi.evar_source with
+ | _,Evar_kinds.ImplicitArg (c,(n,Some id),b) -> id
+ | _,Evar_kinds.VarInstance id -> id
+ | _,Evar_kinds.QuestionMark (_,Name id) -> id
+ | _,Evar_kinds.GoalEvar -> Id.of_string "Goal"
+ | _ ->
+ let env = reset_with_named_context evi.evar_hyps (Global.env()) in
+ Namegen.id_of_name_using_hdchar env sigma (EConstr.of_constr evi.evar_concl) Anonymous
+ in
+ let names = EvMap.mapi base_id (undefined_map sigma) in
+ let id = EvMap.find evk names in
+ let fold evk' id' (seen, n) =
+ if seen then (seen, n)
+ else if Evar.equal evk evk' then (true, n)
+ else if Id.equal id id' then (seen, succ n)
+ else (seen, n)
+ in
+ let (_, n) = EvMap.fold fold names (false, 0) in
+ if n = 0 then id else Nameops.add_suffix id (string_of_int (pred n))
+
+let pr_existential_key sigma evk =
+let open Evd in
+match evar_ident evk sigma with
+| None ->
+ str "?" ++ pr_id (pr_evar_suggested_name evk sigma)
+| Some id ->
+ str "?" ++ pr_id id
+
+let pr_instance_status (sc,typ) =
+ let open Evd in
+ begin match sc with
+ | IsSubType -> str " [or a subtype of it]"
+ | IsSuperType -> str " [or a supertype of it]"
+ | Conv -> mt ()
+ end ++
+ begin match typ with
+ | CoerceToType -> str " [up to coercion]"
+ | TypeNotProcessed -> mt ()
+ | TypeProcessed -> str " [type is checked]"
+ end
+
+let protect f x =
+ try f x
+ with e -> str "EXCEPTION: " ++ str (Printexc.to_string e)
+
+let print_kconstr a =
+ protect (fun c -> print_constr (EConstr.of_constr c)) a
+
+let pr_meta_map evd =
+ let open Evd in
+ let print_constr = print_kconstr in
+ let pr_name = function
+ Name id -> str"[" ++ pr_id id ++ str"]"
+ | _ -> mt() in
+ let pr_meta_binding = function
+ | (mv,Cltyp (na,b)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " : " ++
+ print_constr b.rebus ++ fnl ())
+ | (mv,Clval(na,(b,s),t)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " := " ++
+ print_constr b.rebus ++
+ str " : " ++ print_constr t.rebus ++
+ spc () ++ pr_instance_status s ++ fnl ())
+ in
+ prlist pr_meta_binding (meta_list evd)
+
+let pr_decl (decl,ok) =
+ let open NamedDecl in
+ let print_constr = print_kconstr in
+ match decl with
+ | LocalAssum (id,_) -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
+ | LocalDef (id,c,_) -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
+ print_constr c ++ str (if ok then ")" else "}")
+
+let pr_evar_source = function
+ | Evar_kinds.NamedHole id -> pr_id id
+ | Evar_kinds.QuestionMark _ -> str "underscore"
+ | Evar_kinds.CasesType false -> str "pattern-matching return predicate"
+ | Evar_kinds.CasesType true ->
+ str "subterm of pattern-matching return predicate"
+ | Evar_kinds.BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
+ | Evar_kinds.BinderType Anonymous -> str "type of anonymous binder"
+ | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
+ let open Globnames in
+ let print_constr = print_kconstr in
+ let id = Option.get ido in
+ str "parameter " ++ pr_id id ++ spc () ++ str "of" ++
+ spc () ++ print_constr (printable_constr_of_global c)
+ | Evar_kinds.InternalHole -> str "internal placeholder"
+ | Evar_kinds.TomatchTypeParameter (ind,n) ->
+ let print_constr = print_kconstr in
+ pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind)
+ | Evar_kinds.GoalEvar -> str "goal evar"
+ | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause"
+ | Evar_kinds.MatchingVar _ -> str "matching variable"
+ | Evar_kinds.VarInstance id -> str "instance of " ++ pr_id id
+ | Evar_kinds.SubEvar evk ->
+ let open Evd in
+ str "subterm of " ++ str (string_of_existential evk)
+
+let pr_evar_info evi =
+ let open Evd in
+ let print_constr = print_kconstr in
+ let phyps =
+ try
+ let decls = match Filter.repr (evar_filter evi) with
+ | None -> List.map (fun c -> (c, true)) (evar_context evi)
+ | Some filter -> List.combine (evar_context evi) filter
+ in
+ prlist_with_sep spc pr_decl (List.rev decls)
+ with Invalid_argument _ -> str "Ill-formed filtered context" in
+ let pty = print_constr evi.evar_concl in
+ let pb =
+ match evi.evar_body with
+ | Evar_empty -> mt ()
+ | Evar_defined c -> spc() ++ str"=> " ++ print_constr c
+ in
+ let candidates =
+ match evi.evar_body, evi.evar_candidates with
+ | Evar_empty, Some l ->
+ spc () ++ str "{" ++
+ prlist_with_sep (fun () -> str "|") print_constr l ++ str "}"
+ | _ ->
+ mt ()
+ in
+ let src = str "(" ++ pr_evar_source (snd evi.evar_source) ++ str ")" in
+ hov 2
+ (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
+ candidates ++ spc() ++ src)
+
+let compute_evar_dependency_graph sigma =
+ let open Evd in
+ (* Compute the map binding ev to the evars whose body depends on ev *)
+ let fold evk evi acc =
+ let fold_ev evk' acc =
+ let tab =
+ try EvMap.find evk' acc
+ with Not_found -> Evar.Set.empty
+ in
+ EvMap.add evk' (Evar.Set.add evk tab) acc
+ in
+ match evar_body evi with
+ | Evar_empty -> acc
+ | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc
+ in
+ Evd.fold fold sigma EvMap.empty
+
+let evar_dependency_closure n sigma =
+ let open Evd in
+ (** Create the DAG of depth [n] representing the recursive dependencies of
+ undefined evars. *)
+ let graph = compute_evar_dependency_graph sigma in
+ let rec aux n curr accu =
+ if Int.equal n 0 then Evar.Set.union curr accu
+ else
+ let fold evk accu =
+ try
+ let deps = EvMap.find evk graph in
+ Evar.Set.union deps accu
+ with Not_found -> accu
+ in
+ (** Consider only the newly added evars *)
+ let ncurr = Evar.Set.fold fold curr Evar.Set.empty in
+ (** Merge the others *)
+ let accu = Evar.Set.union curr accu in
+ aux (n - 1) ncurr accu
+ in
+ let undef = EvMap.domain (undefined_map sigma) in
+ aux n undef Evar.Set.empty
+
+let evar_dependency_closure n sigma =
+ let open Evd in
+ let deps = evar_dependency_closure n sigma in
+ let map = EvMap.bind (fun ev -> find sigma ev) deps in
+ EvMap.bindings map
+
+let has_no_evar sigma =
+ try let () = Evd.fold (fun _ _ () -> raise Exit) sigma () in true
+ with Exit -> false
+
+let pr_evd_level evd = UState.pr_uctx_level (Evd.evar_universe_context evd)
+
+let pr_evar_universe_context ctx =
+ let open UState in
+ let open Evd in
+ let prl = pr_uctx_level ctx in
+ if UState.is_empty ctx then mt ()
+ else
+ (str"UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.pr_universe_context_set prl (evar_universe_context_set ctx)) ++ fnl () ++
+ str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.LSet.pr prl (UState.algebraics ctx)) ++ fnl() ++
+ str"UNDEFINED UNIVERSES:"++brk(0,1)++
+ h 0 (Universes.pr_universe_opt_subst (UState.subst ctx)) ++ fnl())
+
+let print_env_short env =
+ let print_constr = print_kconstr in
+ let pr_rel_decl = function
+ | RelDecl.LocalAssum (n,_) -> Name.print n
+ | RelDecl.LocalDef (n,b,_) -> str "(" ++ Name.print n ++ str " := " ++ print_constr b ++ str ")"
+ in
+ let pr_named_decl = NamedDecl.to_rel_decl %> pr_rel_decl in
+ let nc = List.rev (named_context env) in
+ let rc = List.rev (rel_context env) in
+ str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
+ str "[" ++ pr_sequence pr_rel_decl rc ++ str "]"
+
+let pr_evar_constraints sigma pbs =
+ let pr_evconstr (pbty, env, t1, t2) =
+ let env =
+ (** We currently allow evar instances to refer to anonymous de
+ Bruijn indices, so we protect the error printing code in this
+ case by giving names to every de Bruijn variable in the
+ rel_context of the conversion problem. MS: we should rather
+ stop depending on anonymous variables, they can be used to
+ indicate independency. Also, this depends on a strategy for
+ naming/renaming. *)
+ Namegen.make_all_name_different env sigma
+ in
+ print_env_short env ++ spc () ++ str "|-" ++ spc () ++
+ print_constr_env env sigma (EConstr.of_constr t1) ++ spc () ++
+ str (match pbty with
+ | Reduction.CONV -> "=="
+ | Reduction.CUMUL -> "<=") ++
+ spc () ++ print_constr_env env Evd.empty (EConstr.of_constr t2)
+ in
+ prlist_with_sep fnl pr_evconstr pbs
+
+let pr_evar_map_gen with_univs pr_evars sigma =
+ let uvs = Evd.evar_universe_context sigma in
+ let (_, conv_pbs) = Evd.extract_all_conv_pbs sigma in
+ let evs = if has_no_evar sigma then mt () else pr_evars sigma ++ fnl ()
+ and svs = if with_univs then pr_evar_universe_context uvs else mt ()
+ and cstrs =
+ if List.is_empty conv_pbs then mt ()
+ else
+ str "CONSTRAINTS:" ++ brk (0, 1) ++
+ pr_evar_constraints sigma conv_pbs ++ fnl ()
+ and metas =
+ if List.is_empty (Evd.meta_list sigma) then mt ()
+ else
+ str "METAS:" ++ brk (0, 1) ++ pr_meta_map sigma
+ in
+ evs ++ svs ++ cstrs ++ metas
+
+let pr_evar_list sigma l =
+ let open Evd in
+ let pr (ev, evi) =
+ h 0 (str (string_of_existential ev) ++
+ str "==" ++ pr_evar_info evi ++
+ (if evi.evar_body == Evar_empty
+ then str " {" ++ pr_existential_key sigma ev ++ str "}"
+ else mt ()))
+ in
+ h 0 (prlist_with_sep fnl pr l)
+
+let pr_evar_by_depth depth sigma = match depth with
+| None ->
+ (* Print all evars *)
+ let to_list d =
+ let open Evd in
+ (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
+ let l = ref [] in
+ let fold_def evk evi () = match evi.evar_body with
+ | Evar_defined _ -> l := (evk, evi) :: !l
+ | Evar_empty -> ()
+ in
+ let fold_undef evk evi () = match evi.evar_body with
+ | Evar_empty -> l := (evk, evi) :: !l
+ | Evar_defined _ -> ()
+ in
+ Evd.fold fold_def d ();
+ Evd.fold fold_undef d ();
+ !l
+ in
+ str"EVARS:"++brk(0,1)++pr_evar_list sigma (to_list sigma)++fnl()
+| Some n ->
+ (* Print all evars *)
+ str"UNDEFINED EVARS:"++
+ (if Int.equal n 0 then mt() else str" (+level "++int n++str" closure):")++
+ brk(0,1)++
+ pr_evar_list sigma (evar_dependency_closure n sigma)++fnl()
+
+let pr_evar_by_filter filter sigma =
+ let open Evd in
+ let elts = Evd.fold (fun evk evi accu -> (evk, evi) :: accu) sigma [] in
+ let elts = List.rev elts in
+ let is_def (_, evi) = match evi.evar_body with
+ | Evar_defined _ -> true
+ | Evar_empty -> false
+ in
+ let (defined, undefined) = List.partition is_def elts in
+ let filter (evk, evi) = filter evk evi in
+ let defined = List.filter filter defined in
+ let undefined = List.filter filter undefined in
+ let prdef =
+ if List.is_empty defined then mt ()
+ else str "DEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma defined
+ in
+ let prundef =
+ if List.is_empty undefined then mt ()
+ else str "UNDEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma undefined
+ in
+ prdef ++ prundef
+
+let pr_evar_map ?(with_univs=true) depth sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_depth depth sigma) sigma
+
+let pr_evar_map_filter ?(with_univs=true) filter sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_filter filter sigma) sigma
+
+let pr_metaset metas =
+ str "[" ++ pr_sequence pr_meta (Evd.Metaset.elements metas) ++ str "]"
+
let pr_var_decl env decl =
let open NamedDecl in
let pbody = match decl with
| LocalAssum _ -> mt ()
| LocalDef (_,c,_) ->
(* Force evaluation *)
- let pb = print_constr_env env c in
+ let c = EConstr.of_constr c in
+ let pb = print_constr_env env Evd.empty c in
(str" := " ++ pb ++ cut () ) in
- let pt = print_constr_env env (get_type decl) in
+ let pt = print_constr_env env Evd.empty (EConstr.of_constr (get_type decl)) in
let ptyp = (str" : " ++ pt) in
(pr_id (get_id decl) ++ hov 0 (pbody ++ ptyp))
@@ -119,9 +442,10 @@ let pr_rel_decl env decl =
| LocalAssum _ -> mt ()
| LocalDef (_,c,_) ->
(* Force evaluation *)
- let pb = print_constr_env env c in
+ let c = EConstr.of_constr c in
+ let pb = print_constr_env env Evd.empty c in
(str":=" ++ spc () ++ pb ++ spc ()) in
- let ptyp = print_constr_env env (get_type decl) in
+ let ptyp = print_constr_env env Evd.empty (EConstr.of_constr (get_type decl)) in
match get_name decl with
| Anonymous -> hov 0 (str"<>" ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
| Name id -> hov 0 (pr_id id ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
@@ -157,6 +481,7 @@ let print_env env =
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
let rel_list n m =
+ let open EConstr in
let rec reln l p =
if p>m then l else reln (mkRel(n+p)::l) (p+1)
in
@@ -164,6 +489,7 @@ let rel_list n m =
let push_rel_assum (x,t) env =
let open RelDecl in
+ let open EConstr in
push_rel (LocalAssum (x,t)) env
let push_rels_assum assums =
@@ -177,7 +503,7 @@ let push_named_rec_types (lna,typarray,_) env =
(fun i na t ->
match na with
| Name id -> LocalAssum (id, lift i t)
- | Anonymous -> anomaly (Pp.str "Fix declarations must be named"))
+ | Anonymous -> anomaly (Pp.str "Fix declarations must be named."))
lna typarray in
Array.fold_left
(fun e assum -> push_named assum e) env ctxt
@@ -197,21 +523,17 @@ let lookup_rel_id id sign =
lookrec 1 sign
(* Constructs either [forall x:t, c] or [let x:=b:t in c] *)
-let mkProd_or_LetIn decl c =
- let open RelDecl in
- match decl with
- | LocalAssum (na,t) -> mkProd (na, t, c)
- | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
-
+let mkProd_or_LetIn = EConstr.mkProd_or_LetIn
(* Constructs either [forall x:t, c] or [c] in which [x] is replaced by [b] *)
let mkProd_wo_LetIn decl c =
+ let open EConstr in
let open RelDecl in
match decl with
| LocalAssum (na,t) -> mkProd (na, t, c)
- | LocalDef (_,b,_) -> subst1 b c
+ | LocalDef (_,b,_) -> Vars.subst1 b c
-let it_mkProd init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
-let it_mkLambda init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
+let it_mkProd init = List.fold_left (fun c (n,t) -> EConstr.mkProd (n, t, c)) init
+let it_mkLambda init = List.fold_left (fun c (n,t) -> EConstr.mkLambda (n, t, c)) init
let it_named_context_quantifier f ~init =
List.fold_left (fun c d -> f d c) init
@@ -219,9 +541,9 @@ let it_named_context_quantifier f ~init =
let it_mkProd_or_LetIn init = it_named_context_quantifier mkProd_or_LetIn ~init
let it_mkProd_wo_LetIn init = it_named_context_quantifier mkProd_wo_LetIn ~init
let it_mkLambda_or_LetIn init = it_named_context_quantifier mkLambda_or_LetIn ~init
-let it_mkNamedProd_or_LetIn init = it_named_context_quantifier mkNamedProd_or_LetIn ~init
+let it_mkNamedProd_or_LetIn init = it_named_context_quantifier EConstr.mkNamedProd_or_LetIn ~init
let it_mkNamedProd_wo_LetIn init = it_named_context_quantifier mkNamedProd_wo_LetIn ~init
-let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier mkNamedLambda_or_LetIn ~init
+let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier EConstr.mkNamedLambda_or_LetIn ~init
let it_mkLambda_or_LetIn_from_no_LetIn c decls =
let open RelDecl in
@@ -234,37 +556,39 @@ let it_mkLambda_or_LetIn_from_no_LetIn c decls =
(* *)
(* strips head casts and flattens head applications *)
-let rec strip_head_cast c = match kind_of_term c with
+let rec strip_head_cast sigma c = match EConstr.kind sigma c with
| App (f,cl) ->
- let rec collapse_rec f cl2 = match kind_of_term f with
+ let rec collapse_rec f cl2 = match EConstr.kind sigma f with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
| Cast (c,_,_) -> collapse_rec c cl2
- | _ -> if Int.equal (Array.length cl2) 0 then f else mkApp (f,cl2)
+ | _ -> if Int.equal (Array.length cl2) 0 then f else EConstr.mkApp (f,cl2)
in
collapse_rec f cl
- | Cast (c,_,_) -> strip_head_cast c
+ | Cast (c,_,_) -> strip_head_cast sigma c
| _ -> c
-let rec drop_extra_implicit_args c = match kind_of_term c with
+let rec drop_extra_implicit_args sigma c = match EConstr.kind sigma c with
(* Removed trailing extra implicit arguments, what improves compatibility
for constants with recently added maximal implicit arguments *)
- | App (f,args) when isEvar (Array.last args) ->
- drop_extra_implicit_args
+ | App (f,args) when EConstr.isEvar sigma (Array.last args) ->
+ let open EConstr in
+ drop_extra_implicit_args sigma
(mkApp (f,fst (Array.chop (Array.length args - 1) args)))
| _ -> c
(* Get the last arg of an application *)
-let last_arg c = match kind_of_term c with
+let last_arg sigma c = match EConstr.kind sigma c with
| App (f,cl) -> Array.last cl
- | _ -> anomaly (Pp.str "last_arg")
+ | _ -> anomaly (Pp.str "last_arg.")
(* Get the last arg of an application *)
-let decompose_app_vect c =
- match kind_of_term c with
+let decompose_app_vect sigma c =
+ match EConstr.kind sigma c with
| App (f,cl) -> (f, cl)
| _ -> (c,[||])
let adjust_app_list_size f1 l1 f2 l2 =
+ let open EConstr in
let len1 = List.length l1 and len2 = List.length l2 in
if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
@@ -275,38 +599,15 @@ let adjust_app_list_size f1 l1 f2 l2 =
(applist (f1,extras), restl1, f2, l2)
let adjust_app_array_size f1 l1 f2 l2 =
+ let open EConstr in
let len1 = Array.length l1 and len2 = Array.length l2 in
if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
let extras,restl2 = Array.chop (len2-len1) l2 in
- (f1, l1, appvect (f2,extras), restl2)
+ (f1, l1, mkApp (f2,extras), restl2)
else
let extras,restl1 = Array.chop (len1-len2) l1 in
- (appvect (f1,extras), restl1, f2, l2)
-
-(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
- subterms of [c]; it carries an extra data [l] (typically a name
- list) which is processed by [g na] (which typically cons [na] to
- [l]) at each binder traversal (with name [na]); it is not recursive
- and the order with which subterms are processed is not specified *)
-
-let map_constr_with_named_binders g f l c = match kind_of_term c with
- | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
- | Construct _) -> c
- | Cast (c,k,t) -> mkCast (f l c, k, f l t)
- | Prod (na,t,c) -> mkProd (na, f l t, f (g na l) c)
- | Lambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c)
- | LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c)
- | App (c,al) -> mkApp (f l c, Array.map (f l) al)
- | Proj (p,c) -> mkProj (p, f l c)
- | Evar (e,al) -> mkEvar (e, Array.map (f l) al)
- | Case (ci,p,c,bl) -> mkCase (ci, f l p, f l c, Array.map (f l) bl)
- | Fix (ln,(lna,tl,bl)) ->
- let l' = Array.fold_left (fun l na -> g na l) l lna in
- mkFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
- | CoFix(ln,(lna,tl,bl)) ->
- let l' = Array.fold_left (fun l na -> g na l) l lna in
- mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl))
+ (mkApp (f1,extras), restl1, f2, l2)
(* [map_constr_with_binders_left_to_right g f n c] maps [f n] on the
immediate subterms of [c]; it carries an extra data [n] (typically
@@ -319,6 +620,8 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with
(co-)fixpoint) *)
let fold_rec_types g (lna,typarray,_) e =
+ let open EConstr in
+ let open Vars in
let ctxt = Array.map2_i (fun i na t -> RelDecl.LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> g assum e) e ctxt
@@ -334,9 +637,10 @@ let map_left2 f a g b =
r, s
end
-let map_constr_with_binders_left_to_right g f l c =
+let map_constr_with_binders_left_to_right sigma g f l c =
let open RelDecl in
- match kind_of_term c with
+ let open EConstr in
+ match EConstr.kind sigma c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
| Cast (b,k,t) ->
@@ -398,9 +702,10 @@ let map_constr_with_binders_left_to_right g f l c =
else mkCoFix (ln,(lna,tl',bl'))
(* strong *)
-let map_constr_with_full_binders g f l cstr =
+let map_constr_with_full_binders sigma g f l cstr =
+ let open EConstr in
let open RelDecl in
- match kind_of_term cstr with
+ match EConstr.kind sigma cstr with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> cstr
| Cast (c,k, t) ->
@@ -409,16 +714,16 @@ let map_constr_with_full_binders g f l cstr =
if c==c' && t==t' then cstr else mkCast (c', k, t')
| Prod (na,t,c) ->
let t' = f l t in
- let c' = f (g (LocalAssum (na,t)) l) c in
+ let c' = f (g (LocalAssum (na, t)) l) c in
if t==t' && c==c' then cstr else mkProd (na, t', c')
| Lambda (na,t,c) ->
let t' = f l t in
- let c' = f (g (LocalAssum (na,t)) l) c in
+ let c' = f (g (LocalAssum (na, t)) l) c in
if t==t' && c==c' then cstr else mkLambda (na, t', c')
| LetIn (na,b,t,c) ->
let b' = f l b in
let t' = f l t in
- let c' = f (g (LocalDef (na,b,t)) l) c in
+ let c' = f (g (LocalDef (na, b, t)) l) c in
if b==b' && t==t' && c==c' then cstr else mkLetIn (na, b', t', c')
| App (c,al) ->
let c' = f l c in
@@ -439,7 +744,7 @@ let map_constr_with_full_binders g f l cstr =
| Fix (ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na, t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -447,7 +752,7 @@ let map_constr_with_full_binders g f l cstr =
| CoFix(ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na, t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -460,30 +765,31 @@ let map_constr_with_full_binders g f l cstr =
index) which is processed by [g] (which typically add 1 to [n]) at
each binder traversal; it is not recursive *)
-let fold_constr_with_full_binders g f n acc c =
+let fold_constr_with_full_binders sigma g f n acc c =
let open RelDecl in
- match kind_of_term c with
+ let inj c = EConstr.Unsafe.to_constr c in
+ match EConstr.kind sigma c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> acc
| Cast (c,_, t) -> f n (f n acc c) t
- | Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
- | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
- | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c
+ | Prod (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
+ | Lambda (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g (LocalDef (na, inj b, inj t)) n) (f n (f n acc b) t) c
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
| Proj (p,c) -> f n acc c
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
| CoFix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
-let fold_constr_with_binders g f n acc c =
- fold_constr_with_full_binders (fun _ x -> g x) f n acc c
+let fold_constr_with_binders sigma g f n acc c =
+ fold_constr_with_full_binders sigma (fun _ x -> g x) f n acc c
(* [iter_constr_with_full_binders g f acc c] iters [f acc] on the immediate
subterms of [c]; it carries an extra data [acc] which is processed by [g] at
@@ -518,29 +824,29 @@ let iter_constr_with_full_binders g f l c =
exception Occur
-let occur_meta c =
- let rec occrec c = match kind_of_term c with
+let occur_meta sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Meta _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_existential c =
- let rec occrec c = match kind_of_term c with
+let occur_existential sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Evar _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_meta_or_existential c =
- let rec occrec c = match kind_of_term c with
+let occur_meta_or_existential sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Evar _ -> raise Occur
| Meta _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_evar n c =
- let rec occur_rec c = match kind_of_term c with
+let occur_evar sigma n c =
+ let rec occur_rec c = match EConstr.kind sigma c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
- | _ -> iter_constr occur_rec c
+ | _ -> EConstr.iter sigma occur_rec c
in
try occur_rec c; false with Occur -> true
@@ -548,55 +854,55 @@ let occur_in_global env id constr =
let vars = vars_of_global env constr in
if Id.Set.mem id vars then raise Occur
-let occur_var env id c =
+let occur_var env sigma id c =
let rec occur_rec c =
- match kind_of_term c with
- | Var _ | Const _ | Ind _ | Construct _ -> occur_in_global env id c
- | _ -> iter_constr occur_rec c
+ match EConstr.kind sigma c with
+ | Var _ | Const _ | Ind _ | Construct _ -> occur_in_global env id (EConstr.to_constr sigma c)
+ | _ -> EConstr.iter sigma occur_rec c
in
try occur_rec c; false with Occur -> true
-let occur_var_in_decl env hyp decl =
+let occur_var_in_decl env sigma hyp decl =
let open NamedDecl in
match decl with
- | LocalAssum (_,typ) -> occur_var env hyp typ
+ | LocalAssum (_,typ) -> occur_var env sigma hyp typ
| LocalDef (_, body, typ) ->
- occur_var env hyp typ ||
- occur_var env hyp body
+ occur_var env sigma hyp typ ||
+ occur_var env sigma hyp body
-let local_occur_var id c =
- let rec occur c = match kind_of_term c with
+let local_occur_var sigma id c =
+ let rec occur c = match EConstr.kind sigma c with
| Var id' -> if Id.equal id id' then raise Occur
- | _ -> Constr.iter occur c
+ | _ -> EConstr.iter sigma occur c
in
try occur c; false with Occur -> true
(* returns the list of free debruijn indices in a term *)
-let free_rels m =
- let rec frec depth acc c = match kind_of_term c with
+let free_rels sigma m =
+ let rec frec depth acc c = match EConstr.kind sigma c with
| Rel n -> if n >= depth then Int.Set.add (n-depth+1) acc else acc
- | _ -> fold_constr_with_binders succ frec depth acc c
+ | _ -> fold_constr_with_binders sigma succ frec depth acc c
in
frec 1 Int.Set.empty m
(* collects all metavar occurrences, in left-to-right order, preserving
* repetitions and all. *)
-let collect_metas c =
+let collect_metas sigma c =
let rec collrec acc c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Meta mv -> List.add_set Int.equal mv acc
- | _ -> fold_constr collrec acc c
+ | _ -> EConstr.fold sigma collrec acc c
in
List.rev (collrec [] c)
(* collects all vars; warning: this is only visible vars, not dependencies in
all section variables; for the latter, use global_vars_set *)
-let collect_vars c =
- let rec aux vars c = match kind_of_term c with
+let collect_vars sigma c =
+ let rec aux vars c = match EConstr.kind sigma c with
| Var id -> Id.Set.add id vars
- | _ -> fold_constr aux vars c in
+ | _ -> EConstr.fold sigma aux vars c in
aux Id.Set.empty c
let vars_of_global_reference env gr =
@@ -606,51 +912,56 @@ let vars_of_global_reference env gr =
(* Tests whether [m] is a subterm of [t]:
[m] is appropriately lifted through abstractions of [t] *)
-let dependent_main noevar univs m t =
- let eqc x y = if univs then fst (Universes.eq_constr_universes x y) else eq_constr_nounivs x y in
+let dependent_main noevar univs sigma m t =
+ let open EConstr in
+ let eqc x y =
+ if univs then not (Option.is_empty (eq_constr_universes sigma x y))
+ else eq_constr_nounivs sigma x y
+ in
let rec deprec m t =
if eqc m t then
raise Occur
else
- match kind_of_term m, kind_of_term t with
+ match EConstr.kind sigma m, EConstr.kind sigma t with
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
deprec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
CArray.Fun1.iter deprec m
(Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
- | _, Cast (c,_,_) when noevar && isMeta c -> ()
+ | _, Cast (c,_,_) when noevar && isMeta sigma c -> ()
| _, Evar _ when noevar -> ()
- | _ -> iter_constr_with_binders (fun c -> lift 1 c) deprec m t
+ | _ -> EConstr.iter_with_binders sigma (fun c -> Vars.lift 1 c) deprec m t
in
try deprec m t; false with Occur -> true
-let dependent = dependent_main false false
-let dependent_no_evar = dependent_main true false
+let dependent sigma c t = dependent_main false false sigma c t
+let dependent_no_evar sigma c t = dependent_main true false sigma c t
-let dependent_univs = dependent_main false true
-let dependent_univs_no_evar = dependent_main true true
+let dependent_univs sigma c t = dependent_main false true sigma c t
+let dependent_univs_no_evar sigma c t = dependent_main true true sigma c t
-let dependent_in_decl a decl =
+let dependent_in_decl sigma a decl =
let open NamedDecl in
match decl with
- | LocalAssum (_,t) -> dependent a t
- | LocalDef (_, body, t) -> dependent a body || dependent a t
+ | LocalAssum (_,t) -> dependent sigma a t
+ | LocalDef (_, body, t) -> dependent sigma a body || dependent sigma a t
-let count_occurrences m t =
+let count_occurrences sigma m t =
+ let open EConstr in
let n = ref 0 in
let rec countrec m t =
- if eq_constr m t then
+ if EConstr.eq_constr sigma m t then
incr n
else
- match kind_of_term m, kind_of_term t with
+ match EConstr.kind sigma m, EConstr.kind sigma t with
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
countrec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
Array.iter (countrec m)
(Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
- | _, Cast (c,_,_) when isMeta c -> ()
+ | _, Cast (c,_,_) when isMeta sigma c -> ()
| _, Evar _ -> ()
- | _ -> iter_constr_with_binders (lift 1) countrec m t
+ | _ -> EConstr.iter_with_binders sigma (Vars.lift 1) countrec m t
in
countrec m t;
!n
@@ -658,7 +969,7 @@ let count_occurrences m t =
(* Synonymous *)
let occur_term = dependent
-let pop t = lift (-1) t
+let pop t = EConstr.Vars.lift (-1) t
(***************************)
(* bindings functions *)
@@ -668,35 +979,55 @@ type meta_type_map = (metavariable * types) list
type meta_value_map = (metavariable * constr) list
+let isMetaOf sigma mv c =
+ match EConstr.kind sigma c with Meta mv' -> Int.equal mv mv' | _ -> false
+
let rec subst_meta bl c =
match kind_of_term c with
| Meta i -> (try Int.List.assoc i bl with Not_found -> c)
| _ -> map_constr (subst_meta bl) c
+let rec strip_outer_cast sigma c = match EConstr.kind sigma c with
+ | Cast (c,_,_) -> strip_outer_cast sigma c
+ | _ -> c
+
+(* flattens application lists throwing casts in-between *)
+let collapse_appl sigma c = match EConstr.kind sigma c with
+ | App (f,cl) ->
+ let rec collapse_rec f cl2 =
+ match EConstr.kind sigma (strip_outer_cast sigma f) with
+ | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
+ | _ -> EConstr.mkApp (f,cl2)
+ in
+ collapse_rec f cl
+ | _ -> c
+
(* First utilities for avoiding telescope computation for subst_term *)
-let prefix_application eq_fun (k,c) (t : constr) =
- let c' = collapse_appl c and t' = collapse_appl t in
- match kind_of_term c', kind_of_term t' with
+let prefix_application sigma eq_fun (k,c) t =
+ let open EConstr in
+ let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
+ match EConstr.kind sigma c', EConstr.kind sigma t' with
| App (f1,cl1), App (f2,cl2) ->
let l1 = Array.length cl1
and l2 = Array.length cl2 in
if l1 <= l2
- && eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
+ && eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
else
None
| _ -> None
-let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
- let c' = collapse_appl c and t' = collapse_appl t in
- match kind_of_term c', kind_of_term t' with
+let my_prefix_application sigma eq_fun (k,c) by_c t =
+ let open EConstr in
+ let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
+ match EConstr.kind sigma c', EConstr.kind sigma t' with
| App (f1,cl1), App (f2,cl2) ->
let l1 = Array.length cl1
and l2 = Array.length cl2 in
if l1 <= l2
- && eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
- Some (mkApp ((lift k by_c), Array.sub cl2 l1 (l2 - l1)))
+ && eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
+ Some (mkApp ((Vars.lift k by_c), Array.sub cl2 l1 (l2 - l1)))
else
None
| _ -> None
@@ -705,35 +1036,37 @@ let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
substitutes [(Rel 1)] for all occurrences of term [c] in a term [t];
works if [c] has rels *)
-let subst_term_gen eq_fun c t =
+let subst_term_gen sigma eq_fun c t =
+ let open EConstr in
+ let open Vars in
let rec substrec (k,c as kc) t =
- match prefix_application eq_fun kc t with
+ match prefix_application sigma eq_fun kc t with
| Some x -> x
| None ->
- if eq_fun c t then mkRel k
+ if eq_fun sigma c t then mkRel k
else
- map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
+ EConstr.map_with_binders sigma (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
in
substrec (1,c) t
-let subst_term = subst_term_gen eq_constr
+let subst_term sigma c t = subst_term_gen sigma EConstr.eq_constr c t
(* Recognizing occurrences of a given subterm in a term :
[replace_term c1 c2 t] substitutes [c2] for all occurrences of
term [c1] in a term [t]; works if [c1] and [c2] have rels *)
-let replace_term_gen eq_fun c by_c in_t =
+let replace_term_gen sigma eq_fun c by_c in_t =
let rec substrec (k,c as kc) t =
- match my_prefix_application eq_fun kc by_c t with
+ match my_prefix_application sigma eq_fun kc by_c t with
| Some x -> x
| None ->
- (if eq_fun c t then (lift k by_c) else
- map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c))
+ (if eq_fun sigma c t then (EConstr.Vars.lift k by_c) else
+ EConstr.map_with_binders sigma (fun (k,c) -> (k+1,EConstr.Vars.lift 1 c))
substrec kc t)
in
substrec (0,c) in_t
-let replace_term = replace_term_gen eq_constr
+let replace_term sigma c byc t = replace_term_gen sigma EConstr.eq_constr c byc t
let vars_of_env env =
let s =
@@ -784,15 +1117,37 @@ let is_section_variable id =
try let _ = Global.lookup_named id in true
with Not_found -> false
-let isGlobalRef c =
- match kind_of_term c with
+let global_of_constr sigma c =
+ let open Globnames in
+ match EConstr.kind sigma c with
+ | Const (c, u) -> ConstRef c, u
+ | Ind (i, u) -> IndRef i, u
+ | Construct (c, u) -> ConstructRef c, u
+ | Var id -> VarRef id, EConstr.EInstance.empty
+ | _ -> raise Not_found
+
+let is_global sigma c t =
+ let open Globnames in
+ match c, EConstr.kind sigma t with
+ | ConstRef c, Const (c', _) -> Constant.equal c c'
+ | IndRef i, Ind (i', _) -> eq_ind i i'
+ | ConstructRef i, Construct (i', _) -> eq_constructor i i'
+ | VarRef id, Var id' -> Id.equal id id'
+ | _ -> false
+
+let isGlobalRef sigma c =
+ match EConstr.kind sigma c with
| Const _ | Ind _ | Construct _ | Var _ -> true
| _ -> false
-let is_template_polymorphic env f =
- match kind_of_term f with
- | Ind ind -> Environ.template_polymorphic_pind ind env
- | Const c -> Environ.template_polymorphic_pconstant c env
+let is_template_polymorphic env sigma f =
+ match EConstr.kind sigma f with
+ | Ind (ind, u) ->
+ if not (EConstr.EInstance.is_empty u) then false
+ else Environ.template_polymorphic_ind ind env
+ | Const (cst, u) ->
+ if not (EConstr.EInstance.is_empty u) then false
+ else Environ.template_polymorphic_constant cst env
| _ -> false
let base_sort_cmp pb s0 s1 =
@@ -802,21 +1157,36 @@ let base_sort_cmp pb s0 s1 =
| (Type u1, Type u2) -> true
| _ -> false
+let rec is_Prop sigma c = match EConstr.kind sigma c with
+ | Sort u ->
+ begin match EConstr.ESorts.kind sigma u with
+ | Prop Null -> true
+ | _ -> false
+ end
+ | Cast (c,_,_) -> is_Prop sigma c
+ | _ -> false
+
(* eq_constr extended with universe erasure *)
-let compare_constr_univ f cv_pb t1 t2 =
- match kind_of_term t1, kind_of_term t2 with
- Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 s2
+let compare_constr_univ sigma f cv_pb t1 t2 =
+ let open EConstr in
+ match EConstr.kind sigma t1, EConstr.kind sigma t2 with
+ Sort s1, Sort s2 -> base_sort_cmp cv_pb (ESorts.kind sigma s1) (ESorts.kind sigma s2)
| Prod (_,t1,c1), Prod (_,t2,c2) ->
f Reduction.CONV t1 t2 && f cv_pb c1 c2
- | _ -> compare_constr (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2
+ | Const (c, u), Const (c', u') -> Constant.equal c c'
+ | Ind (i, _), Ind (i', _) -> eq_ind i i'
+ | Construct (i, _), Construct (i', _) -> eq_constructor i i'
+ | _ -> EConstr.compare_constr sigma (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2
-let rec constr_cmp cv_pb t1 t2 = compare_constr_univ constr_cmp cv_pb t1 t2
+let constr_cmp sigma cv_pb t1 t2 =
+ let rec compare cv_pb t1 t2 = compare_constr_univ sigma compare cv_pb t1 t2 in
+ compare cv_pb t1 t2
-let eq_constr t1 t2 = constr_cmp Reduction.CONV t1 t2
+let eq_constr sigma t1 t2 = constr_cmp sigma Reduction.CONV t1 t2
(* App(c,[t1,...tn]) -> ([c,t1,...,tn-1],tn)
App(c,[||]) -> ([],c) *)
-let split_app c = match kind_of_term c with
+let split_app sigma c = match EConstr.kind sigma c with
App(c,l) ->
let len = Array.length l in
if Int.equal len 0 then ([],c) else
@@ -825,28 +1195,30 @@ let split_app c = match kind_of_term c with
c::(Array.to_list prev), last
| _ -> assert false
-type subst = (Context.Rel.t * constr) Evar.Map.t
+type subst = (EConstr.rel_context * EConstr.constr) Evar.Map.t
exception CannotFilter
-let filtering env cv_pb c1 c2 =
+let filtering sigma env cv_pb c1 c2 =
+ let open EConstr in
+ let open Vars in
let evm = ref Evar.Map.empty in
let define cv_pb e1 ev c1 =
try let (e2,c2) = Evar.Map.find ev !evm in
let shift = List.length e1 - List.length e2 in
- if constr_cmp cv_pb c1 (lift shift c2) then () else raise CannotFilter
+ if constr_cmp sigma cv_pb c1 (lift shift c2) then () else raise CannotFilter
with Not_found ->
evm := Evar.Map.add ev (e1,c1) !evm
in
let rec aux env cv_pb c1 c2 =
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind sigma c1, EConstr.kind sigma c2 with
| App _, App _ ->
- let ((p1,l1),(p2,l2)) = (split_app c1),(split_app c2) in
+ let ((p1,l1),(p2,l2)) = (split_app sigma c1),(split_app sigma c2) in
let () = aux env cv_pb l1 l2 in
begin match p1, p2 with
| [], [] -> ()
| (h1 :: p1), (h2 :: p2) ->
- aux env cv_pb (applistc h1 p1) (applistc h2 p2)
+ aux env cv_pb (applist (h1, p1)) (applist (h2, p2))
| _ -> assert false
end
| Prod (n,t1,c1), Prod (_,t2,c2) ->
@@ -855,52 +1227,52 @@ let filtering env cv_pb c1 c2 =
| _, Evar (ev,_) -> define cv_pb env ev c1
| Evar (ev,_), _ -> define cv_pb env ev c2
| _ ->
- if compare_constr_univ
+ if compare_constr_univ sigma
(fun pb c1 c2 -> aux env pb c1 c2; true) cv_pb c1 c2 then ()
else raise CannotFilter
(* TODO: le reste des binders *)
in
aux env cv_pb c1 c2; !evm
-let decompose_prod_letin : constr -> int * Context.Rel.t * constr =
- let rec prodec_rec i l c = match kind_of_term c with
+let decompose_prod_letin sigma c =
+ let rec prodec_rec i l c = match EConstr.kind sigma c with
| Prod (n,t,c) -> prodec_rec (succ i) (RelDecl.LocalAssum (n,t)::l) c
| LetIn (n,d,t,c) -> prodec_rec (succ i) (RelDecl.LocalDef (n,d,t)::l) c
| Cast (c,_,_) -> prodec_rec i l c
| _ -> i,l,c in
- prodec_rec 0 []
+ prodec_rec 0 [] c
(* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction
* gives n (casts are ignored) *)
-let nb_lam =
- let rec nbrec n c = match kind_of_term c with
+let nb_lam sigma c =
+ let rec nbrec n c = match EConstr.kind sigma c with
| Lambda (_,_,c) -> nbrec (n+1) c
| Cast (c,_,_) -> nbrec n c
| _ -> n
in
- nbrec 0
+ nbrec 0 c
(* similar to nb_lam, but gives the number of products instead *)
-let nb_prod =
- let rec nbrec n c = match kind_of_term c with
+let nb_prod sigma c =
+ let rec nbrec n c = match EConstr.kind sigma c with
| Prod (_,_,c) -> nbrec (n+1) c
| Cast (c,_,_) -> nbrec n c
| _ -> n
in
- nbrec 0
+ nbrec 0 c
-let nb_prod_modulo_zeta x =
+let nb_prod_modulo_zeta sigma x =
let rec count n c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
Prod(_,_,t) -> count (n+1) t
- | LetIn(_,a,_,t) -> count n (subst1 a t)
+ | LetIn(_,a,_,t) -> count n (EConstr.Vars.subst1 a t)
| Cast(c,_,_) -> count n c
| _ -> n
in count 0 x
-let align_prod_letin c a : Context.Rel.t * constr =
- let (lc,_,_) = decompose_prod_letin c in
- let (la,l,a) = decompose_prod_letin a in
+let align_prod_letin sigma c a =
+ let (lc,_,_) = decompose_prod_letin sigma c in
+ let (la,l,a) = decompose_prod_letin sigma a in
if not (la >= lc) then invalid_arg "align_prod_letin";
let (l1,l2) = Util.List.chop lc l in
l2,it_mkProd_or_LetIn a l1
@@ -909,21 +1281,23 @@ let align_prod_letin c a : Context.Rel.t * constr =
(* [x1:c1;...;xn:cn]@(f;a1...an;x1;...;xn) --> @(f;a1...an) *)
(* Remplace 2 earlier buggish versions *)
-let rec eta_reduce_head c =
- match kind_of_term c with
+let rec eta_reduce_head sigma c =
+ let open EConstr in
+ let open Vars in
+ match EConstr.kind sigma c with
| Lambda (_,c1,c') ->
- (match kind_of_term (eta_reduce_head c') with
+ (match EConstr.kind sigma (eta_reduce_head sigma c') with
| App (f,cl) ->
let lastn = (Array.length cl) - 1 in
- if lastn < 0 then anomaly (Pp.str "application without arguments")
+ if lastn < 0 then anomaly (Pp.str "application without arguments.")
else
- (match kind_of_term cl.(lastn) with
+ (match EConstr.kind sigma cl.(lastn) with
| Rel 1 ->
let c' =
if Int.equal lastn 0 then f
else mkApp (f, Array.sub cl 0 lastn)
in
- if noccurn 1 c'
+ if noccurn sigma 1 c'
then lift (-1) c'
else c
| _ -> c)
@@ -934,7 +1308,7 @@ let rec eta_reduce_head c =
(* iterator on rel context *)
let process_rel_context f env =
let sign = named_context_val env in
- let rels = rel_context env in
+ let rels = EConstr.rel_context env in
let env0 = reset_with_named_context sign env in
Context.Rel.fold_outside f rels ~init:env0
@@ -982,18 +1356,35 @@ let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init
let mem_named_context_val id ctxt =
try ignore(Environ.lookup_named_val id ctxt); true with Not_found -> false
-let compact_named_context_reverse sign =
- let compact l decl =
- let (i1,c1,t1) = NamedDecl.to_tuple decl in
- match l with
- | [] -> [[i1],c1,t1]
- | (l2,c2,t2)::q ->
- if Option.equal Constr.equal c1 c2 && Constr.equal t1 t2
- then (i1::l2,c2,t2)::q
- else ([i1],c1,t1)::l
- in Context.Named.fold_inside compact ~init:[] sign
+let map_rel_decl f = function
+| RelDecl.LocalAssum (id, t) -> RelDecl.LocalAssum (id, f t)
+| RelDecl.LocalDef (id, b, t) -> RelDecl.LocalDef (id, f b, f t)
+
+let map_named_decl f = function
+| NamedDecl.LocalAssum (id, t) -> NamedDecl.LocalAssum (id, f t)
+| NamedDecl.LocalDef (id, b, t) -> NamedDecl.LocalDef (id, f b, f t)
-let compact_named_context sign = List.rev (compact_named_context_reverse sign)
+let compact_named_context sign =
+ let compact l decl =
+ match decl, l with
+ | NamedDecl.LocalAssum (i,t), [] ->
+ [CompactedDecl.LocalAssum ([i],t)]
+ | NamedDecl.LocalDef (i,c,t), [] ->
+ [CompactedDecl.LocalDef ([i],c,t)]
+ | NamedDecl.LocalAssum (i1,t1), CompactedDecl.LocalAssum (li,t2) :: q ->
+ if Constr.equal t1 t2
+ then CompactedDecl.LocalAssum (i1::li, t2) :: q
+ else CompactedDecl.LocalAssum ([i1],t1) :: CompactedDecl.LocalAssum (li,t2) :: q
+ | NamedDecl.LocalDef (i1,c1,t1), CompactedDecl.LocalDef (li,c2,t2) :: q ->
+ if Constr.equal c1 c2 && Constr.equal t1 t2
+ then CompactedDecl.LocalDef (i1::li, c2, t2) :: q
+ else CompactedDecl.LocalDef ([i1],c1,t1) :: CompactedDecl.LocalDef (li,c2,t2) :: q
+ | NamedDecl.LocalAssum (i,t), q ->
+ CompactedDecl.LocalAssum ([i],t) :: q
+ | NamedDecl.LocalDef (i,c,t), q ->
+ CompactedDecl.LocalDef ([i],c,t) :: q
+ in
+ sign |> Context.Named.fold_inside compact ~init:[] |> List.rev
let clear_named_body id env =
let open NamedDecl in
@@ -1002,28 +1393,58 @@ let clear_named_body id env =
| d -> push_named d in
fold_named_context aux env ~init:(reset_context env)
-let global_vars env ids = Id.Set.elements (global_vars_set env ids)
+let global_vars_set env sigma constr =
+ let rec filtrec acc c =
+ let acc = match EConstr.kind sigma c with
+ | Var _ | Const _ | Ind _ | Construct _ ->
+ Id.Set.union (vars_of_global env (EConstr.to_constr sigma c)) acc
+ | _ -> acc
+ in
+ EConstr.fold sigma filtrec acc c
+ in
+ filtrec Id.Set.empty constr
+
+let global_vars env sigma ids = Id.Set.elements (global_vars_set env sigma ids)
-let global_vars_set_of_decl env = function
- | NamedDecl.LocalAssum (_,t) -> global_vars_set env t
+let global_vars_set_of_decl env sigma = function
+ | NamedDecl.LocalAssum (_,t) -> global_vars_set env sigma t
| NamedDecl.LocalDef (_,c,t) ->
- Id.Set.union (global_vars_set env t)
- (global_vars_set env c)
+ Id.Set.union (global_vars_set env sigma t)
+ (global_vars_set env sigma c)
-let dependency_closure env sign hyps =
+let dependency_closure env sigma sign hyps =
if Id.Set.is_empty hyps then [] else
let (_,lh) =
Context.Named.fold_inside
(fun (hs,hl) d ->
let x = NamedDecl.get_id d in
if Id.Set.mem x hs then
- (Id.Set.union (global_vars_set_of_decl env d) (Id.Set.remove x hs),
+ (Id.Set.union (global_vars_set_of_decl env sigma d) (Id.Set.remove x hs),
x::hl)
else (hs,hl))
~init:(hyps,[])
sign in
List.rev lh
+let global_app_of_constr sigma c =
+ let open Globnames in
+ match EConstr.kind sigma c with
+ | Const (c, u) -> (ConstRef c, u), None
+ | Ind (i, u) -> (IndRef i, u), None
+ | Construct (c, u) -> (ConstructRef c, u), None
+ | Var id -> (VarRef id, EConstr.EInstance.empty), None
+ | Proj (p, c) -> (ConstRef (Projection.constant p), EConstr.EInstance.empty), Some c
+ | _ -> raise Not_found
+
+let prod_applist sigma c l =
+ let open EConstr in
+ let rec app subst c l =
+ match EConstr.kind sigma c, l with
+ | Prod(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> Vars.substl subst c
+ | _ -> anomaly (Pp.str "Not enough prod's.") in
+ app [] c l
+
(* Combinators on judgments *)
let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type }
@@ -1037,33 +1458,13 @@ let context_chop k ctx =
| (0, l2) -> (List.rev acc, l2)
| (n, (RelDecl.LocalDef _ as h)::t) -> chop_aux (h::acc) (n, t)
| (n, (h::t)) -> chop_aux (h::acc) (pred n, t)
- | (_, []) -> anomaly (Pp.str "context_chop")
+ | (_, []) -> anomaly (Pp.str "context_chop.")
in chop_aux [] (k,ctx)
(* Do not skip let-in's *)
let env_rel_context_chop k env =
+ let open EConstr in
let rels = rel_context env in
let ctx1,ctx2 = List.chop k rels in
push_rel_context ctx2 (reset_with_named_context (named_context_val env) env),
ctx1
-
-(*******************************************)
-(* Functions to deal with impossible cases *)
-(*******************************************)
-let impossible_default_case = ref None
-
-let set_impossible_default_clause c = impossible_default_case := Some c
-
-let coq_unit_judge =
- let na1 = Name (Id.of_string "A") in
- let na2 = Name (Id.of_string "H") in
- fun () ->
- match !impossible_default_case with
- | Some fn ->
- let (id,type_of_id), ctx = fn () in
- make_judge id type_of_id, ctx
- | None ->
- (* In case the constants id/ID are not defined *)
- make_judge (mkLambda (na1,mkProp,mkLambda(na2,mkRel 1,mkRel 1)))
- (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2))),
- Univ.ContextSet.empty
diff --git a/engine/termops.mli b/engine/termops.mli
index 9ab8e70ec..c19a2d15a 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -13,63 +13,54 @@ open Pp
open Names
open Term
open Environ
+open EConstr
(** printers *)
val print_sort : sorts -> std_ppcmds
val pr_sort_family : sorts_family -> std_ppcmds
-val pr_fix : (constr -> std_ppcmds) -> fixpoint -> std_ppcmds
-
-(** debug printer: do not use to display terms to the casual user... *)
-val set_print_constr : (env -> constr -> std_ppcmds) -> unit
-val print_constr : constr -> std_ppcmds
-val print_constr_env : env -> constr -> std_ppcmds
-val print_named_context : env -> std_ppcmds
-val pr_rel_decl : env -> Context.Rel.Declaration.t -> std_ppcmds
-val print_rel_context : env -> std_ppcmds
-val print_env : env -> std_ppcmds
+val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> std_ppcmds
(** about contexts *)
val push_rel_assum : Name.t * types -> env -> env
-val push_rels_assum : (Name.t * types) list -> env -> env
-val push_named_rec_types : Name.t array * types array * 'a -> env -> env
+val push_rels_assum : (Name.t * Constr.types) list -> env -> env
+val push_named_rec_types : Name.t array * Constr.types array * 'a -> env -> env
-val lookup_rel_id : Id.t -> Context.Rel.t -> int * constr option * types
+val lookup_rel_id : Id.t -> ('c, 't) Context.Rel.pt -> int * 'c option * 't
(** Associates the contents of an identifier in a [rel_context]. Raise
[Not_found] if there is no such identifier. *)
(** Functions that build argument lists matching a block of binders or a context.
[rel_vect n m] builds [|Rel (n+m);...;Rel(n+1)|]
*)
-val rel_vect : int -> int -> constr array
+val rel_vect : int -> int -> Constr.constr array
val rel_list : int -> int -> constr list
(** iterators/destructors on terms *)
-val mkProd_or_LetIn : Context.Rel.Declaration.t -> types -> types
-val mkProd_wo_LetIn : Context.Rel.Declaration.t -> types -> types
+val mkProd_or_LetIn : rel_declaration -> types -> types
+val mkProd_wo_LetIn : rel_declaration -> types -> types
val it_mkProd : types -> (Name.t * types) list -> types
val it_mkLambda : constr -> (Name.t * types) list -> constr
-val it_mkProd_or_LetIn : types -> Context.Rel.t -> types
-val it_mkProd_wo_LetIn : types -> Context.Rel.t -> types
-val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr
-val it_mkNamedProd_or_LetIn : types -> Context.Named.t -> types
-val it_mkNamedProd_wo_LetIn : types -> Context.Named.t -> types
-val it_mkNamedLambda_or_LetIn : constr -> Context.Named.t -> constr
+val it_mkProd_or_LetIn : types -> rel_context -> types
+val it_mkProd_wo_LetIn : types -> rel_context -> types
+val it_mkLambda_or_LetIn : Constr.constr -> Context.Rel.t -> Constr.constr
+val it_mkNamedProd_or_LetIn : types -> named_context -> types
+val it_mkNamedProd_wo_LetIn : Constr.types -> Context.Named.t -> Constr.types
+val it_mkNamedLambda_or_LetIn : constr -> named_context -> constr
(* Ad hoc version reinserting letin, assuming the body is defined in
the context where the letins are expanded *)
-val it_mkLambda_or_LetIn_from_no_LetIn : constr -> Context.Rel.t -> constr
+val it_mkLambda_or_LetIn_from_no_LetIn : Constr.constr -> Context.Rel.t -> Constr.constr
(** {6 Generic iterators on constr} *)
-val map_constr_with_named_binders :
- (Name.t -> 'a -> 'a) ->
- ('a -> constr -> constr) -> 'a -> constr -> constr
val map_constr_with_binders_left_to_right :
- (Context.Rel.Declaration.t -> 'a -> 'a) ->
+ Evd.evar_map ->
+ (rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) ->
'a -> constr -> constr
val map_constr_with_full_binders :
- (Context.Rel.Declaration.t -> 'a -> 'a) ->
+ Evd.evar_map ->
+ (rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) -> 'a -> constr -> constr
(** [fold_constr_with_binders g f n acc c] folds [f n] on the immediate
@@ -79,57 +70,58 @@ val map_constr_with_full_binders :
index) which is processed by [g] (which typically add 1 to [n]) at
each binder traversal; it is not recursive *)
-val fold_constr_with_binders :
+val fold_constr_with_binders : Evd.evar_map ->
('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
-val fold_constr_with_full_binders :
+val fold_constr_with_full_binders : Evd.evar_map ->
(Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) ->
'a -> 'b -> constr -> 'b
val iter_constr_with_full_binders :
- (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> constr -> unit) -> 'a ->
- constr -> unit
+ (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> Constr.constr -> unit) -> 'a ->
+ Constr.constr -> unit
(**********************************************************************)
-val strip_head_cast : constr -> constr
-val drop_extra_implicit_args : constr -> constr
+val strip_head_cast : Evd.evar_map -> constr -> constr
+val drop_extra_implicit_args : Evd.evar_map -> constr -> constr
(** occur checks *)
exception Occur
-val occur_meta : types -> bool
-val occur_existential : types -> bool
-val occur_meta_or_existential : types -> bool
-val occur_evar : existential_key -> types -> bool
-val occur_var : env -> Id.t -> types -> bool
+val occur_meta : Evd.evar_map -> constr -> bool
+val occur_existential : Evd.evar_map -> constr -> bool
+val occur_meta_or_existential : Evd.evar_map -> constr -> bool
+val occur_evar : Evd.evar_map -> existential_key -> constr -> bool
+val occur_var : env -> Evd.evar_map -> Id.t -> constr -> bool
val occur_var_in_decl :
- env ->
- Id.t -> Context.Named.Declaration.t -> bool
+ env -> Evd.evar_map ->
+ Id.t -> named_declaration -> bool
(** As {!occur_var} but assume the identifier not to be a section variable *)
-val local_occur_var : Id.t -> types -> bool
+val local_occur_var : Evd.evar_map -> Id.t -> constr -> bool
-val free_rels : constr -> Int.Set.t
+val free_rels : Evd.evar_map -> constr -> Int.Set.t
(** [dependent m t] tests whether [m] is a subterm of [t] *)
-val dependent : constr -> constr -> bool
-val dependent_no_evar : constr -> constr -> bool
-val dependent_univs : constr -> constr -> bool
-val dependent_univs_no_evar : constr -> constr -> bool
-val dependent_in_decl : constr -> Context.Named.Declaration.t -> bool
-val count_occurrences : constr -> constr -> int
-val collect_metas : constr -> int list
-val collect_vars : constr -> Id.Set.t (** for visible vars only *)
+val dependent : Evd.evar_map -> constr -> constr -> bool
+val dependent_no_evar : Evd.evar_map -> constr -> constr -> bool
+val dependent_univs : Evd.evar_map -> constr -> constr -> bool
+val dependent_univs_no_evar : Evd.evar_map -> constr -> constr -> bool
+val dependent_in_decl : Evd.evar_map -> constr -> named_declaration -> bool
+val count_occurrences : Evd.evar_map -> constr -> constr -> int
+val collect_metas : Evd.evar_map -> constr -> int list
+val collect_vars : Evd.evar_map -> constr -> Id.Set.t (** for visible vars only *)
val vars_of_global_reference : env -> Globnames.global_reference -> Id.Set.t
-val occur_term : constr -> constr -> bool (** Synonymous
- of dependent
- Substitution of metavariables *)
-type meta_value_map = (metavariable * constr) list
-val subst_meta : meta_value_map -> constr -> constr
+val occur_term : Evd.evar_map -> constr -> constr -> bool (** Synonymous of dependent *)
+
+(* Substitution of metavariables *)
+type meta_value_map = (metavariable * Constr.constr) list
+val subst_meta : meta_value_map -> Constr.constr -> Constr.constr
+val isMetaOf : Evd.evar_map -> metavariable -> constr -> bool
(** Type assignment for metavariables *)
-type meta_type_map = (metavariable * types) list
+type meta_type_map = (metavariable * Constr.types) list
(** [pop c] lifts by -1 the positive indexes in [c] *)
val pop : constr -> constr
@@ -140,29 +132,38 @@ val pop : constr -> constr
(** [subst_term_gen eq d c] replaces [d] by [Rel 1] in [c] using [eq]
as equality *)
-val subst_term_gen :
- (constr -> constr -> bool) -> constr -> constr -> constr
+val subst_term_gen : Evd.evar_map ->
+ (Evd.evar_map -> constr -> constr -> bool) -> constr -> constr -> constr
(** [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq]
as equality *)
val replace_term_gen :
- (constr -> constr -> bool) ->
+ Evd.evar_map -> (Evd.evar_map -> constr -> constr -> bool) ->
constr -> constr -> constr -> constr
(** [subst_term d c] replaces [d] by [Rel 1] in [c] *)
-val subst_term : constr -> constr -> constr
+val subst_term : Evd.evar_map -> constr -> constr -> constr
(** [replace_term d e c] replaces [d] by [e] in [c] *)
-val replace_term : constr -> constr -> constr -> constr
+val replace_term : Evd.evar_map -> constr -> constr -> constr -> constr
(** Alternative term equalities *)
val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool
-val compare_constr_univ : (Reduction.conv_pb -> constr -> constr -> bool) ->
+val compare_constr_univ : Evd.evar_map -> (Reduction.conv_pb -> constr -> constr -> bool) ->
Reduction.conv_pb -> constr -> constr -> bool
-val constr_cmp : Reduction.conv_pb -> constr -> constr -> bool
-val eq_constr : constr -> constr -> bool (* FIXME rename: erases universes*)
+val constr_cmp : Evd.evar_map -> Reduction.conv_pb -> constr -> constr -> bool
+val eq_constr : Evd.evar_map -> constr -> constr -> bool (* FIXME rename: erases universes*)
+
+val eta_reduce_head : Evd.evar_map -> constr -> constr
+
+(** Flattens application lists *)
+val collapse_appl : Evd.evar_map -> constr -> constr
-val eta_reduce_head : constr -> constr
+val prod_applist : Evd.evar_map -> constr -> constr list -> constr
+
+(** Remove recursively the casts around a term i.e.
+ [strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
+val strip_outer_cast : Evd.evar_map -> constr -> constr
exception CannotFilter
@@ -172,27 +173,27 @@ exception CannotFilter
(context,term), or raises [CannotFilter].
Warning: Outer-kernel sort subtyping are taken into account: c1 has
to be smaller than c2 wrt. sorts. *)
-type subst = (Context.Rel.t * constr) Evar.Map.t
-val filtering : Context.Rel.t -> Reduction.conv_pb -> constr -> constr -> subst
+type subst = (rel_context * constr) Evar.Map.t
+val filtering : Evd.evar_map -> rel_context -> Reduction.conv_pb -> constr -> constr -> subst
-val decompose_prod_letin : constr -> int * Context.Rel.t * constr
-val align_prod_letin : constr -> constr -> Context.Rel.t * constr
+val decompose_prod_letin : Evd.evar_map -> constr -> int * rel_context * constr
+val align_prod_letin : Evd.evar_map -> constr -> constr -> rel_context * constr
(** [nb_lam] {% $ %}[x_1:T_1]...[x_n:T_n]c{% $ %} where {% $ %}c{% $ %} is not an abstraction
gives {% $ %}n{% $ %} (casts are ignored) *)
-val nb_lam : constr -> int
+val nb_lam : Evd.evar_map -> constr -> int
(** Similar to [nb_lam], but gives the number of products instead *)
-val nb_prod : constr -> int
+val nb_prod : Evd.evar_map -> constr -> int
(** Similar to [nb_prod], but zeta-contracts let-in on the way *)
-val nb_prod_modulo_zeta : constr -> int
+val nb_prod_modulo_zeta : Evd.evar_map -> constr -> int
(** Get the last arg of a constr intended to be an application *)
-val last_arg : constr -> constr
+val last_arg : Evd.evar_map -> constr -> constr
(** Force the decomposition of a term as an applicative one *)
-val decompose_app_vect : constr -> constr * constr array
+val decompose_app_vect : Evd.evar_map -> constr -> constr * constr array
val adjust_app_list_size : constr -> constr list -> constr -> constr list ->
(constr * constr list * constr * constr list)
@@ -205,8 +206,8 @@ val add_name : Name.t -> names_context -> names_context
val lookup_name_of_rel : int -> names_context -> Name.t
val lookup_rel_of_name : Id.t -> names_context -> int
val empty_names_context : names_context
-val ids_of_rel_context : Context.Rel.t -> Id.t list
-val ids_of_named_context : Context.Named.t -> Id.t list
+val ids_of_rel_context : ('c, 't) Context.Rel.pt -> Id.t list
+val ids_of_named_context : ('c, 't) Context.Named.pt -> Id.t list
val ids_of_context : env -> Id.t list
val names_of_rel_context : env -> names_context
@@ -218,52 +219,85 @@ val context_chop : int -> Context.Rel.t -> Context.Rel.t * Context.Rel.t
(* [env_rel_context_chop n env] extracts out the [n] top declarations
of the rel_context part of [env], counting both local definitions and
hypotheses *)
-val env_rel_context_chop : int -> env -> env * Context.Rel.t
+val env_rel_context_chop : int -> env -> env * rel_context
(** Set of local names *)
val vars_of_env: env -> Id.Set.t
val add_vname : Id.Set.t -> Name.t -> Id.Set.t
(** other signature iterators *)
-val process_rel_context : (Context.Rel.Declaration.t -> env -> env) -> env -> env
-val assums_of_rel_context : Context.Rel.t -> (Name.t * constr) list
+val process_rel_context : (rel_declaration -> env -> env) -> env -> env
+val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Name.t * 't) list
val lift_rel_context : int -> Context.Rel.t -> Context.Rel.t
-val substl_rel_context : constr list -> Context.Rel.t -> Context.Rel.t
+val substl_rel_context : Constr.constr list -> Context.Rel.t -> Context.Rel.t
val smash_rel_context : Context.Rel.t -> Context.Rel.t (** expand lets in context *)
val map_rel_context_in_env :
- (env -> constr -> constr) -> env -> Context.Rel.t -> Context.Rel.t
+ (env -> Constr.constr -> Constr.constr) -> env -> Context.Rel.t -> Context.Rel.t
val map_rel_context_with_binders :
- (int -> constr -> constr) -> Context.Rel.t -> Context.Rel.t
+ (int -> 'c -> 'c) -> ('c, 'c) Context.Rel.pt -> ('c, 'c) Context.Rel.pt
val fold_named_context_both_sides :
('a -> Context.Named.Declaration.t -> Context.Named.Declaration.t list -> 'a) ->
Context.Named.t -> init:'a -> 'a
val mem_named_context_val : Id.t -> named_context_val -> bool
-val compact_named_context : Context.Named.t -> Context.NamedList.t
-val compact_named_context_reverse : Context.Named.t -> Context.NamedList.t
+val compact_named_context : Context.Named.t -> Context.Compacted.t
+
+val map_rel_decl : ('a -> 'b) -> ('a, 'a) Context.Rel.Declaration.pt -> ('b, 'b) Context.Rel.Declaration.pt
+val map_named_decl : ('a -> 'b) -> ('a, 'a) Context.Named.Declaration.pt -> ('b, 'b) Context.Named.Declaration.pt
val clear_named_body : Id.t -> env -> env
-val global_vars : env -> constr -> Id.t list
-val global_vars_set_of_decl : env -> Context.Named.Declaration.t -> Id.Set.t
+val global_vars : env -> Evd.evar_map -> constr -> Id.t list
+val global_vars_set : env -> Evd.evar_map -> constr -> Id.Set.t
+val global_vars_set_of_decl : env -> Evd.evar_map -> named_declaration -> Id.Set.t
+val global_app_of_constr : Evd.evar_map -> constr -> (Globnames.global_reference * EInstance.t) * constr option
(** Gives an ordered list of hypotheses, closed by dependencies,
containing a given set *)
-val dependency_closure : env -> Context.Named.t -> Id.Set.t -> Id.t list
+val dependency_closure : env -> Evd.evar_map -> named_context -> Id.Set.t -> Id.t list
(** Test if an identifier is the basename of a global reference *)
val is_section_variable : Id.t -> bool
-val isGlobalRef : constr -> bool
+val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference * EInstance.t
+
+val is_global : Evd.evar_map -> Globnames.global_reference -> constr -> bool
-val is_template_polymorphic : env -> constr -> bool
+val isGlobalRef : Evd.evar_map -> constr -> bool
+
+val is_template_polymorphic : env -> Evd.evar_map -> constr -> bool
+
+val is_Prop : Evd.evar_map -> constr -> bool
(** Combinators on judgments *)
-val on_judgment : (types -> types) -> unsafe_judgment -> unsafe_judgment
-val on_judgment_value : (types -> types) -> unsafe_judgment -> unsafe_judgment
-val on_judgment_type : (types -> types) -> unsafe_judgment -> unsafe_judgment
+val on_judgment : ('a -> 'b) -> ('a, 'a) punsafe_judgment -> ('b, 'b) punsafe_judgment
+val on_judgment_value : ('c -> 'c) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment
+val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment
+
+(** {5 Debug pretty-printers} *)
+
+open Evd
+
+val pr_existential_key : evar_map -> evar -> Pp.std_ppcmds
+
+val pr_evar_suggested_name : existential_key -> evar_map -> Id.t
+
+val pr_evar_info : evar_info -> Pp.std_ppcmds
+val pr_evar_constraints : evar_map -> evar_constraint list -> Pp.std_ppcmds
+val pr_evar_map : ?with_univs:bool -> int option -> evar_map -> Pp.std_ppcmds
+val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> evar_info -> bool) ->
+ evar_map -> Pp.std_ppcmds
+val pr_metaset : Metaset.t -> Pp.std_ppcmds
+val pr_evar_universe_context : evar_universe_context -> Pp.std_ppcmds
+val pr_evd_level : evar_map -> Univ.Level.t -> Pp.std_ppcmds
-(** {6 Functions to deal with impossible cases } *)
-val set_impossible_default_clause : (unit -> (constr * types) Univ.in_universe_context_set) -> unit
-val coq_unit_judge : unit -> unsafe_judgment Univ.in_universe_context_set
+(** debug printer: do not use to display terms to the casual user... *)
+
+val set_print_constr : (env -> Evd.evar_map -> constr -> std_ppcmds) -> unit
+val print_constr : constr -> std_ppcmds
+val print_constr_env : env -> Evd.evar_map -> constr -> std_ppcmds
+val print_named_context : env -> std_ppcmds
+val pr_rel_decl : env -> Context.Rel.Declaration.t -> std_ppcmds
+val print_rel_context : env -> std_ppcmds
+val print_env : env -> std_ppcmds
diff --git a/engine/uState.ml b/engine/uState.ml
index 146a386a2..63bd247d5 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -131,84 +131,87 @@ let instantiate_variable l b v =
exception UniversesDiffer
let process_universe_constraints ctx cstrs =
+ let open Univ in
let univs = ctx.uctx_universes in
let vars = ref ctx.uctx_univ_variables in
let normalize = Universes.normalize_universe_opt_subst vars in
- let rec unify_universes fo l d r local =
+ let is_local l = Univ.LMap.mem l !vars in
+ let varinfo x =
+ match Univ.Universe.level x with
+ | None -> Inl x
+ | Some l -> Inr l
+ in
+ let equalize_variables fo l l' r r' local =
+ (** Assumes l = [l',0] and r = [r',0] *)
+ let () =
+ if is_local l' then
+ instantiate_variable l' r vars
+ else if is_local r' then
+ instantiate_variable r' l vars
+ else if not (UGraph.check_eq_level univs l' r') then
+ (* Two rigid/global levels, none of them being local,
+ one of them being Prop/Set, disallow *)
+ if Univ.Level.is_small l' || Univ.Level.is_small r' then
+ raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
+ else if fo then
+ raise UniversesDiffer
+ in
+ Univ.enforce_eq_level l' r' local
+ in
+ let equalize_universes l r local = match varinfo l, varinfo r with
+ | Inr l', Inr r' -> equalize_variables false l l' r r' local
+ | Inr l, Inl r | Inl r, Inr l ->
+ let alg = Univ.LSet.mem l ctx.uctx_univ_algebraic in
+ let inst = Univ.univ_level_rem l r r in
+ if alg then (instantiate_variable l inst vars; local)
+ else
+ let lu = Univ.Universe.make l in
+ if Univ.univ_level_mem l r then
+ Univ.enforce_leq inst lu local
+ else raise (Univ.UniverseInconsistency (Univ.Eq, lu, r, None))
+ | Inl _, Inl _ (* both are algebraic *) ->
+ if UGraph.check_eq univs l r then local
+ else raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
+ in
+ let unify_universes (l, d, r) local =
let l = normalize l and r = normalize r in
if Univ.Universe.equal l r then local
else
- let varinfo x =
- match Univ.Universe.level x with
- | None -> Inl x
- | Some l -> Inr (l, Univ.LMap.mem l !vars, Univ.LSet.mem l ctx.uctx_univ_algebraic)
- in
- if d == Universes.ULe then
+ match d with
+ | Universes.ULe ->
if UGraph.check_leq univs l r then
(** Keep Prop/Set <= var around if var might be instantiated by prop or set
later. *)
- if Univ.Universe.is_level l then
- match Univ.Universe.level r with
- | Some r ->
- Univ.Constraint.add (Option.get (Univ.Universe.level l),Univ.Le,r) local
- | _ -> local
- else local
+ match Univ.Universe.level l, Univ.Universe.level r with
+ | Some l, Some r ->
+ Univ.Constraint.add (l, Univ.Le, r) local
+ | _ -> local
else
- match Univ.Universe.level r with
- | None -> error ("Algebraic universe on the right")
- | Some rl ->
- if Univ.Level.is_small rl then
+ begin match Univ.Universe.level r with
+ | None -> user_err Pp.(str "Algebraic universe on the right")
+ | Some r' ->
+ if Univ.Level.is_small r' then
let levels = Univ.Universe.levels l in
- Univ.LSet.fold (fun l local ->
- if Univ.Level.is_small l || Univ.LMap.mem l !vars then
- unify_universes fo (Univ.Universe.make l) Universes.UEq r local
- else raise (Univ.UniverseInconsistency (Univ.Le, Univ.Universe.make l, r, None)))
- levels local
+ let fold l' local =
+ let l = Univ.Universe.make l' in
+ if Univ.Level.is_small l' || is_local l' then
+ equalize_variables false l l' r r' local
+ else raise (Univ.UniverseInconsistency (Univ.Le, l, r, None))
+ in
+ Univ.LSet.fold fold levels local
else
Univ.enforce_leq l r local
- else if d == Universes.ULub then
- match varinfo l, varinfo r with
- | (Inr (l, true, _), Inr (r, _, _))
- | (Inr (r, _, _), Inr (l, true, _)) ->
- instantiate_variable l (Univ.Universe.make r) vars;
- Univ.enforce_eq_level l r local
- | Inr (_, _, _), Inr (_, _, _) ->
- unify_universes true l Universes.UEq r local
+ end
+ | Universes.ULub ->
+ begin match Universe.level l, Universe.level r with
+ | Some l', Some r' ->
+ equalize_variables true l l' r r' local
| _, _ -> assert false
- else (* d = Universes.UEq *)
- match varinfo l, varinfo r with
- | Inr (l', lloc, _), Inr (r', rloc, _) ->
- let () =
- if lloc then
- instantiate_variable l' r vars
- else if rloc then
- instantiate_variable r' l vars
- else if not (UGraph.check_eq univs l r) then
- (* Two rigid/global levels, none of them being local,
- one of them being Prop/Set, disallow *)
- if Univ.Level.is_small l' || Univ.Level.is_small r' then
- raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
- else
- if fo then
- raise UniversesDiffer
- in
- Univ.enforce_eq_level l' r' local
- | Inr (l, loc, alg), Inl r
- | Inl r, Inr (l, loc, alg) ->
- let inst = Univ.univ_level_rem l r r in
- if alg then (instantiate_variable l inst vars; local)
- else
- let lu = Univ.Universe.make l in
- if Univ.univ_level_mem l r then
- Univ.enforce_leq inst lu local
- else raise (Univ.UniverseInconsistency (Univ.Eq, lu, r, None))
- | _, _ (* One of the two is algebraic or global *) ->
- if UGraph.check_eq univs l r then local
- else raise (Univ.UniverseInconsistency (Univ.Eq, l, r, None))
+ end
+ | Universes.UEq -> equalize_universes l r local
in
let local =
- Universes.Constraints.fold (fun (l,d,r) local -> unify_universes false l d r local)
- cstrs Univ.Constraint.empty
+ Universes.Constraints.fold unify_universes cstrs Univ.Constraint.empty
in
!vars, local
@@ -255,8 +258,8 @@ let universe_context ?names ctx =
let l =
try UNameMap.find (Id.to_string id) (fst ctx.uctx_names)
with Not_found ->
- user_err_loc (loc, "universe_context",
- str"Universe " ++ Nameops.pr_id id ++ str" is not bound anymore.")
+ user_err ?loc ~hdr:"universe_context"
+ (str"Universe " ++ Nameops.pr_id id ++ str" is not bound anymore.")
in (l :: newinst, (id, l) :: map, Univ.LSet.remove l acc))
pl ([], [], levels)
in
@@ -266,11 +269,11 @@ let universe_context ?names ctx =
try
let info =
Univ.LMap.find (Univ.LSet.choose left) (snd ctx.uctx_names) in
- Option.default Loc.ghost info.uloc
- with Not_found -> Loc.ghost
+ info.uloc
+ with Not_found -> None
in
- user_err_loc (loc, "universe_context",
- (str(CString.plural n "Universe") ++ spc () ++
+ user_err ?loc ~hdr:"universe_context"
+ ((str(CString.plural n "Universe") ++ spc () ++
Univ.LSet.pr (pr_uctx_level ctx) left ++
spc () ++ str (CString.conjugate_verb_to_be n) ++
str" unbound."))
@@ -281,7 +284,7 @@ let universe_context ?names ctx =
in map, ctx
let restrict ctx vars =
- let uctx' = Universes.restrict_universe_context ctx.uctx_local vars in
+ let uctx' = Univops.restrict_universe_context ctx.uctx_local vars in
{ ctx with uctx_local = uctx' }
type rigid =
diff --git a/engine/universes.ml b/engine/universes.ml
new file mode 100644
index 000000000..28058aeed
--- /dev/null
+++ b/engine/universes.ml
@@ -0,0 +1,1140 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Pp
+open Names
+open Term
+open Environ
+open Univ
+open Globnames
+
+let pr_with_global_universes l =
+ try Nameops.pr_id (LMap.find l (snd (Global.global_universe_names ())))
+ with Not_found -> Level.pr l
+
+(** Local universe names of polymorphic references *)
+
+type universe_binders = (Id.t * Univ.universe_level) list
+
+let universe_binders_table = Summary.ref Refmap.empty ~name:"universe binders"
+
+let universe_binders_of_global ref =
+ try
+ let l = Refmap.find ref !universe_binders_table in l
+ with Not_found -> []
+
+let register_universe_binders ref l =
+ universe_binders_table := Refmap.add ref l !universe_binders_table
+
+(* To disallow minimization to Set *)
+
+let set_minimization = ref true
+let is_set_minimization () = !set_minimization
+
+type universe_constraint_type = ULe | UEq | ULub
+
+type universe_constraint = universe * universe_constraint_type * universe
+
+module Constraints = struct
+ module S = Set.Make(
+ struct
+ type t = universe_constraint
+
+ let compare_type c c' =
+ match c, c' with
+ | ULe, ULe -> 0
+ | ULe, _ -> -1
+ | _, ULe -> 1
+ | UEq, UEq -> 0
+ | UEq, _ -> -1
+ | ULub, ULub -> 0
+ | ULub, _ -> 1
+
+ let compare (u,c,v) (u',c',v') =
+ let i = compare_type c c' in
+ if Int.equal i 0 then
+ let i' = Universe.compare u u' in
+ if Int.equal i' 0 then Universe.compare v v'
+ else
+ if c != ULe && Universe.compare u v' = 0 && Universe.compare v u' = 0 then 0
+ else i'
+ else i
+ end)
+
+ include S
+
+ let add (l,d,r as cst) s =
+ if Universe.equal l r then s
+ else add cst s
+
+ let tr_dir = function
+ | ULe -> Le
+ | UEq -> Eq
+ | ULub -> Eq
+
+ let op_str = function ULe -> " <= " | UEq -> " = " | ULub -> " /\\ "
+
+ let pr c =
+ fold (fun (u1,op,u2) pp_std ->
+ pp_std ++ Universe.pr u1 ++ str (op_str op) ++
+ Universe.pr u2 ++ fnl ()) c (str "")
+
+ let equal x y =
+ x == y || equal x y
+
+end
+
+type universe_constraints = Constraints.t
+type 'a constraint_accumulator = universe_constraints -> 'a -> 'a option
+type 'a universe_constrained = 'a * universe_constraints
+
+type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints
+
+let enforce_eq_instances_univs strict x y c =
+ let d = if strict then ULub else UEq in
+ let ax = Instance.to_array x and ay = Instance.to_array y in
+ if Array.length ax != Array.length ay then
+ CErrors.anomaly (Pp.str "Invalid argument: enforce_eq_instances_univs called with" ++
+ Pp.str " instances of different lengths.");
+ CArray.fold_right2
+ (fun x y -> Constraints.add (Universe.make x, d, Universe.make y))
+ ax ay c
+
+let subst_univs_universe_constraint fn (u,d,v) =
+ let u' = subst_univs_universe fn u and v' = subst_univs_universe fn v in
+ if Universe.equal u' v' then None
+ else Some (u',d,v')
+
+let subst_univs_universe_constraints subst csts =
+ Constraints.fold
+ (fun c -> Option.fold_right Constraints.add (subst_univs_universe_constraint subst c))
+ csts Constraints.empty
+
+
+let to_constraints g s =
+ let tr (x,d,y) acc =
+ let add l d l' acc = Constraint.add (l,Constraints.tr_dir d,l') acc in
+ match Universe.level x, d, Universe.level y with
+ | Some l, (ULe | UEq | ULub), Some l' -> add l d l' acc
+ | _, ULe, Some l' -> enforce_leq x y acc
+ | _, ULub, _ -> acc
+ | _, d, _ ->
+ let f = if d == ULe then UGraph.check_leq else UGraph.check_eq in
+ if f g x y then acc else
+ raise (Invalid_argument
+ "to_constraints: non-trivial algebraic constraint between universes")
+ in Constraints.fold tr s Constraint.empty
+
+let test_constr_univs_infer leq univs fold m n accu =
+ if m == n then Some accu
+ else
+ let cstrs = ref accu in
+ let eq_universes strict l l' = UGraph.check_eq_instances univs l l' in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
+ match fold (Constraints.singleton (u1, UEq, u2)) !cstrs with
+ | None -> false
+ | Some accu -> cstrs := accu; true
+ in
+ let leq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
+ match fold (Constraints.singleton (u1, ULe, u2)) !cstrs with
+ | None -> false
+ | Some accu -> cstrs := accu; true
+ in
+ let rec eq_constr' m n =
+ m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ let res =
+ if leq then
+ let rec compare_leq m n =
+ Constr.compare_head_gen_leq eq_universes leq_sorts
+ eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ compare_leq m n
+ else Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ if res then Some !cstrs else None
+
+let eq_constr_univs_infer univs fold m n accu =
+ test_constr_univs_infer false univs fold m n accu
+
+let leq_constr_univs_infer univs fold m n accu =
+ test_constr_univs_infer true univs fold m n accu
+
+(** Variant of [eq_constr_univs_infer] taking kind-of-term functions,
+ to expose subterms of [m] and [n], arguments. *)
+let eq_constr_univs_infer_with kind1 kind2 univs fold m n accu =
+ (* spiwack: duplicates the code of [eq_constr_univs_infer] because I
+ haven't find a way to factor the code without destroying
+ pointer-equality optimisations in [eq_constr_univs_infer].
+ Pointer equality is not sufficient to ensure equality up to
+ [kind1,kind2], because [kind1] and [kind2] may be different,
+ typically evaluating [m] and [n] in different evar maps. *)
+ let cstrs = ref accu in
+ let eq_universes strict = UGraph.check_eq_instances univs in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
+ match fold (Constraints.singleton (u1, UEq, u2)) !cstrs with
+ | None -> false
+ | Some accu -> cstrs := accu; true
+ in
+ let rec eq_constr' m n =
+ Constr.compare_head_gen_with kind1 kind2 eq_universes eq_sorts eq_constr' m n
+ in
+ let res = Constr.compare_head_gen_with kind1 kind2 eq_universes eq_sorts eq_constr' m n in
+ if res then Some !cstrs else None
+
+let test_constr_universes leq m n =
+ if m == n then Some Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,UEq,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let leq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,ULe,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n =
+ m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ let res =
+ if leq then
+ let rec compare_leq m n =
+ Constr.compare_head_gen_leq eq_universes leq_sorts eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ compare_leq m n
+ else
+ Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n
+ in
+ if res then Some !cstrs else None
+
+let eq_constr_universes m n = test_constr_universes false m n
+let leq_constr_universes m n = test_constr_universes true m n
+
+let compare_head_gen_proj env equ eqs eqc' m n =
+ match kind_of_term m, kind_of_term n with
+ | Proj (p, c), App (f, args)
+ | App (f, args), Proj (p, c) ->
+ (match kind_of_term f with
+ | Const (p', u) when eq_constant (Projection.constant p) p' ->
+ let pb = Environ.lookup_projection p env in
+ let npars = pb.Declarations.proj_npars in
+ if Array.length args == npars + 1 then
+ eqc' c args.(npars)
+ else false
+ | _ -> false)
+ | _ -> Constr.compare_head_gen equ eqs eqc' m n
+
+let eq_constr_universes_proj env m n =
+ if m == n then true, Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1, UEq, Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen_proj env eq_universes eq_sorts eq_constr' m n
+ in
+ let res = eq_constr' m n in
+ res, !cstrs
+
+(* Generator of levels *)
+let new_univ_level, set_remote_new_univ_level =
+ RemoteCounter.new_counter ~name:"Universes" 0 ~incr:((+) 1)
+ ~build:(fun n -> Univ.Level.make (Global.current_dirpath ()) n)
+
+let new_univ_level _ = new_univ_level ()
+ (* Univ.Level.make db (new_univ_level ()) *)
+
+let fresh_level () = new_univ_level (Global.current_dirpath ())
+
+(* TODO: remove *)
+let new_univ dp = Univ.Universe.make (new_univ_level dp)
+let new_Type dp = mkType (new_univ dp)
+let new_Type_sort dp = Type (new_univ dp)
+
+let fresh_universe_instance ctx =
+ Instance.subst_fn (fun _ -> new_univ_level (Global.current_dirpath ()))
+ (AUContext.instance ctx)
+
+let fresh_instance_from_context ctx =
+ let inst = fresh_universe_instance ctx in
+ let constraints = UContext.constraints (subst_instance_context inst ctx) in
+ inst, constraints
+
+let fresh_instance ctx =
+ let ctx' = ref LSet.empty in
+ let inst =
+ Instance.subst_fn (fun v ->
+ let u = new_univ_level (Global.current_dirpath ()) in
+ ctx' := LSet.add u !ctx'; u)
+ (AUContext.instance ctx)
+ in !ctx', inst
+
+let existing_instance ctx inst =
+ let () =
+ let a1 = Instance.to_array inst
+ and a2 = Instance.to_array (AUContext.instance ctx) in
+ let len1 = Array.length a1 and len2 = Array.length a2 in
+ if not (len1 == len2) then
+ CErrors.user_err ~hdr:"Universes"
+ (str "Polymorphic constant expected " ++ int len2 ++
+ str" levels but was given " ++ int len1)
+ else ()
+ in LSet.empty, inst
+
+let fresh_instance_from ctx inst =
+ let ctx', inst =
+ match inst with
+ | Some inst -> existing_instance ctx inst
+ | None -> fresh_instance ctx
+ in
+ let constraints = UContext.constraints (subst_instance_context inst ctx) in
+ inst, (ctx', constraints)
+
+let unsafe_instance_from ctx =
+ (Univ.AUContext.instance ctx, Univ.instantiate_univ_context ctx)
+
+(** Fresh universe polymorphic construction *)
+
+let fresh_constant_instance env c inst =
+ let cb = lookup_constant c env in
+ match cb.Declarations.const_universes with
+ | Declarations.Monomorphic_const _ -> ((c,Instance.empty), ContextSet.empty)
+ | Declarations.Polymorphic_const auctx ->
+ let inst, ctx =
+ fresh_instance_from auctx inst
+ in
+ ((c, inst), ctx)
+
+let fresh_inductive_instance env ind inst =
+ let mib, mip = Inductive.lookup_mind_specif env ind in
+ match mib.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ ->
+ ((ind,Instance.empty), ContextSet.empty)
+ | Declarations.Polymorphic_ind uactx ->
+ let inst, ctx = (fresh_instance_from uactx) inst in
+ ((ind,inst), ctx)
+ | Declarations.Cumulative_ind acumi ->
+ let inst, ctx =
+ fresh_instance_from (Univ.ACumulativityInfo.univ_context acumi) inst
+ in ((ind,inst), ctx)
+
+let fresh_constructor_instance env (ind,i) inst =
+ let mib, mip = Inductive.lookup_mind_specif env ind in
+ match mib.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ -> (((ind,i),Instance.empty), ContextSet.empty)
+ | Declarations.Polymorphic_ind auctx ->
+ let inst, ctx = fresh_instance_from auctx inst in
+ (((ind,i),inst), ctx)
+ | Declarations.Cumulative_ind acumi ->
+ let inst, ctx = fresh_instance_from (ACumulativityInfo.univ_context acumi) inst in
+ (((ind,i),inst), ctx)
+
+let unsafe_constant_instance env c =
+ let cb = lookup_constant c env in
+ match cb.Declarations.const_universes with
+ | Declarations.Monomorphic_const _ ->
+ ((c,Instance.empty), UContext.empty)
+ | Declarations.Polymorphic_const auctx ->
+ let inst, ctx = unsafe_instance_from auctx in ((c, inst), ctx)
+
+let unsafe_inductive_instance env ind =
+ let mib, mip = Inductive.lookup_mind_specif env ind in
+ match mib.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ -> ((ind,Instance.empty), UContext.empty)
+ | Declarations.Polymorphic_ind auctx ->
+ let inst, ctx = unsafe_instance_from auctx in ((ind,inst), ctx)
+ | Declarations.Cumulative_ind acumi ->
+ let inst, ctx = unsafe_instance_from (ACumulativityInfo.univ_context acumi) in
+ ((ind,inst), ctx)
+
+let unsafe_constructor_instance env (ind,i) =
+ let mib, mip = Inductive.lookup_mind_specif env ind in
+ match mib.Declarations.mind_universes with
+ | Declarations.Monomorphic_ind _ -> (((ind, i),Instance.empty), UContext.empty)
+ | Declarations.Polymorphic_ind auctx ->
+ let inst, ctx = unsafe_instance_from auctx in (((ind, i),inst), ctx)
+ | Declarations.Cumulative_ind acumi ->
+ let inst, ctx = unsafe_instance_from (ACumulativityInfo.univ_context acumi) in
+ (((ind, i),inst), ctx)
+
+open Globnames
+
+let fresh_global_instance ?names env gr =
+ match gr with
+ | VarRef id -> mkVar id, ContextSet.empty
+ | ConstRef sp ->
+ let c, ctx = fresh_constant_instance env sp names in
+ mkConstU c, ctx
+ | ConstructRef sp ->
+ let c, ctx = fresh_constructor_instance env sp names in
+ mkConstructU c, ctx
+ | IndRef sp ->
+ let c, ctx = fresh_inductive_instance env sp names in
+ mkIndU c, ctx
+
+let fresh_constant_instance env sp =
+ fresh_constant_instance env sp None
+
+let fresh_inductive_instance env sp =
+ fresh_inductive_instance env sp None
+
+let fresh_constructor_instance env sp =
+ fresh_constructor_instance env sp None
+
+let unsafe_global_instance env gr =
+ match gr with
+ | VarRef id -> mkVar id, UContext.empty
+ | ConstRef sp ->
+ let c, ctx = unsafe_constant_instance env sp in
+ mkConstU c, ctx
+ | ConstructRef sp ->
+ let c, ctx = unsafe_constructor_instance env sp in
+ mkConstructU c, ctx
+ | IndRef sp ->
+ let c, ctx = unsafe_inductive_instance env sp in
+ mkIndU c, ctx
+
+let constr_of_global gr =
+ let c, ctx = fresh_global_instance (Global.env ()) gr in
+ if not (Univ.ContextSet.is_empty ctx) then
+ if Univ.LSet.is_empty (Univ.ContextSet.levels ctx) then
+ (* Should be an error as we might forget constraints, allow for now
+ to make firstorder work with "using" clauses *)
+ c
+ else CErrors.user_err ~hdr:"constr_of_global"
+ Pp.(str "globalization of polymorphic reference " ++ Nametab.pr_global_env Id.Set.empty gr ++
+ str " would forget universes.")
+ else c
+
+let constr_of_reference = constr_of_global
+
+let unsafe_constr_of_global gr =
+ unsafe_global_instance (Global.env ()) gr
+
+let constr_of_global_univ (gr,u) =
+ match gr with
+ | VarRef id -> mkVar id
+ | ConstRef sp -> mkConstU (sp,u)
+ | ConstructRef sp -> mkConstructU (sp,u)
+ | IndRef sp -> mkIndU (sp,u)
+
+let fresh_global_or_constr_instance env = function
+ | IsConstr c -> c, ContextSet.empty
+ | IsGlobal gr -> fresh_global_instance env gr
+
+let global_of_constr c =
+ match kind_of_term c with
+ | Const (c, u) -> ConstRef c, u
+ | Ind (i, u) -> IndRef i, u
+ | Construct (c, u) -> ConstructRef c, u
+ | Var id -> VarRef id, Instance.empty
+ | _ -> raise Not_found
+
+open Declarations
+
+let type_of_reference env r =
+ match r with
+ | VarRef id -> Environ.named_type id env, ContextSet.empty
+ | ConstRef c ->
+ let cb = Environ.lookup_constant c env in
+ let ty = Typeops.type_of_constant_type env cb.const_type in
+ begin
+ match cb.const_universes with
+ | Monomorphic_const _ -> ty, ContextSet.empty
+ | Polymorphic_const auctx ->
+ let inst, ctx = fresh_instance_from auctx None in
+ Vars.subst_instance_constr inst ty, ctx
+ end
+ | IndRef ind ->
+ let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in
+ begin
+ match mib.mind_universes with
+ | Monomorphic_ind _ ->
+ let ty = Inductive.type_of_inductive env (specif, Univ.Instance.empty) in
+ ty, ContextSet.empty
+ | Polymorphic_ind auctx ->
+ let inst, ctx = fresh_instance_from auctx None in
+ let ty = Inductive.type_of_inductive env (specif, inst) in
+ ty, ctx
+ | Cumulative_ind cumi ->
+ let inst, ctx =
+ fresh_instance_from (ACumulativityInfo.univ_context cumi) None
+ in
+ let ty = Inductive.type_of_inductive env (specif, inst) in
+ ty, ctx
+ end
+
+ | ConstructRef cstr ->
+ let (mib,oib as specif) =
+ Inductive.lookup_mind_specif env (inductive_of_constructor cstr)
+ in
+ begin
+ match mib.mind_universes with
+ | Monomorphic_ind _ ->
+ Inductive.type_of_constructor (cstr,Instance.empty) specif, ContextSet.empty
+ | Polymorphic_ind auctx ->
+ let inst, ctx = fresh_instance_from auctx None in
+ Inductive.type_of_constructor (cstr,inst) specif, ctx
+ | Cumulative_ind cumi ->
+ let inst, ctx =
+ fresh_instance_from (ACumulativityInfo.univ_context cumi) None
+ in
+ Inductive.type_of_constructor (cstr,inst) specif, ctx
+ end
+
+let type_of_global t = type_of_reference (Global.env ()) t
+
+let unsafe_type_of_reference env r =
+ match r with
+ | VarRef id -> Environ.named_type id env
+ | ConstRef c ->
+ let cb = Environ.lookup_constant c env in
+ Typeops.type_of_constant_type env cb.const_type
+
+ | IndRef ind ->
+ let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in
+ let (_, inst), _ = unsafe_inductive_instance env ind in
+ Inductive.type_of_inductive env (specif, inst)
+
+ | ConstructRef (ind, _ as cstr) ->
+ let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in
+ let (_, inst), _ = unsafe_inductive_instance env ind in
+ Inductive.type_of_constructor (cstr,inst) specif
+
+let unsafe_type_of_global t = unsafe_type_of_reference (Global.env ()) t
+
+let fresh_sort_in_family env = function
+ | InProp -> prop_sort, ContextSet.empty
+ | InSet -> set_sort, ContextSet.empty
+ | InType ->
+ let u = fresh_level () in
+ Type (Univ.Universe.make u), ContextSet.singleton u
+
+let new_sort_in_family sf =
+ fst (fresh_sort_in_family (Global.env ()) sf)
+
+let extend_context (a, ctx) (ctx') =
+ (a, ContextSet.union ctx ctx')
+
+let new_global_univ () =
+ let u = fresh_level () in
+ (Univ.Universe.make u, ContextSet.singleton u)
+
+(** Simplification *)
+
+module LevelUnionFind = Unionfind.Make (Univ.LSet) (Univ.LMap)
+
+let add_list_map u t map =
+ try
+ let l = LMap.find u map in
+ LMap.update u (t :: l) map
+ with Not_found ->
+ LMap.add u [t] map
+
+module UF = LevelUnionFind
+
+(** Precondition: flexible <= ctx *)
+let choose_canonical ctx flexible algs s =
+ let global = LSet.diff s ctx in
+ let flexible, rigid = LSet.partition flexible (LSet.inter s ctx) in
+ (** If there is a global universe in the set, choose it *)
+ if not (LSet.is_empty global) then
+ let canon = LSet.choose global in
+ canon, (LSet.remove canon global, rigid, flexible)
+ else (** No global in the equivalence class, choose a rigid one *)
+ if not (LSet.is_empty rigid) then
+ let canon = LSet.choose rigid in
+ canon, (global, LSet.remove canon rigid, flexible)
+ else (** There are only flexible universes in the equivalence
+ class, choose a non-algebraic. *)
+ let algs, nonalgs = LSet.partition (fun x -> LSet.mem x algs) flexible in
+ if not (LSet.is_empty nonalgs) then
+ let canon = LSet.choose nonalgs in
+ canon, (global, rigid, LSet.remove canon flexible)
+ else
+ let canon = LSet.choose algs in
+ canon, (global, rigid, LSet.remove canon flexible)
+
+let subst_univs_fn_puniverses lsubst (c, u as cu) =
+ let u' = Instance.subst_fn lsubst u in
+ if u' == u then cu else (c, u')
+
+let nf_evars_and_universes_opt_subst f subst =
+ let subst = fun l -> match LMap.find l subst with None -> raise Not_found | Some l' -> l' in
+ let lsubst = Univ.level_subst_of subst in
+ let rec aux c =
+ match kind_of_term c with
+ | Evar (evk, args) ->
+ let args = Array.map aux args in
+ (match try f (evk, args) with Not_found -> None with
+ | None -> c
+ | Some c -> aux c)
+ | Const pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkConstU pu'
+ | Ind pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkIndU pu'
+ | Construct pu ->
+ let pu' = subst_univs_fn_puniverses lsubst pu in
+ if pu' == pu then c else mkConstructU pu'
+ | Sort (Type u) ->
+ let u' = Univ.subst_univs_universe subst u in
+ if u' == u then c else mkSort (sort_of_univ u')
+ | _ -> map_constr aux c
+ in aux
+
+let fresh_universe_context_set_instance ctx =
+ if ContextSet.is_empty ctx then LMap.empty, ctx
+ else
+ let (univs, cst) = ContextSet.levels ctx, ContextSet.constraints ctx in
+ let univs',subst = LSet.fold
+ (fun u (univs',subst) ->
+ let u' = fresh_level () in
+ (LSet.add u' univs', LMap.add u u' subst))
+ univs (LSet.empty, LMap.empty)
+ in
+ let cst' = subst_univs_level_constraints subst cst in
+ subst, (univs', cst')
+
+let normalize_univ_variable ~find ~update =
+ let rec aux cur =
+ let b = find cur in
+ let b' = subst_univs_universe aux b in
+ if Universe.equal b' b then b
+ else update cur b'
+ in aux
+
+let normalize_univ_variable_opt_subst ectx =
+ let find l =
+ match Univ.LMap.find l !ectx with
+ | Some b -> b
+ | None -> raise Not_found
+ in
+ let update l b =
+ assert (match Universe.level b with Some l' -> not (Level.equal l l') | None -> true);
+ try ectx := Univ.LMap.add l (Some b) !ectx; b with Not_found -> assert false
+ in normalize_univ_variable ~find ~update
+
+let normalize_univ_variable_subst subst =
+ let find l = Univ.LMap.find l !subst in
+ let update l b =
+ assert (match Universe.level b with Some l' -> not (Level.equal l l') | None -> true);
+ try subst := Univ.LMap.update l b !subst; b with Not_found -> assert false in
+ normalize_univ_variable ~find ~update
+
+let normalize_universe_opt_subst subst =
+ let normlevel = normalize_univ_variable_opt_subst subst in
+ subst_univs_universe normlevel
+
+let normalize_universe_subst subst =
+ let normlevel = normalize_univ_variable_subst subst in
+ subst_univs_universe normlevel
+
+let normalize_opt_subst ctx =
+ let ectx = ref ctx in
+ let normalize = normalize_univ_variable_opt_subst ectx in
+ let () =
+ Univ.LMap.iter (fun u v ->
+ if Option.is_empty v then ()
+ else try ignore(normalize u) with Not_found -> assert(false)) ctx
+ in !ectx
+
+type universe_opt_subst = universe option universe_map
+
+let make_opt_subst s =
+ fun x ->
+ (match Univ.LMap.find x s with
+ | Some u -> u
+ | None -> raise Not_found)
+
+let subst_opt_univs_constr s =
+ let f = make_opt_subst s in
+ Vars.subst_univs_fn_constr f
+
+
+let normalize_univ_variables ctx =
+ let ctx = normalize_opt_subst ctx in
+ let undef, def, subst =
+ Univ.LMap.fold (fun u v (undef, def, subst) ->
+ match v with
+ | None -> (Univ.LSet.add u undef, def, subst)
+ | Some b -> (undef, Univ.LSet.add u def, Univ.LMap.add u b subst))
+ ctx (Univ.LSet.empty, Univ.LSet.empty, Univ.LMap.empty)
+ in ctx, undef, def, subst
+
+let pr_universe_body = function
+ | None -> mt ()
+ | Some v -> str" := " ++ Univ.Universe.pr v
+
+let pr_universe_opt_subst = Univ.LMap.pr pr_universe_body
+
+let compare_constraint_type d d' =
+ match d, d' with
+ | Eq, Eq -> 0
+ | Eq, _ -> -1
+ | _, Eq -> 1
+ | Le, Le -> 0
+ | Le, _ -> -1
+ | _, Le -> 1
+ | Lt, Lt -> 0
+
+type lowermap = constraint_type LMap.t
+
+let lower_union =
+ let merge k a b =
+ match a, b with
+ | Some _, None -> a
+ | None, Some _ -> b
+ | None, None -> None
+ | Some l, Some r ->
+ if compare_constraint_type l r >= 0 then a
+ else b
+ in LMap.merge merge
+
+let lower_add l c m =
+ try let c' = LMap.find l m in
+ if compare_constraint_type c c' > 0 then
+ LMap.add l c m
+ else m
+ with Not_found -> LMap.add l c m
+
+let lower_of_list l =
+ List.fold_left (fun acc (d,l) -> LMap.add l d acc) LMap.empty l
+
+exception Found of Level.t * lowermap
+let find_inst insts v =
+ try LMap.iter (fun k (enf,alg,v',lower) ->
+ if not alg && enf && Universe.equal v' v then raise (Found (k, lower)))
+ insts; raise Not_found
+ with Found (f,l) -> (f,l)
+
+let compute_lbound left =
+ (** The universe variable was not fixed yet.
+ Compute its level using its lower bound. *)
+ let sup l lbound =
+ match lbound with
+ | None -> Some l
+ | Some l' -> Some (Universe.sup l l')
+ in
+ List.fold_left (fun lbound (d, l) ->
+ if d == Le (* l <= ?u *) then sup l lbound
+ else (* l < ?u *)
+ (assert (d == Lt);
+ if not (Universe.level l == None) then
+ sup (Universe.super l) lbound
+ else None))
+ None left
+
+let instantiate_with_lbound u lbound lower alg enforce (ctx, us, algs, insts, cstrs) =
+ if enforce then
+ let inst = Universe.make u in
+ let cstrs' = enforce_leq lbound inst cstrs in
+ (ctx, us, LSet.remove u algs,
+ LMap.add u (enforce,alg,lbound,lower) insts, cstrs'),
+ (enforce, alg, inst, lower)
+ else (* Actually instantiate *)
+ (Univ.LSet.remove u ctx, Univ.LMap.add u (Some lbound) us, algs,
+ LMap.add u (enforce,alg,lbound,lower) insts, cstrs),
+ (enforce, alg, lbound, lower)
+
+type constraints_map = (Univ.constraint_type * Univ.LMap.key) list Univ.LMap.t
+
+let _pr_constraints_map (cmap:constraints_map) =
+ LMap.fold (fun l cstrs acc ->
+ Level.pr l ++ str " => " ++
+ prlist_with_sep spc (fun (d,r) -> pr_constraint_type d ++ Level.pr r) cstrs ++
+ fnl () ++ acc)
+ cmap (mt ())
+
+let remove_alg l (ctx, us, algs, insts, cstrs) =
+ (ctx, us, LSet.remove l algs, insts, cstrs)
+
+let remove_lower u lower =
+ let levels = Universe.levels u in
+ LSet.fold (fun l acc -> LMap.remove l acc) levels lower
+
+let minimize_univ_variables ctx us algs left right cstrs =
+ let left, lbounds =
+ Univ.LMap.fold (fun r lower (left, lbounds as acc) ->
+ if Univ.LMap.mem r us || not (Univ.LSet.mem r ctx) then acc
+ else (* Fixed universe, just compute its glb for sharing *)
+ let lbounds' =
+ match compute_lbound (List.map (fun (d,l) -> d, Universe.make l) lower) with
+ | None -> lbounds
+ | Some lbound -> LMap.add r (true, false, lbound, lower_of_list lower)
+ lbounds
+ in (Univ.LMap.remove r left, lbounds'))
+ left (left, Univ.LMap.empty)
+ in
+ let rec instance (ctx', us, algs, insts, cstrs as acc) u =
+ let acc, left, lower =
+ try
+ let l = LMap.find u left in
+ let acc, left, newlow, lower =
+ List.fold_left
+ (fun (acc, left', newlow, lower') (d, l) ->
+ let acc', (enf,alg,l',lower) = aux acc l in
+ let l' =
+ if enf then Universe.make l
+ else l'
+ in acc', (d, l') :: left',
+ lower_add l d newlow, lower_union lower lower')
+ (acc, [], LMap.empty, LMap.empty) l
+ in
+ let not_lower (d,l) =
+ (* We're checking if (d,l) is already implied by the lower
+ constraints on some level u. If it represents l < u (d is Lt
+ or d is Le and i > 0, the i < 0 case is impossible due to
+ invariants of Univ), and the lower constraints only have l <=
+ u then it is not implied. *)
+ Univ.Universe.exists
+ (fun (l,i) ->
+ let d =
+ if i == 0 then d
+ else match d with
+ | Le -> Lt
+ | d -> d
+ in
+ try let d' = LMap.find l lower in
+ (* If d is stronger than the already implied lower
+ * constraints we must keep it. *)
+ compare_constraint_type d d' > 0
+ with Not_found ->
+ (** No constraint existing on l *) true) l
+ in
+ let left = List.uniquize (List.filter not_lower left) in
+ (acc, left, LMap.union newlow lower)
+ with Not_found -> acc, [], LMap.empty
+ and right =
+ try Some (LMap.find u right)
+ with Not_found -> None
+ in
+ let instantiate_lbound lbound =
+ let alg = LSet.mem u algs in
+ if alg then
+ (* u is algebraic: we instantiate it with its lower bound, if any,
+ or enforce the constraints if it is bounded from the top. *)
+ let lower = remove_lower lbound lower in
+ instantiate_with_lbound u lbound lower true false acc
+ else (* u is non algebraic *)
+ match Universe.level lbound with
+ | Some l -> (* The lowerbound is directly a level *)
+ (* u is not algebraic but has no upper bounds,
+ we instantiate it with its lower bound if it is a
+ different level, otherwise we keep it. *)
+ let lower = LMap.remove l lower in
+ if not (Level.equal l u) then
+ (* Should check that u does not
+ have upper constraints that are not already in right *)
+ let acc' = remove_alg l acc in
+ instantiate_with_lbound u lbound lower false false acc'
+ else acc, (true, false, lbound, lower)
+ | None ->
+ try
+ (* Another universe represents the same lower bound,
+ we can share them with no harm. *)
+ let can, lower = find_inst insts lbound in
+ let lower = LMap.remove can lower in
+ instantiate_with_lbound u (Universe.make can) lower false false acc
+ with Not_found ->
+ (* We set u as the canonical universe representing lbound *)
+ instantiate_with_lbound u lbound lower false true acc
+ in
+ let acc' acc =
+ match right with
+ | None -> acc
+ | Some cstrs ->
+ let dangling = List.filter (fun (d, r) -> not (LMap.mem r us)) cstrs in
+ if List.is_empty dangling then acc
+ else
+ let ((ctx', us, algs, insts, cstrs), (enf,_,inst,lower as b)) = acc in
+ let cstrs' = List.fold_left (fun cstrs (d, r) ->
+ if d == Univ.Le then
+ enforce_leq inst (Universe.make r) cstrs
+ else
+ try let lev = Option.get (Universe.level inst) in
+ Constraint.add (lev, d, r) cstrs
+ with Option.IsNone -> failwith "")
+ cstrs dangling
+ in
+ (ctx', us, algs, insts, cstrs'), b
+ in
+ if not (LSet.mem u ctx) then acc' (acc, (true, false, Universe.make u, lower))
+ else
+ let lbound = compute_lbound left in
+ match lbound with
+ | None -> (* Nothing to do *)
+ acc' (acc, (true, false, Universe.make u, lower))
+ | Some lbound ->
+ try acc' (instantiate_lbound lbound)
+ with Failure _ -> acc' (acc, (true, false, Universe.make u, lower))
+ and aux (ctx', us, algs, seen, cstrs as acc) u =
+ try acc, LMap.find u seen
+ with Not_found -> instance acc u
+ in
+ LMap.fold (fun u v (ctx', us, algs, seen, cstrs as acc) ->
+ if v == None then fst (aux acc u)
+ else LSet.remove u ctx', us, LSet.remove u algs, seen, cstrs)
+ us (ctx, us, algs, lbounds, cstrs)
+
+let normalize_context_set ctx us algs =
+ let (ctx, csts) = ContextSet.levels ctx, ContextSet.constraints ctx in
+ let uf = UF.create () in
+ (** Keep the Prop/Set <= i constraints separate for minimization *)
+ let smallles, csts =
+ Constraint.fold (fun (l,d,r as cstr) (smallles, noneqs) ->
+ if d == Le then
+ if Univ.Level.is_small l then
+ if is_set_minimization () && LSet.mem r ctx then
+ (Constraint.add cstr smallles, noneqs)
+ else (smallles, noneqs)
+ else if Level.is_small r then
+ if Level.is_prop r then
+ raise (Univ.UniverseInconsistency
+ (Le,Universe.make l,Universe.make r,None))
+ else (smallles, Constraint.add (l,Eq,r) noneqs)
+ else (smallles, Constraint.add cstr noneqs)
+ else (smallles, Constraint.add cstr noneqs))
+ csts (Constraint.empty, Constraint.empty)
+ in
+ let csts =
+ (* We first put constraints in a normal-form: all self-loops are collapsed
+ to equalities. *)
+ let g = Univ.LSet.fold (fun v g -> UGraph.add_universe v false g)
+ ctx UGraph.empty_universes
+ in
+ let g =
+ Univ.Constraint.fold
+ (fun (l, d, r) g ->
+ let g =
+ if not (Level.is_small l || LSet.mem l ctx) then
+ try UGraph.add_universe l false g
+ with UGraph.AlreadyDeclared -> g
+ else g
+ in
+ let g =
+ if not (Level.is_small r || LSet.mem r ctx) then
+ try UGraph.add_universe r false g
+ with UGraph.AlreadyDeclared -> g
+ else g
+ in g) csts g
+ in
+ let g = Univ.Constraint.fold UGraph.enforce_constraint csts g in
+ UGraph.constraints_of_universes g
+ in
+ let noneqs =
+ Constraint.fold (fun (l,d,r as cstr) noneqs ->
+ if d == Eq then (UF.union l r uf; noneqs)
+ else (* We ignore the trivial Prop/Set <= i constraints. *)
+ if d == Le && Univ.Level.is_small l then noneqs
+ else if Univ.Level.is_prop l && d == Lt && Univ.Level.is_set r
+ then noneqs
+ else Constraint.add cstr noneqs)
+ csts Constraint.empty
+ in
+ let noneqs = Constraint.union noneqs smallles in
+ let partition = UF.partition uf in
+ let flex x = LMap.mem x us in
+ let ctx, subst, us, eqs = List.fold_left (fun (ctx, subst, us, cstrs) s ->
+ let canon, (global, rigid, flexible) = choose_canonical ctx flex algs s in
+ (* Add equalities for globals which can't be merged anymore. *)
+ let cstrs = LSet.fold (fun g cst ->
+ Constraint.add (canon, Univ.Eq, g) cst) global
+ cstrs
+ in
+ (* Also add equalities for rigid variables *)
+ let cstrs = LSet.fold (fun g cst ->
+ Constraint.add (canon, Univ.Eq, g) cst) rigid
+ cstrs
+ in
+ let subst = LSet.fold (fun f -> LMap.add f canon) rigid subst in
+ let subst = LSet.fold (fun f -> LMap.add f canon) flexible subst in
+ let canonu = Some (Universe.make canon) in
+ let us = LSet.fold (fun f -> LMap.add f canonu) flexible us in
+ (LSet.diff ctx flexible, subst, us, cstrs))
+ (ctx, LMap.empty, us, Constraint.empty) partition
+ in
+ (* Noneqs is now in canonical form w.r.t. equality constraints,
+ and contains only inequality constraints. *)
+ let noneqs = subst_univs_level_constraints subst noneqs in
+ (* Compute the left and right set of flexible variables, constraints
+ mentionning other variables remain in noneqs. *)
+ let noneqs, ucstrsl, ucstrsr =
+ Constraint.fold (fun (l,d,r as cstr) (noneq, ucstrsl, ucstrsr) ->
+ let lus = LMap.mem l us and rus = LMap.mem r us in
+ let ucstrsl' =
+ if lus then add_list_map l (d, r) ucstrsl
+ else ucstrsl
+ and ucstrsr' =
+ add_list_map r (d, l) ucstrsr
+ in
+ let noneqs =
+ if lus || rus then noneq
+ else Constraint.add cstr noneq
+ in (noneqs, ucstrsl', ucstrsr'))
+ noneqs (Constraint.empty, LMap.empty, LMap.empty)
+ in
+ (* Now we construct the instantiation of each variable. *)
+ let ctx', us, algs, inst, noneqs =
+ minimize_univ_variables ctx us algs ucstrsr ucstrsl noneqs
+ in
+ let us = normalize_opt_subst us in
+ (us, algs), (ctx', Constraint.union noneqs eqs)
+
+(* let normalize_conkey = Profile.declare_profile "normalize_context_set" *)
+(* let normalize_context_set a b c = Profile.profile3 normalize_conkey normalize_context_set a b c *)
+
+let simplify_universe_context (univs,csts) =
+ let uf = UF.create () in
+ let noneqs =
+ Constraint.fold (fun (l,d,r) noneqs ->
+ if d == Eq && (LSet.mem l univs || LSet.mem r univs) then
+ (UF.union l r uf; noneqs)
+ else Constraint.add (l,d,r) noneqs)
+ csts Constraint.empty
+ in
+ let partition = UF.partition uf in
+ let flex x = LSet.mem x univs in
+ let subst, univs', csts' = List.fold_left (fun (subst, univs, cstrs) s ->
+ let canon, (global, rigid, flexible) = choose_canonical univs flex LSet.empty s in
+ (* Add equalities for globals which can't be merged anymore. *)
+ let cstrs = LSet.fold (fun g cst ->
+ Constraint.add (canon, Univ.Eq, g) cst) (LSet.union global rigid)
+ cstrs
+ in
+ let subst = LSet.fold (fun f -> LMap.add f canon)
+ flexible subst
+ in (subst, LSet.diff univs flexible, cstrs))
+ (LMap.empty, univs, noneqs) partition
+ in
+ (* Noneqs is now in canonical form w.r.t. equality constraints,
+ and contains only inequality constraints. *)
+ let csts' = subst_univs_level_constraints subst csts' in
+ (univs', csts'), subst
+
+let is_trivial_leq (l,d,r) =
+ Univ.Level.is_prop l && (d == Univ.Le || (d == Univ.Lt && Univ.Level.is_set r))
+
+(* Prop < i <-> Set+1 <= i <-> Set < i *)
+let translate_cstr (l,d,r as cstr) =
+ if Level.equal Level.prop l && d == Univ.Lt && not (Level.equal Level.set r) then
+ (Level.set, d, r)
+ else cstr
+
+let refresh_constraints univs (ctx, cstrs) =
+ let cstrs', univs' =
+ Univ.Constraint.fold (fun c (cstrs', univs as acc) ->
+ let c = translate_cstr c in
+ if is_trivial_leq c then acc
+ else (Univ.Constraint.add c cstrs', UGraph.enforce_constraint c univs))
+ cstrs (Univ.Constraint.empty, univs)
+ in ((ctx, cstrs'), univs')
+
+
+(**********************************************************************)
+(* Tools for sort-polymorphic inductive types *)
+
+(* Miscellaneous functions to remove or test local univ assumed to
+ occur only in the le constraints *)
+
+(*
+ Solve a system of universe constraint of the form
+
+ u_s11, ..., u_s1p1, w1 <= u1
+ ...
+ u_sn1, ..., u_snpn, wn <= un
+
+where
+
+ - the ui (1 <= i <= n) are universe variables,
+ - the sjk select subsets of the ui for each equations,
+ - the wi are arbitrary complex universes that do not mention the ui.
+*)
+
+let is_direct_sort_constraint s v = match s with
+ | Some u -> univ_level_mem u v
+ | None -> false
+
+let solve_constraints_system levels level_bounds level_min =
+ let open Univ in
+ let levels =
+ Array.mapi (fun i o ->
+ match o with
+ | Some u ->
+ (match Universe.level u with
+ | Some u -> Some u
+ | _ -> level_bounds.(i) <- Universe.sup level_bounds.(i) u; None)
+ | None -> None)
+ levels in
+ let v = Array.copy level_bounds in
+ let nind = Array.length v in
+ let clos = Array.map (fun _ -> Int.Set.empty) levels in
+ (* First compute the transitive closure of the levels dependencies *)
+ for i=0 to nind-1 do
+ for j=0 to nind-1 do
+ if not (Int.equal i j) && is_direct_sort_constraint levels.(j) v.(i) then
+ clos.(i) <- Int.Set.add j clos.(i);
+ done;
+ done;
+ let rec closure () =
+ let continue = ref false in
+ Array.iteri (fun i deps ->
+ let deps' =
+ Int.Set.fold (fun j acc -> Int.Set.union acc clos.(j)) deps deps
+ in
+ if Int.Set.equal deps deps' then ()
+ else (clos.(i) <- deps'; continue := true))
+ clos;
+ if !continue then closure ()
+ else ()
+ in
+ closure ();
+ for i=0 to nind-1 do
+ for j=0 to nind-1 do
+ if not (Int.equal i j) && Int.Set.mem j clos.(i) then
+ (v.(i) <- Universe.sup v.(i) level_bounds.(j));
+ done;
+ done;
+ v
+
+
+(** Operations for universe_info_ind *)
+
+(** Given a universe context representing constraints of an inductive
+ this function produces a UInfoInd.t that with the trivial subtyping relation. *)
+let univ_inf_ind_from_universe_context univcst =
+ let freshunivs = Instance.of_array
+ (Array.map (fun _ -> new_univ_level ())
+ (Instance.to_array (UContext.instance univcst)))
+ in CumulativityInfo.from_universe_context univcst freshunivs
diff --git a/engine/universes.mli b/engine/universes.mli
new file mode 100644
index 000000000..5f4d212b6
--- /dev/null
+++ b/engine/universes.mli
@@ -0,0 +1,231 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Names
+open Term
+open Environ
+open Univ
+
+val set_minimization : bool ref
+val is_set_minimization : unit -> bool
+
+(** Universes *)
+
+val pr_with_global_universes : Level.t -> Pp.std_ppcmds
+
+(** Local universe name <-> level mapping *)
+
+type universe_binders = (Id.t * Univ.universe_level) list
+
+val register_universe_binders : Globnames.global_reference -> universe_binders -> unit
+val universe_binders_of_global : Globnames.global_reference -> universe_binders
+
+(** The global universe counter *)
+val set_remote_new_univ_level : universe_level RemoteCounter.installer
+
+(** Side-effecting functions creating new universe levels. *)
+
+val new_univ_level : Names.dir_path -> universe_level
+val new_univ : Names.dir_path -> universe
+val new_Type : Names.dir_path -> types
+val new_Type_sort : Names.dir_path -> sorts
+
+val new_global_univ : unit -> universe in_universe_context_set
+val new_sort_in_family : sorts_family -> sorts
+
+(** {6 Constraints for type inference}
+
+ When doing conversion of universes, not only do we have =/<= constraints but
+ also Lub constraints which correspond to unification of two levels which might
+ not be necessary if unfolding is performed.
+*)
+
+type universe_constraint_type = ULe | UEq | ULub
+
+type universe_constraint = universe * universe_constraint_type * universe
+module Constraints : sig
+ include Set.S with type elt = universe_constraint
+
+ val pr : t -> Pp.std_ppcmds
+end
+
+type universe_constraints = Constraints.t
+type 'a constraint_accumulator = universe_constraints -> 'a -> 'a option
+type 'a universe_constrained = 'a * universe_constraints
+type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints
+
+val subst_univs_universe_constraints : universe_subst_fn ->
+ universe_constraints -> universe_constraints
+
+val enforce_eq_instances_univs : bool -> universe_instance universe_constraint_function
+
+val to_constraints : UGraph.t -> universe_constraints -> constraints
+
+(** [eq_constr_univs_infer u a b] is [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping, the universe constraints in [u] and additional constraints [c]. *)
+val eq_constr_univs_infer : UGraph.t -> 'a constraint_accumulator ->
+ constr -> constr -> 'a -> 'a option
+
+(** [eq_constr_univs_infer_With kind1 kind2 univs m n] is a variant of
+ {!eq_constr_univs_infer} taking kind-of-term functions, to expose
+ subterms of [m] and [n], arguments. *)
+val eq_constr_univs_infer_with :
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ UGraph.t -> 'a constraint_accumulator -> constr -> constr -> 'a -> 'a option
+
+(** [leq_constr_univs u a b] is [true, c] if [a] is convertible to [b]
+ modulo alpha, casts, application grouping, the universe constraints
+ in [u] and additional constraints [c]. *)
+val leq_constr_univs_infer : UGraph.t -> 'a constraint_accumulator ->
+ constr -> constr -> 'a -> 'a option
+
+(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe constraints in [c]. *)
+val eq_constr_universes : constr -> constr -> universe_constraints option
+
+(** [leq_constr_universes a b] [true, c] if [a] is convertible to [b] modulo
+ alpha, casts, application grouping and the universe constraints in [c]. *)
+val leq_constr_universes : constr -> constr -> universe_constraints option
+
+(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts,
+ application grouping and the universe constraints in [c]. *)
+val eq_constr_universes_proj : env -> constr -> constr -> bool universe_constrained
+
+(** Build a fresh instance for a given context, its associated substitution and
+ the instantiated constraints. *)
+
+val fresh_instance_from_context : abstract_universe_context ->
+ universe_instance constrained
+
+val fresh_instance_from : abstract_universe_context -> universe_instance option ->
+ universe_instance in_universe_context_set
+
+val fresh_sort_in_family : env -> sorts_family ->
+ sorts in_universe_context_set
+val fresh_constant_instance : env -> constant ->
+ pconstant in_universe_context_set
+val fresh_inductive_instance : env -> inductive ->
+ pinductive in_universe_context_set
+val fresh_constructor_instance : env -> constructor ->
+ pconstructor in_universe_context_set
+
+val fresh_global_instance : ?names:Univ.Instance.t -> env -> Globnames.global_reference ->
+ constr in_universe_context_set
+
+val fresh_global_or_constr_instance : env -> Globnames.global_reference_or_constr ->
+ constr in_universe_context_set
+
+(** Get fresh variables for the universe context.
+ Useful to make tactics that manipulate constrs in universe contexts polymorphic. *)
+val fresh_universe_context_set_instance : universe_context_set ->
+ universe_level_subst * universe_context_set
+
+(** Raises [Not_found] if not a global reference. *)
+val global_of_constr : constr -> Globnames.global_reference puniverses
+
+val constr_of_global_univ : Globnames.global_reference puniverses -> constr
+
+val extend_context : 'a in_universe_context_set -> universe_context_set ->
+ 'a in_universe_context_set
+
+(** Simplification and pruning of constraints:
+ [normalize_context_set ctx us]
+
+ - Instantiate the variables in [us] with their most precise
+ universe levels respecting the constraints.
+
+ - Normalizes the context [ctx] w.r.t. equality constraints,
+ choosing a canonical universe in each equivalence class
+ (a global one if there is one) and transitively saturate
+ the constraints w.r.t to the equalities. *)
+
+module UF : Unionfind.PartitionSig with type elt = universe_level
+
+type universe_opt_subst = universe option universe_map
+
+val make_opt_subst : universe_opt_subst -> universe_subst_fn
+
+val subst_opt_univs_constr : universe_opt_subst -> constr -> constr
+
+val normalize_context_set : universe_context_set ->
+ universe_opt_subst (* The defined and undefined variables *) ->
+ universe_set (* univ variables that can be substituted by algebraics *) ->
+ (universe_opt_subst * universe_set) in_universe_context_set
+
+val normalize_univ_variables : universe_opt_subst ->
+ universe_opt_subst * universe_set * universe_set * universe_subst
+
+val normalize_univ_variable :
+ find:(universe_level -> universe) ->
+ update:(universe_level -> universe -> universe) ->
+ universe_level -> universe
+
+val normalize_univ_variable_opt_subst : universe_opt_subst ref ->
+ (universe_level -> universe)
+
+val normalize_univ_variable_subst : universe_subst ref ->
+ (universe_level -> universe)
+
+val normalize_universe_opt_subst : universe_opt_subst ref ->
+ (universe -> universe)
+
+val normalize_universe_subst : universe_subst ref ->
+ (universe -> universe)
+
+(** Create a fresh global in the global environment, without side effects.
+ BEWARE: this raises an ANOMALY on polymorphic constants/inductives:
+ the constraints should be properly added to an evd.
+ See Evd.fresh_global, Evarutil.new_global, and pf_constr_of_global for
+ the proper way to get a fresh copy of a global reference. *)
+val constr_of_global : Globnames.global_reference -> constr
+
+(** ** DEPRECATED ** synonym of [constr_of_global] *)
+val constr_of_reference : Globnames.global_reference -> constr
+
+(** [unsafe_constr_of_global gr] turns [gr] into a constr, works on polymorphic
+ references by taking the original universe instance that is not recorded
+ anywhere. The constraints are forgotten as well. DO NOT USE in new code. *)
+val unsafe_constr_of_global : Globnames.global_reference -> constr in_universe_context
+
+(** Returns the type of the global reference, by creating a fresh instance of polymorphic
+ references and computing their instantiated universe context. (side-effect on the
+ universe counter, use with care). *)
+val type_of_global : Globnames.global_reference -> types in_universe_context_set
+
+(** [unsafe_type_of_global gr] returns [gr]'s type, works on polymorphic
+ references by taking the original universe instance that is not recorded
+ anywhere. The constraints are forgotten as well.
+ USE with care. *)
+val unsafe_type_of_global : Globnames.global_reference -> types
+
+(** Full universes substitutions into terms *)
+
+val nf_evars_and_universes_opt_subst : (existential -> constr option) ->
+ universe_opt_subst -> constr -> constr
+
+val simplify_universe_context : universe_context_set ->
+ universe_context_set * universe_level_subst
+
+val refresh_constraints : UGraph.t -> universe_context_set -> universe_context_set * UGraph.t
+
+(** Pretty-printing *)
+
+val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds
+
+(** {6 Support for template polymorphism } *)
+
+val solve_constraints_system : universe option array -> universe array -> universe array ->
+ universe array
+
+(** Operations for universe_info_ind *)
+
+(** Given a universe context representing constraints of an inductive
+ this function produces a UInfoInd.t that with the trivial subtyping relation. *)
+val univ_inf_ind_from_universe_context : universe_context -> cumulativity_info